Library Flocq.Pff.Pff



Require Export Reals.
Require Export ZArith.
Require Export List.
Require Export Div2.
Require Export Even.
Require Import Psatz.

Ltac omega ::= lia.



Global Set Asymmetric Patterns.



Ltac CaseEq name :=
  generalize (refl_equal name); pattern name at -1 in |- *; case name.

Ltac Casec name := case name; clear name.

Ltac Elimc name := elim name; clear name.


Theorem lt_next : forall n m : nat, n < m -> m = S n \/ S n < m.
intros n m H'; elim H'; auto with arith.
Qed.

Theorem le_next : forall n m : nat, n <= m -> m = n \/ S n <= m.
intros n m H'; case (le_lt_or_eq _ _ H'); auto with arith.
Qed.

Theorem min_or :
 forall n m : nat, min n m = n /\ n <= m \/ min n m = m /\ m < n.
intros n; elim n; simpl in |- *; auto with arith.
intros n' Rec m; case m; simpl in |- *; auto with arith.
intros m'; elim (Rec m'); intros H'0; case H'0; clear H'0; intros H'0 H'1;
 rewrite H'0; auto with arith.
Qed.

Theorem minus_inv_lt_aux : forall n m : nat, n - m = 0 -> n - S m = 0.
intros n; elim n; simpl in |- *; auto with arith.
intros n0 H' m; case m; auto with arith.
intros H'0; discriminate.
Qed.

Theorem minus_inv_lt : forall n m : nat, m <= n -> m - n = 0.
intros n m H'; elim H'; simpl in |- *; auto with arith.
intros m0 H'0 H'1; apply minus_inv_lt_aux; auto.
Qed.

Theorem minus_le : forall m n p q : nat, m <= n -> p <= q -> m - q <= n - p.
intros m n p q H' H'0.
case (le_or_lt m q); intros H'1.
rewrite minus_inv_lt with (1 := H'1); auto with arith.
apply (fun p n m : nat => plus_le_reg_l n m p) with (p := q).
rewrite le_plus_minus_r; auto with arith.
rewrite (le_plus_minus p q); auto.
rewrite (plus_comm p).
rewrite plus_assoc_reverse.
rewrite le_plus_minus_r; auto with arith.
apply le_trans with (1 := H'); auto with arith.
apply le_trans with (1 := H'0); auto with arith.
apply le_trans with (2 := H'); auto with arith.
Qed.


Theorem convert_not_O : forall p : positive, nat_of_P p <> 0.
intros p; elim p.
intros p0 H'; unfold nat_of_P in |- *; simpl in |- *; rewrite ZL6.
generalize H'; case (nat_of_P p0); auto.
intros p0 H'; unfold nat_of_P in |- *; simpl in |- *; rewrite ZL6.
generalize H'; case (nat_of_P p0); simpl in |- *; auto.
unfold nat_of_P in |- *; simpl in |- *; auto with arith.
Qed.

Theorem inj_abs :
 forall x : Z, (0 <= x)%Z -> Z_of_nat (Z.abs_nat x) = x.
intros x; elim x; auto.
unfold Z.abs_nat in |- *.
intros p.
pattern p at 1 3 in |- *;
 rewrite <- (pred_o_P_of_succ_nat_o_nat_of_P_eq_id p).
generalize (convert_not_O p); case (nat_of_P p); simpl in |- *;
 auto with arith.
intros H'; case H'; auto.
intros n H' H'0; rewrite Pos.pred_succ; auto.
intros p H'; contradict H'; auto.
Qed.

Theorem inject_nat_convert :
 forall (p : Z) (q : positive),
 p = Zpos q -> Z_of_nat (nat_of_P q) = p.
intros p q H'; rewrite H'.
CaseEq (nat_of_P q); simpl in |- *.
elim q; unfold nat_of_P in |- *; simpl in |- *; intros;
 try discriminate.
absurd (0%Z = Zpos p0); auto.
red in |- *; intros H'0; try discriminate.
apply H; auto.
change (nat_of_P p0 = 0) in |- *.
generalize H0; rewrite ZL6; case (nat_of_P p0); simpl in |- *;
 auto; intros; try discriminate.
intros n; rewrite <- nat_of_P_o_P_of_succ_nat_eq_succ.
intros H'0; apply f_equal with (f := Zpos).
apply nat_of_P_inj; auto.
Qed.

Theorem ZleLe : forall x y : nat, (Z_of_nat x <= Z_of_nat y)%Z -> x <= y.
intros x y H'.
case (le_or_lt x y); auto with arith.
intros H'0; contradict H'; auto with zarith.
Qed.

Theorem Zcompare_EGAL :
 forall p q : Z, (p ?= q)%Z = Datatypes.Eq -> p = q.
intros p q; case p; case q; simpl in |- *; auto with arith;
 try (intros; discriminate); intros q1 p1.
intros H1; rewrite (Pcompare_Eq_eq p1 q1); auto.
unfold Pos.compare.
generalize (Pcompare_Eq_eq p1 q1);
 case (Pcompare p1 q1 Datatypes.Eq); simpl in |- *;
 intros H H1; try discriminate; rewrite H; auto.
Qed.

Theorem Zlt_Zopp : forall x y : Z, (x < y)%Z -> (- y < - x)%Z.
intros x y; case x; case y; simpl in |- *; auto with zarith; intros p p0;
 unfold Z.lt in |- *; simpl in |- *; unfold Pos.compare; rewrite <- ZC4;
 auto.
Qed.

Theorem Zle_Zopp : forall x y : Z, (x <= y)%Z -> (- y <= - x)%Z.
intros x y H'; case (Zle_lt_or_eq _ _ H'); auto with zarith.
Qed.

Theorem absolu_INR : forall n : nat, Z.abs_nat (Z_of_nat n) = n.
intros n; case n; simpl in |- *; auto with arith.
intros n0; rewrite nat_of_P_o_P_of_succ_nat_eq_succ; auto with arith.
Qed.

Theorem Zabs_absolu : forall z : Z, Z.abs z = Z_of_nat (Z.abs_nat z).
intros z; case z; simpl in |- *; auto; intros p; apply sym_equal;
 apply inject_nat_convert; auto.
Qed.

Theorem Zmin_sym : forall m n : Z, Z.min n m = Z.min m n.
intros m n; unfold Z.min in |- *.
case n; case m; simpl in |- *; auto; unfold Pos.compare.
intros p p0; rewrite (ZC4 p p0).
generalize (Pcompare_Eq_eq p0 p).
case (Pcompare p0 p Datatypes.Eq); simpl in |- *; auto.
intros H'; rewrite H'; auto.
intros p p0; rewrite (ZC4 p p0).
generalize (Pcompare_Eq_eq p0 p).
case (Pcompare p0 p Datatypes.Eq); simpl in |- *; auto.
intros H'; rewrite H'; auto.
Qed.

Theorem Zpower_nat_O : forall z : Z, Zpower_nat z 0 = Z_of_nat 1.
intros z; unfold Zpower_nat in |- *; simpl in |- *; auto.
Qed.

Theorem Zpower_nat_1 : forall z : Z, Zpower_nat z 1 = z.
intros z; unfold Zpower_nat in |- *; simpl in |- *; rewrite Zmult_1_r; auto.
Qed.

Theorem Zmin_le1 : forall z1 z2 : Z, (z1 <= z2)%Z -> Z.min z1 z2 = z1.
intros z1 z2; unfold Z.le, Z.min in |- *; case (z1 ?= z2)%Z; auto; intros H;
 contradict H; auto.
Qed.

Theorem Zmin_le2 : forall z1 z2 : Z, (z2 <= z1)%Z -> Z.min z1 z2 = z2.
intros z1 z2 H; rewrite Zmin_sym; apply Zmin_le1; auto.
Qed.

Theorem Zmin_Zle :
 forall z1 z2 z3 : Z,
 (z1 <= z2)%Z -> (z1 <= z3)%Z -> (z1 <= Z.min z2 z3)%Z.
intros z1 z2 z3 H' H'0; unfold Z.min in |- *.
case (z2 ?= z3)%Z; auto.
Qed.

Theorem Zminus_n_predm :
 forall n m : Z, Z.succ (n - m) = (n - Z.pred m)%Z.
intros n m.
unfold Z.pred in |- *; unfold Z.succ in |- *; ring.
Qed.

Theorem Zopp_Zpred_Zs : forall z : Z, (- Z.pred z)%Z = Z.succ (- z).
intros z; unfold Z.pred, Z.succ in |- *; ring.
Qed.

Theorem Zle_mult_gen :
 forall x y : Z, (0 <= x)%Z -> (0 <= y)%Z -> (0 <= x * y)%Z.
intros x y H' H'0; case (Zle_lt_or_eq _ _ H').
intros H'1; rewrite Zmult_comm; apply Zmult_gt_0_le_0_compat; auto;
 apply Z.lt_gt; auto.
intros H'1; rewrite <- H'1; simpl in |- *; auto with zarith.
Qed.

Definition Zmax : forall x_ x_ : Z, Z :=
  fun n m : Z =>
  match (n ?= m)%Z with
  | Datatypes.Eq => m
  | Datatypes.Lt => m
  | Datatypes.Gt => n
  end.

Theorem ZmaxLe1 : forall z1 z2 : Z, (z1 <= Zmax z1 z2)%Z.
intros z1 z2; unfold Zmax in |- *; CaseEq (z1 ?= z2)%Z; simpl in |- *;
 auto with zarith.
unfold Z.le in |- *; intros H; rewrite H; red in |- *; intros; discriminate.
rewrite Z.compare_lt_iff; auto with zarith.
Qed.

Theorem ZmaxSym : forall z1 z2 : Z, Zmax z1 z2 = Zmax z2 z1.
intros z1 z2; unfold Zmax in |- *; CaseEq (z1 ?= z2)%Z; CaseEq (z2 ?= z1)%Z;
 intros H1 H2; try case (Zcompare_EGAL _ _ H1); auto;
 try case (Zcompare_EGAL _ _ H2); auto; contradict H1.
case (Zcompare.Zcompare_Gt_Lt_antisym z2 z1); auto.
intros H' H'0; rewrite H'0; auto; red in |- *; intros; discriminate.
case (Zcompare.Zcompare_Gt_Lt_antisym z1 z2); auto.
intros H'; rewrite H'; auto; intros; red in |- *; intros; discriminate.
Qed.

Theorem ZmaxLe2 : forall z1 z2 : Z, (z2 <= Zmax z1 z2)%Z.
intros z1 z2; rewrite ZmaxSym; apply ZmaxLe1.
Qed.

Theorem Zeq_Zs :
 forall p q : Z, (p <= q)%Z -> (q < Z.succ p)%Z -> p = q.
intros p q H' H'0; apply Zle_antisym; auto.
apply Zlt_succ_le; auto.
Qed.

Theorem Zmin_Zmax : forall z1 z2 : Z, (Z.min z1 z2 <= Zmax z1 z2)%Z.
intros z1 z2; case (Zle_or_lt z1 z2); unfold Z.le, Z.lt, Z.min, Zmax in |- *;
 CaseEq (z1 ?= z2)%Z; auto; intros H1 H2; try rewrite H1;
 try rewrite H2; red in |- *; intros; discriminate.
Qed.

Theorem Zabs_Zmult :
 forall z1 z2 : Z, Z.abs (z1 * z2) = (Z.abs z1 * Z.abs z2)%Z.
intros z1 z2; case z1; case z2; simpl in |- *; auto with zarith.
Qed.

Theorem Zle_Zmult_comp_r :
 forall x y z : Z, (0 <= z)%Z -> (x <= y)%Z -> (x * z <= y * z)%Z.
intros x y z H' H'0; case (Zle_lt_or_eq _ _ H'); intros Zlt1.
apply Zmult_gt_0_le_compat_r; auto.
apply Z.lt_gt; auto.
rewrite <- Zlt1; repeat rewrite <- Zmult_0_r_reverse; auto with zarith.
Qed.

Theorem Zle_Zmult_comp_l :
 forall x y z : Z, (0 <= z)%Z -> (x <= y)%Z -> (z * x <= z * y)%Z.
intros x y z H' H'0; repeat rewrite (Zmult_comm z);
 apply Zle_Zmult_comp_r; auto.
Qed.


Theorem Rlt_IZR : forall p q : Z, (p < q)%Z -> (IZR p < IZR q)%R.
intros p q H; case (Rle_or_lt (IZR q) (IZR p)); auto.
intros H1; contradict H; apply Zle_not_lt.
apply le_IZR; auto.
Qed.

Theorem Rle_IZR : forall x y : Z, (x <= y)%Z -> (IZR x <= IZR y)%R.
intros x y H'.
case (Zle_lt_or_eq _ _ H'); clear H'; intros H'.
apply Rlt_le; auto with real.
now apply Rlt_IZR.
rewrite <- H'; auto with real.
Qed.

Theorem lt_Rlt : forall n m : nat, (INR n < INR m)%R -> n < m.
intros n m H'; case (le_or_lt m n); auto; intros H0; contradict H';
 auto with real.
case (le_lt_or_eq _ _ H0); intros H1; auto with real.
rewrite H1; apply Rlt_irrefl.
Qed.

Theorem INR_inv : forall n m : nat, INR n = INR m -> n = m.
intros n; elim n; auto; try rewrite S_INR.
intros m; case m; auto.
intros m' H1; contradict H1; auto.
rewrite S_INR.
apply Rlt_dichotomy_converse; left.
apply Rle_lt_0_plus_1.
apply pos_INR.
intros n' H' m; case m.
intros H'0; contradict H'0; auto.
rewrite S_INR.
apply Rlt_dichotomy_converse; right.
red in |- *; apply Rle_lt_0_plus_1.
apply pos_INR.
intros m' H'0.
rewrite (H' m'); auto.
repeat rewrite S_INR in H'0.
apply Rplus_eq_reg_l with (r := 1%R); repeat rewrite (Rplus_comm 1);
 auto with real.
Qed.

Theorem Rle_INR : forall x y : nat, x <= y -> (INR x <= INR y)%R.
intros x y H; repeat rewrite INR_IZR_INZ.
apply Rle_IZR; auto with zarith.
Qed.

Theorem le_Rle : forall n m : nat, (INR n <= INR m)%R -> n <= m.
intros n m H'; case H'; auto.
intros H'0; apply lt_le_weak; apply lt_Rlt; auto.
intros H'0; rewrite <- (INR_inv _ _ H'0); auto with arith.
Qed.

Theorem absolu_Zs :
 forall z : Z, (0 <= z)%Z -> Z.abs_nat (Z.succ z) = S (Z.abs_nat z).
intros z; case z.
3: intros p H'; contradict H'; auto with zarith.
replace (Z.succ 0) with (Z_of_nat 1).
intros H'; rewrite absolu_INR; simpl in |- *; auto.
simpl in |- *; auto.
intros p H'; rewrite <- Zpos_succ_morphism; simpl in |- *; auto with zarith.
Qed.

Theorem Zlt_next :
 forall n m : Z, (n < m)%Z -> m = Z.succ n \/ (Z.succ n < m)%Z.
intros n m H'; case (Zle_lt_or_eq (Z.succ n) m); auto with zarith.
Qed.

Theorem Zle_next :
 forall n m : Z, (n <= m)%Z -> m = n \/ (Z.succ n <= m)%Z.
intros n m H'; case (Zle_lt_or_eq _ _ H'); auto with zarith.
Qed.

Theorem inj_pred :
 forall n : nat, n <> 0 -> Z_of_nat (pred n) = Z.pred (Z_of_nat n).
intros n; case n; auto.
intros H'; contradict H'; auto.
intros n0 H'; rewrite inj_S; rewrite <- Zpred_succ; auto.
Qed.

Theorem Zle_abs : forall p : Z, (p <= Z_of_nat (Z.abs_nat p))%Z.
intros p; case p; simpl in |- *; auto with zarith; intros q;
 rewrite inject_nat_convert with (p := Zpos q);
 auto with zarith.
Qed.

Theorem lt_Zlt_inv : forall n m : nat, (Z_of_nat n < Z_of_nat m)%Z -> n < m.
intros n m H'; case (le_or_lt n m); auto.
intros H'0.
case (le_lt_or_eq _ _ H'0); auto with zarith.
intros H'1.
contradict H'.
apply Zle_not_lt; auto with zarith.
Qed.

Theorem NconvertO : forall p : positive, nat_of_P p <> 0.
intros p; elim p; unfold nat_of_P in |- *; simpl in |- *.
intros p0 H'; red in |- *; intros H'0; discriminate.
intros p0; rewrite ZL6; unfold nat_of_P in |- *.
case (Pmult_nat p0 1); simpl in |- *; auto.
red in |- *; intros H'; discriminate.
Qed.

Theorem absolu_lt_nz : forall z : Z, z <> 0%Z -> 0 < Z.abs_nat z.
intros z; case z; simpl in |- *; auto; try (intros H'; case H'; auto; fail);
 intros p; generalize (NconvertO p); auto with zarith.
Qed.

Theorem Rledouble : forall r : R, (0 <= r)%R -> (r <= INR 2 * r)%R.
intros r H'.
replace (INR 2 * r)%R with (r + r)%R; [ idtac | simpl in |- *; ring ].
pattern r at 1 in |- *; replace r with (r + 0)%R; [ idtac | ring ].
apply Rplus_le_compat_l; auto.
Qed.

Theorem Rltdouble : forall r : R, (0 < r)%R -> (r < INR 2 * r)%R.
intros r H'.
pattern r at 1 in |- *; replace r with (r + 0)%R; try ring.
replace (INR 2 * r)%R with (r + r)%R; simpl in |- *; try ring; auto with real.
Qed.

Theorem Rle_Rinv : forall x y : R, (0 < x)%R -> (x <= y)%R -> (/ y <= / x)%R.
intros x y H H1; case H1; intros H2.
left; apply Rinv_lt_contravar; auto.
apply Rmult_lt_0_compat; auto.
apply Rlt_trans with (2 := H2); auto.
rewrite H2; auto with real.
Qed.

Theorem Zlt_Rlt : forall z1 z2 : Z, (IZR z1 < IZR z2)%R -> (z1 < z2)%Z.
intros z1 z2 H; case (Zle_or_lt z2 z1); auto.
intros H1; contradict H; auto with real zarith.
apply Rle_not_lt, Rle_IZR; auto with real zarith.
Qed.

Theorem Zle_Rle :
 forall z1 z2 : Z, (IZR z1 <= IZR z2)%R -> (z1 <= z2)%Z.
intros z1 z2 H; case (Zle_or_lt z1 z2); auto.
intros H1; contradict H; auto with real zarith.
apply Rlt_not_le, Rlt_IZR; auto with real zarith.
Qed.

Theorem Zabs_eq_opp : forall x, (x <= 0)%Z -> Z.abs x = (- x)%Z.
intros x; case x; simpl in |- *; auto.
intros p H; contradict H; auto with zarith.
Qed.

Theorem Zabs_Zs : forall z : Z, (Z.abs (Z.succ z) <= Z.succ (Z.abs z))%Z.
intros z; case z; auto.
simpl in |- *; auto with zarith.
repeat rewrite Z.abs_eq; auto with zarith.
intros p; rewrite Zabs_eq_opp; auto with zarith.
Qed.

Theorem Zle_Zpred : forall x y : Z, (x < y)%Z -> (x <= Z.pred y)%Z.
intros x y H; apply Zlt_succ_le.
rewrite <- Zsucc_pred; auto.
Qed.

Theorem Zabs_Zopp : forall z : Z, Z.abs (- z) = Z.abs z.
intros z; case z; simpl in |- *; auto.
Qed.

Theorem Zle_Zabs : forall z : Z, (z <= Z.abs z)%Z.
intros z; case z; simpl in |- *; red in |- *; simpl in |- *; auto;
 try (red in |- *; intros; discriminate; fail).
intros p; elim p; simpl in |- *; auto;
 try (red in |- *; intros; discriminate; fail).
Qed.

Theorem Zlt_mult_simpl_l :
 forall a b c : Z, (0 < c)%Z -> (c * a < c * b)%Z -> (a < b)%Z.
intros a b0 c H H0; apply Z.gt_lt.
apply Zmult_gt_reg_r with (p := c); try apply Z.lt_gt; auto with zarith.
Qed.

Fixpoint pos_eq_bool (a b : positive) {struct b} : bool :=
  match a, b with
  | xH, xH => true
  | xI a', xI b' => pos_eq_bool a' b'
  | xO a', xO b' => pos_eq_bool a' b'
  | _, _ => false
  end.

Theorem pos_eq_bool_correct :
 forall p q : positive,
 match pos_eq_bool p q with
 | true => p = q
 | false => p <> q
 end.
intros p q; generalize p; elim q; simpl in |- *; auto; clear p q.
intros p Rec q; case q; simpl in |- *;
 try (intros; red in |- *; intros; discriminate; fail).
intros q'; generalize (Rec q'); case (pos_eq_bool q' p); simpl in |- *; auto.
intros H1; rewrite H1; auto.
intros H1; contradict H1; injection H1; auto.
intros p Rec q; case q; simpl in |- *;
 try (intros; red in |- *; intros; discriminate; fail).
intros q'; generalize (Rec q'); case (pos_eq_bool q' p); simpl in |- *; auto.
intros H1; rewrite H1; auto.
intros H1; contradict H1; injection H1; auto.
intros q; case q; simpl in |- *;
 try (intros; red in |- *; intros; discriminate; fail);
 auto.
Qed.

Definition Z_eq_bool a b :=
  match a, b with
  | Z0, Z0 => true
  | Zpos a', Zpos b' => pos_eq_bool a' b'
  | Zneg a', Zneg b' => pos_eq_bool a' b'
  | _, _ => false
  end.

Theorem Z_eq_bool_correct :
 forall p q : Z,
 match Z_eq_bool p q with
 | true => p = q
 | false => p <> q
 end.
intros p q; case p; case q; simpl in |- *; auto;
 try (intros; red in |- *; intros; discriminate; fail).
intros p' q'; generalize (pos_eq_bool_correct q' p');
 case (pos_eq_bool q' p'); simpl in |- *; auto.
intros H1; rewrite H1; auto.
intros H1; contradict H1; injection H1; auto.
intros p' q'; generalize (pos_eq_bool_correct q' p');
 case (pos_eq_bool q' p'); simpl in |- *; auto.
intros H1; rewrite H1; auto.
intros H1; contradict H1; injection H1; auto.
Qed.

Theorem Zlt_mult_ZERO :
 forall x y : Z, (0 < x)%Z -> (0 < y)%Z -> (0 < x * y)%Z.
intros x y; case x; case y; unfold Z.lt in |- *; simpl in |- *; auto.
Qed.

Theorem Zle_Zminus_ZERO :
 forall z1 z2 : Z, (z2 <= z1)%Z -> (0 <= z1 - z2)%Z.
intros z1 z2; rewrite (Zminus_diag_reverse z2); auto with zarith.
Qed.

Theorem Zle_Zpred_Zpred :
 forall z1 z2 : Z, (z1 <= z2)%Z -> (Z.pred z1 <= Z.pred z2)%Z.
intros z1 z2 H; apply Zsucc_le_reg.
repeat rewrite <- Zsucc_pred; auto.
Qed.

Theorem Zle_ZERO_Zabs : forall z : Z, (0 <= Z.abs z)%Z.
intros z; case z; simpl in |- *; auto with zarith.
Qed.

Theorem Zlt_Zabs_inv1 :
 forall z1 z2 : Z, (Z.abs z1 < z2)%Z -> (- z2 < z1)%Z.
intros z1 z2 H; case (Zle_or_lt 0 z1); intros H1.
apply Z.lt_le_trans with (- (0))%Z; auto with zarith.
rewrite <- (Z.opp_involutive z1); rewrite <- (Zabs_eq_opp z1);
 auto with zarith.
Qed.

Theorem Zlt_Zabs_inv2 :
 forall z1 z2 : Z, (Z.abs z1 < Z.abs z2)%Z -> (z1 < Z.abs z2)%Z.
intros z1 z2; case (Zle_or_lt 0 z1); intros H1.
rewrite Z.abs_eq; omega.
rewrite (Zabs_eq_opp z1); omega.
Qed.

Theorem Zle_Zabs_inv1 :
 forall z1 z2 : Z, (Z.abs z1 <= z2)%Z -> (- z2 <= z1)%Z.
intros z1 z2 H; case (Zle_or_lt 0 z1); intros H1.
apply Z.le_trans with (- (0))%Z; auto with zarith.
rewrite <- (Z.opp_involutive z1); rewrite <- (Zabs_eq_opp z1);
 auto with zarith.
Qed.

Theorem Zle_Zabs_inv2 :
 forall z1 z2 : Z, (Z.abs z1 <= z2)%Z -> (z1 <= z2)%Z.
intros z1 z2 H; case (Zle_or_lt 0 z1); intros H1.
rewrite <- (Z.abs_eq z1); auto.
apply Z.le_trans with (Z.abs z1); auto with zarith.
Qed.

Theorem Zlt_Zabs_Zpred :
 forall z1 z2 : Z,
 (Z.abs z1 < z2)%Z -> z1 <> Z.pred z2 -> (Z.abs (Z.succ z1) < z2)%Z.
intros z1 z2 H H0; case (Zle_or_lt 0 z1); intros H1.
rewrite Z.abs_eq; auto with zarith.
repeat rewrite Zabs_eq_opp; auto with zarith.
Qed.

Theorem Zle_n_Zpred :
 forall z1 z2 : Z, (Z.pred z1 <= Z.pred z2)%Z -> (z1 <= z2)%Z.
intros z1 z2 H; rewrite (Zsucc_pred z1); rewrite (Zsucc_pred z2);
 auto with zarith.
Qed.

Theorem Zpred_Zopp_Zs : forall z : Z, Z.pred (- z) = (- Z.succ z)%Z.
intros z; unfold Z.pred, Z.succ in |- *; ring.
Qed.

Theorem Zlt_1_O : forall z : Z, (1 <= z)%Z -> (0 < z)%Z.
intros z H; apply Zsucc_lt_reg; simpl in |- *; auto with zarith.
Qed.

Theorem Zlt_not_eq : forall p q : Z, (p < q)%Z -> p <> q.
intros p q H; contradict H; rewrite H; auto with zarith.
Qed.

Theorem Zlt_not_eq_rev : forall p q : Z, (q < p)%Z -> p <> q.
intros p q H; contradict H; rewrite H; auto with zarith.
Qed.

Theorem Zle_Zpred_inv :
 forall z1 z2 : Z, (z1 <= Z.pred z2)%Z -> (z1 < z2)%Z.
intros z1 z2 H; rewrite (Zsucc_pred z2); auto with zarith.
Qed.

Theorem Zabs_intro :
 forall (P : Z -> Prop) (z : Z), P (- z)%Z -> P z -> P (Z.abs z).
intros P z; case z; simpl in |- *; auto.
Qed.

Theorem Zpred_Zle_Zabs_intro :
 forall z1 z2 : Z,
 (- Z.pred z2 <= z1)%Z -> (z1 <= Z.pred z2)%Z -> (Z.abs z1 < z2)%Z.
intros z1 z2 H H0; apply Zle_Zpred_inv.
apply Zabs_intro with (P := fun x => (x <= Z.pred z2)%Z); auto with zarith.
Qed.

Theorem Zlt_ZERO_Zle_ONE : forall z : Z, (0 < z)%Z -> (1 <= z)%Z.
intros z H; replace 1%Z with (Z.succ 0); auto with zarith; simpl in |- *; auto.
Qed.

Theorem Zlt_Zabs_intro :
 forall z1 z2 : Z, (- z2 < z1)%Z -> (z1 < z2)%Z -> (Z.abs z1 < z2)%Z.
intros z1 z2; case z1; case z2; simpl in |- *; auto with zarith.
Qed.

Section Pdigit.
Variable n : Z.
Hypothesis nMoreThan1 : (1 < n)%Z.

Let nMoreThanOne := Zlt_1_O _ (Zlt_le_weak _ _ nMoreThan1).

Theorem Zpower_nat_less : forall q : nat, (0 < Zpower_nat n q)%Z.
intros q; elim q; simpl in |- *; auto with zarith.
intros; apply Z.mul_pos_pos; auto with zarith.
Qed.

Theorem Zpower_nat_monotone_S :
 forall p : nat, (Zpower_nat n p < Zpower_nat n (S p))%Z.
intros p; rewrite <- (Zmult_1_l (Zpower_nat n p)); replace (S p) with (1 + p);
 [ rewrite Zpower_nat_is_exp | auto with zarith ].
rewrite Zpower_nat_1; auto with zarith.
apply Zmult_gt_0_lt_compat_r; auto with zarith.
apply Z.lt_gt; auto with zarith.
apply Zpower_nat_less.
Qed.

Theorem Zpower_nat_monotone_lt :
 forall p q : nat, p < q -> (Zpower_nat n p < Zpower_nat n q)%Z.
intros p q H'; elim H'; simpl in |- *; auto.
apply Zpower_nat_monotone_S.
intros m H H0; apply Z.lt_trans with (1 := H0).
apply Zpower_nat_monotone_S.
Qed.

Theorem Zpower_nat_anti_monotone_lt :
 forall p q : nat, (Zpower_nat n p < Zpower_nat n q)%Z -> p < q.
intros p q H'.
case (le_or_lt q p); auto; (intros H'1; generalize H'; case H'1).
intros H'0; contradict H'0; auto with zarith.
intros m H'0 H'2; contradict H'2; auto with zarith.
apply Zle_not_lt, Zlt_le_weak, Zpower_nat_monotone_lt; auto with zarith.
Qed.

Theorem Zpower_nat_monotone_le :
 forall p q : nat, p <= q -> (Zpower_nat n p <= Zpower_nat n q)%Z.
intros p q H'; case (le_lt_or_eq _ _ H'); auto with zarith.
intros H1; apply Zlt_le_weak, Zpower_nat_monotone_lt; auto with zarith.
intros H1; rewrite H1; auto with zarith.
Qed.


Fixpoint digitAux (v r : Z) (q : positive) {struct q} : nat :=
  match q with
  | xH => 0
  | xI q' =>
      match (n * r)%Z with
      | r' =>
          match (r ?= v)%Z with
          | Datatypes.Gt => 0
          | _ => S (digitAux v r' q')
          end
      end
  | xO q' =>
      match (n * r)%Z with
      | r' =>
          match (r ?= v)%Z with
          | Datatypes.Gt => 0
          | _ => S (digitAux v r' q')
          end
      end
  end.

Definition digit (q : Z) :=
  match q with
  | Z0 => 0
  | Zpos q' => digitAux (Z.abs q) 1 (xO q')
  | Zneg q' => digitAux (Z.abs q) 1 (xO q')
  end.

Theorem digitAux1 :
 forall p r, (Zpower_nat n (S p) * r)%Z = (Zpower_nat n p * (n * r))%Z.
intros p r; replace (S p) with (1 + p);
 [ rewrite Zpower_nat_is_exp | auto with arith ].
rewrite Zpower_nat_1; ring.
Qed.

Theorem Zcompare_correct :
 forall p q : Z,
 match (p ?= q)%Z with
 | Datatypes.Gt => (q < p)%Z
 | Datatypes.Lt => (p < q)%Z
 | Datatypes.Eq => p = q
 end.
intros p q; unfold Z.lt in |- *; generalize (Zcompare_EGAL p q);
 (CaseEq (p ?= q)%Z; simpl in |- *; auto).
intros H H0; case (Zcompare_Gt_Lt_antisym p q); auto.
Qed.

Theorem digitAuxLess :
 forall (v r : Z) (q : positive),
 match digitAux v r q with
 | S r' => (Zpower_nat n r' * r <= v)%Z
 | O => True
 end.
intros v r q; generalize r; elim q; clear r q; simpl in |- *; auto.
intros q' Rec r; generalize (Zcompare_correct r v); case (r ?= v)%Z; auto.
intros H1; generalize (Rec (n * r)%Z); case (digitAux v (n * r) q').
intros; rewrite H1; rewrite Zpower_nat_O; auto with zarith.
intros r'; rewrite digitAux1; auto.
intros H1; generalize (Rec (n * r)%Z); case (digitAux v (n * r) q').
intros; rewrite Zpower_nat_O; auto with zarith.
intros r'; rewrite digitAux1; auto.
intros q' Rec r; generalize (Zcompare_correct r v); case (r ?= v)%Z; auto.
intros H1; generalize (Rec (n * r)%Z); case (digitAux v (n * r) q').
intros; rewrite H1; rewrite Zpower_nat_O; auto with zarith.
intros r'; rewrite digitAux1; auto.
intros H1; generalize (Rec (n * r)%Z); case (digitAux v (n * r) q').
intros; rewrite Zpower_nat_O; auto with zarith.
intros r'; rewrite digitAux1; auto.
Qed.

Theorem digitLess :
 forall q : Z, q <> 0%Z -> (Zpower_nat n (pred (digit q)) <= Z.abs q)%Z.
intros q; case q.
intros H; contradict H; auto with zarith.
intros p H; unfold digit in |- *;
 generalize (digitAuxLess (Z.abs (Zpos p)) 1 (xO p));
 case (digitAux (Z.abs (Zpos p)) 1 (xO p)); simpl in |- *;
 auto with zarith.
intros p H; unfold digit in |- *;
 generalize (digitAuxLess (Z.abs (Zneg p)) 1 (xO p));
 case (digitAux (Z.abs (Zneg p)) 1 (xO p)); simpl in |- *;
 auto with zarith.
Qed.

Fixpoint pos_length (p : positive) : nat :=
  match p with
  | xH => 0
  | xO p' => S (pos_length p')
  | xI p' => S (pos_length p')
  end.

Theorem digitAuxMore :
 forall (v r : Z) (q : positive),
 (0 < r)%Z ->
 (v < Zpower_nat n (pos_length q) * r)%Z ->
 (v < Zpower_nat n (digitAux v r q) * r)%Z.
intros v r q; generalize r; elim q; clear r q; simpl in |- *.
intros p Rec r Hr; generalize (Zcompare_correct r v); case (r ?= v)%Z; auto.
intros H1 H2; rewrite <- H1.
apply Z.le_lt_trans with (Zpower_nat n 0 * r)%Z; auto with zarith arith.
rewrite Zpower_nat_O; rewrite Zmult_1_l; auto with zarith.
apply Zmult_lt_compat_r; auto with zarith.
apply Zpower_nat_monotone_lt; auto with zarith.
intros H1 H2; rewrite digitAux1.
apply Rec.
apply Zlt_mult_ZERO; auto with zarith.
rewrite <- digitAux1; auto.
rewrite Zpower_nat_O; rewrite Zmult_1_l; auto with zarith.
intros p Rec r Hr; generalize (Zcompare_correct r v); case (r ?= v)%Z; auto.
intros H1 H2; rewrite <- H1.
apply Z.le_lt_trans with (Zpower_nat n 0 * r)%Z; auto with zarith arith.
rewrite Zpower_nat_O; rewrite Zmult_1_l; auto with zarith.
apply Zmult_lt_compat_r; auto with zarith.
apply Zpower_nat_monotone_lt; auto with zarith.
intros H1 H2; rewrite digitAux1.
apply Rec.
apply Zlt_mult_ZERO; auto with zarith.
rewrite <- digitAux1; auto.
rewrite Zpower_nat_O; rewrite Zmult_1_l; auto with zarith.
auto.
Qed.

Theorem pos_length_pow :
 forall p : positive, (Zpos p < Zpower_nat n (S (pos_length p)))%Z.
intros p; elim p; simpl in |- *; auto.
intros p0 H; rewrite Zpos_xI.
apply Z.lt_le_trans with (2 * (n * Zpower_nat n (pos_length p0)))%Z;
auto with zarith.
intros p0 H; rewrite Zpos_xO.
apply Z.lt_le_trans with (2 * (n * Zpower_nat n (pos_length p0)))%Z;
auto with zarith.
auto with zarith.
Qed.

Theorem digitMore : forall q : Z, (Z.abs q < Zpower_nat n (digit q))%Z.
intros q; case q.
easy.
intros q'; rewrite <- (Zmult_1_r (Zpower_nat n (digit (Zpos q')))).
unfold digit in |- *; apply digitAuxMore; auto with zarith.
rewrite Zmult_1_r.
simpl in |- *; apply pos_length_pow.
intros q'; rewrite <- (Zmult_1_r (Zpower_nat n (digit (Zneg q')))).
unfold digit in |- *; apply digitAuxMore; auto with zarith.
rewrite Zmult_1_r.
simpl in |- *; apply pos_length_pow.
Qed.


Theorem digitInv :
 forall (q : Z) (r : nat),
 (Zpower_nat n (pred r) <= Z.abs q)%Z ->
 (Z.abs q < Zpower_nat n r)%Z -> digit q = r.
intros q r H' H'0; case (le_or_lt (digit q) r).
intros H'1; case (le_lt_or_eq _ _ H'1); auto; intros H'2.
absurd (Z.abs q < Zpower_nat n (digit q))%Z; auto with zarith.
apply Zle_not_lt; auto with zarith.
apply Z.le_trans with (m := Zpower_nat n (pred r)); auto with zarith.
apply Zpower_nat_monotone_le.
generalize H'2; case r; auto with arith.
apply digitMore.
intros H'1.
absurd (Zpower_nat n (pred (digit q)) <= Z.abs q)%Z; auto with zarith.
apply Zlt_not_le; auto with zarith.
apply Z.lt_le_trans with (m := Zpower_nat n r); auto.
apply Zpower_nat_monotone_le.
generalize H'1; case (digit q); auto with arith.
apply digitLess; auto with zarith.
generalize H'1; case q; unfold digit in |- *; intros tmp; intros; red in |- *;
 intros; try discriminate; contradict tmp; auto with arith.
Qed.


Theorem digit_monotone :
 forall p q : Z, (Z.abs p <= Z.abs q)%Z -> digit p <= digit q.
intros p q H; case (le_or_lt (digit p) (digit q)); auto; intros H1;
 contradict H.
apply Zlt_not_le.
cut (p <> 0%Z); [ intros H2 | idtac ].
apply Z.lt_le_trans with (2 := digitLess p H2).
cut (digit q <= pred (digit p)); [ intros H3 | idtac ].
apply Z.lt_le_trans with (2 := Zpower_nat_monotone_le _ _ H3);
 auto with zarith.
apply digitMore.
generalize H1; case (digit p); simpl in |- *; auto with arith.
generalize H1; case p; simpl in |- *; intros tmp; intros; red in |- *; intros;
 try discriminate; contradict tmp; auto with arith.
Qed.


Theorem digitNotZero : forall q : Z, q <> 0%Z -> 0 < digit q.
intros q H'.
apply lt_le_trans with (m := digit 1); auto with zarith.
apply digit_monotone.
generalize H'; case q; simpl in |- *; auto with zarith; intros q'; case q';
 simpl in |- *; auto with zarith arith; intros; red in |- *;
 simpl in |- *; red in |- *; intros; discriminate.
Qed.

Theorem digitAdd :
 forall (q : Z) (r : nat),
 q <> 0%Z -> digit (q * Zpower_nat n r) = digit q + r.
intros q r H0.
apply digitInv.
replace (pred (digit q + r)) with (pred (digit q) + r).
rewrite Zpower_nat_is_exp; rewrite Zabs_Zmult;
 rewrite (fun x => Z.abs_eq (Zpower_nat n x)); auto with zarith arith.
apply Zmult_le_compat_r.
apply digitLess; auto with zarith.
apply Zpower_NR0; auto with zarith.
apply Zpower_NR0; auto with zarith.
generalize (digitNotZero _ H0); case (digit q); auto with arith.
intros H'; contradict H'; auto with arith.
rewrite Zpower_nat_is_exp; rewrite Zabs_Zmult;
 rewrite (fun x => Z.abs_eq (Zpower_nat n x)); auto with zarith arith.
apply Zmult_lt_compat_r.
apply Zpower_nat_less; auto with zarith.
apply digitMore; auto with zarith.
apply Zpower_NR0; auto with zarith.
Qed.

Theorem digit_abs : forall p : Z, digit (Z.abs p) = digit p.
intros p; case p; simpl in |- *; auto.
Qed.

Theorem digit_anti_monotone_lt :
 (1 < n)%Z -> forall p q : Z, digit p < digit q -> (Z.abs p < Z.abs q)%Z.
intros H' p q H'0.
case (Zle_or_lt (Z.abs q) (Z.abs p)); auto; intros H'1.
contradict H'0.
case (Zle_lt_or_eq _ _ H'1); intros H'2.
apply le_not_lt; auto with arith.
now apply digit_monotone.
rewrite <- (digit_abs p); rewrite <- (digit_abs q); rewrite H'2;
 auto with arith.
Qed.
End Pdigit.



Theorem pow_NR0 : forall (e : R) (n : nat), e <> 0%R -> (e ^ n)%R <> 0%R.
intros e n; elim n; simpl in |- *; auto with real.
Qed.

Theorem pow_add :
 forall (e : R) (n m : nat), (e ^ (n + m))%R = (e ^ n * e ^ m)%R.
intros e n; elim n; simpl in |- *; auto with real.
intros n0 H' m; rewrite H'; auto with real.
Qed.

Theorem pow_RN_plus :
 forall (e : R) (n m : nat),
 e <> 0%R -> (e ^ n)%R = (e ^ (n + m) * / e ^ m)%R.
intros e n; elim n; simpl in |- *; auto with real.
intros n0 H' m H'0.
rewrite Rmult_assoc; rewrite <- H'; auto.
Qed.

Theorem pow_lt : forall (e : R) (n : nat), (0 < e)%R -> (0 < e ^ n)%R.
intros e n; elim n; simpl in |- *; auto with real.
intros n0 H' H'0; replace 0%R with (e * 0)%R; auto with real.
Qed.

Theorem Rlt_pow_R1 :
 forall (e : R) (n : nat), (1 < e)%R -> 0 < n -> (1 < e ^ n)%R.
intros e n; elim n; simpl in |- *; auto with real.
intros H' H'0; contradict H'0; auto with arith.
intros n0; case n0.
simpl in |- *; rewrite Rmult_1_r; auto.
intros n1 H' H'0 H'1.
replace 1%R with (1 * 1)%R; auto with real.
apply Rlt_trans with (r2 := (e * 1)%R); auto with real.
apply Rmult_lt_compat_l; auto with real.
apply Rlt_trans with (r2 := 1%R); auto with real.
apply H'; auto with arith.
Qed.

Theorem Rlt_pow :
 forall (e : R) (n m : nat), (1 < e)%R -> n < m -> (e ^ n < e ^ m)%R.
intros e n m H' H'0; replace m with (m - n + n).
rewrite pow_add.
pattern (e ^ n)%R at 1 in |- *; replace (e ^ n)%R with (1 * e ^ n)%R;
 auto with real.
apply Rminus_lt.
repeat rewrite (fun x : R => Rmult_comm x (e ^ n));
 rewrite <- Rmult_minus_distr_l.
replace 0%R with (e ^ n * 0)%R; auto with real.
apply Rmult_lt_compat_l; auto with real.
apply pow_lt; auto with real.
apply Rlt_trans with (r2 := 1%R); auto with real.
apply Rlt_minus; auto with real.
apply Rlt_pow_R1; auto with arith.
apply plus_lt_reg_l with (p := n); auto with arith.
rewrite le_plus_minus_r; auto with arith; rewrite <- plus_n_O; auto.
rewrite plus_comm; auto with arith.
Qed.

Theorem pow_R1 :
 forall (r : R) (n : nat), (r ^ n)%R = 1%R -> Rabs r = 1%R \/ n = 0.
intros r n H'.
case (Req_dec (Rabs r) 1); auto; intros H'1.
case (Rdichotomy _ _ H'1); intros H'2.
generalize H'; case n; auto.
intros n0 H'0.
cut (r <> 0%R); [ intros Eq1 | idtac ].
2: contradict H'0; auto with arith.
2: simpl in |- *; rewrite H'0; rewrite Rmult_0_l; auto with real.
cut (Rabs r <> 0%R); [ intros Eq2 | apply Rabs_no_R0 ]; auto.
absurd (Rabs (/ r) ^ 0 < Rabs (/ r) ^ S n0)%R; auto.
replace (Rabs (/ r) ^ S n0)%R with 1%R.
simpl in |- *; apply Rlt_irrefl; auto.
rewrite Rabs_Rinv; auto.
rewrite <- Rinv_pow; auto.
rewrite RPow_abs; auto.
rewrite H'0; rewrite Rabs_right; auto with real.
apply Rlt_pow; auto with arith.
rewrite Rabs_Rinv; auto.
apply Rmult_lt_reg_l with (r := Rabs r).
case (Rabs_pos r); auto.
intros H'3; case Eq2; auto.
rewrite Rmult_1_r; rewrite Rinv_r; auto with real.
generalize H'; case n; auto.
intros n0 H'0.
cut (r <> 0%R); [ intros Eq1 | auto with real ].
2: contradict H'0; simpl in |- *; rewrite H'0; rewrite Rmult_0_l;
    auto with real.
cut (Rabs r <> 0%R); [ intros Eq2 | apply Rabs_no_R0 ]; auto.
absurd (Rabs r ^ 0 < Rabs r ^ S n0)%R; auto with real arith.
repeat rewrite RPow_abs; rewrite H'0; simpl in |- *; auto with real.
Qed.

Theorem Zpower_NR0 :
 forall (e : Z) (n : nat), (0 <= e)%Z -> (0 <= Zpower_nat e n)%Z.
intros e n; elim n; unfold Zpower_nat in |- *; simpl in |- *;
 auto with zarith.
Qed.

Theorem Zpower_NR1 :
 forall (e : Z) (n : nat), (1 <= e)%Z -> (1 <= Zpower_nat e n)%Z.
intros e n; elim n; simpl in |- *;
 auto with zarith.
intros m H1 H2; rewrite <- (Zmult_1_l 1).
apply Z.mul_le_mono_nonneg; auto with zarith.
Qed.



Theorem powerRZ_O : forall e : R, powerRZ e 0 = 1%R.
simpl in |- *; auto.
Qed.

Theorem powerRZ_1 : forall e : R, powerRZ e (Z.succ 0) = e.
simpl in |- *; auto with real.
Qed.

Theorem powerRZ_NOR : forall (e : R) (z : Z), e <> 0%R -> powerRZ e z <> 0%R.
intros e z; case z; simpl in |- *; auto with real.
Qed.

Theorem powerRZ_add :
 forall (e : R) (n m : Z),
 e <> 0%R -> powerRZ e (n + m) = (powerRZ e n * powerRZ e m)%R.
intros e n m; case n; case m; simpl in |- *; auto with real.
intros n1 m1; rewrite nat_of_P_plus_morphism; auto with real.
intros n1 m1. rewrite Z.pos_sub_spec; unfold Pos.compare.
CaseEq (Pcompare m1 n1 Datatypes.Eq); simpl in |- *; auto with real.
intros H' H'0; rewrite Pcompare_Eq_eq with (1 := H'); auto with real.
intros H' H'0; rewrite (nat_of_P_minus_morphism n1 m1); auto with real.
rewrite (pow_RN_plus e (nat_of_P n1 - nat_of_P m1) (nat_of_P m1));
 auto with real.
rewrite plus_comm; rewrite le_plus_minus_r; auto with real.
rewrite Rinv_mult_distr; auto with real.
rewrite Rinv_involutive; auto with real.
apply lt_le_weak.
apply nat_of_P_lt_Lt_compare_morphism; auto.
apply ZC2; auto.
intros H' H'0; rewrite (nat_of_P_minus_morphism m1 n1); auto with real.
rewrite (pow_RN_plus e (nat_of_P m1 - nat_of_P n1) (nat_of_P n1));
 auto with real.
rewrite plus_comm; rewrite le_plus_minus_r; auto with real.
apply lt_le_weak.
change (nat_of_P m1 > nat_of_P n1) in |- *.
apply nat_of_P_gt_Gt_compare_morphism; auto.
intros n1 m1. rewrite Z.pos_sub_spec; unfold Pos.compare.
CaseEq (Pcompare n1 m1 Datatypes.Eq); simpl in |- *; auto with real.
intros H' H'0; rewrite Pcompare_Eq_eq with (1 := H'); auto with real.
intros H' H'0; rewrite (nat_of_P_minus_morphism m1 n1); auto with real.
rewrite (pow_RN_plus e (nat_of_P m1 - nat_of_P n1) (nat_of_P n1));
 auto with real.
rewrite plus_comm; rewrite le_plus_minus_r; auto with real.
rewrite Rinv_mult_distr; auto with real.
apply lt_le_weak.
apply nat_of_P_lt_Lt_compare_morphism; auto.
apply ZC2; auto.
intros H' H'0; rewrite (nat_of_P_minus_morphism n1 m1); auto with real.
rewrite (pow_RN_plus e (nat_of_P n1 - nat_of_P m1) (nat_of_P m1));
 auto with real.
rewrite plus_comm; rewrite le_plus_minus_r; auto with real.
apply lt_le_weak.
change (nat_of_P n1 > nat_of_P m1) in |- *.
apply nat_of_P_gt_Gt_compare_morphism; auto.
intros n1 m1; rewrite nat_of_P_plus_morphism; auto with real.
intros H'; rewrite pow_add; auto with real.
apply Rinv_mult_distr; auto.
apply pow_NR0; auto.
apply pow_NR0; auto.
Qed.

Theorem powerRZ_Zopp :
 forall (e : R) (z : Z), e <> 0%R -> powerRZ e (- z) = (/ powerRZ e z)%R.
intros e z H; case z; simpl in |- *; auto with real.
intros p; apply sym_eq; apply Rinv_involutive.
apply pow_nonzero; auto.
Qed.

Theorem powerRZ_Zs :
 forall (e : R) (n : Z),
 e <> 0%R -> powerRZ e (Z.succ n) = (e * powerRZ e n)%R.
intros e n H'0.
replace (Z.succ n) with (n + Z.succ 0)%Z.
rewrite powerRZ_add; auto.
rewrite powerRZ_1.
rewrite Rmult_comm; auto.
auto with zarith.
Qed.

Theorem Zpower_nat_Z_powerRZ :
 forall (n : Z) (m : nat),
 IZR (Zpower_nat n m) = powerRZ (IZR n) (Z_of_nat m).
intros n m; elim m; simpl in |- *; auto with real.
intros m1 H'; rewrite nat_of_P_o_P_of_succ_nat_eq_succ; simpl in |- *.
replace (Zpower_nat n (S m1)) with (n * Zpower_nat n m1)%Z.
rewrite mult_IZR; auto with real.
rewrite H'; simpl in |- *.
case m1; simpl in |- *; auto with real.
intros m2; rewrite nat_of_P_o_P_of_succ_nat_eq_succ; auto.
unfold Zpower_nat in |- *; auto.
Qed.

Theorem powerRZ_lt : forall (e : R) (z : Z), (0 < e)%R -> (0 < powerRZ e z)%R.
intros e z; case z; simpl in |- *; auto with real.
Qed.

Theorem powerRZ_le :
 forall (e : R) (z : Z), (0 < e)%R -> (0 <= powerRZ e z)%R.
intros e z H'; apply Rlt_le; now apply powerRZ_lt.
Qed.

Theorem Rlt_powerRZ :
 forall (e : R) (n m : Z),
 (1 < e)%R -> (n < m)%Z -> (powerRZ e n < powerRZ e m)%R.
intros e n m; case n; case m; simpl in |- *;
 try (unfold Z.lt in |- *; intros; discriminate); auto with real.
intros p p0 H' H'0; apply Rlt_pow; auto with real.
apply nat_of_P_lt_Lt_compare_morphism; auto.
intros p H' H'0; replace 1%R with (/ 1)%R; auto with real.
intros p p0 H' H'0; apply Rlt_trans with (r2 := 1%R).
replace 1%R with (/ 1)%R; auto with real.
apply Rlt_pow_R1; auto with real.
intros p p0 H' H'0; apply Rinv_1_lt_contravar; auto with real.
apply Rlt_pow; auto with real.
apply nat_of_P_lt_Lt_compare_morphism; rewrite ZC4; auto.
Qed.

Theorem Zpower_nat_powerRZ_absolu :
 forall n m : Z,
 (0 <= m)%Z -> IZR (Zpower_nat n (Z.abs_nat m)) = powerRZ (IZR n) m.
intros n m; case m; simpl in |- *; auto with zarith.
intros p H'; elim (nat_of_P p); simpl in |- *; auto with zarith.
intros n0 H'0; rewrite <- H'0; simpl in |- *; auto with zarith.
rewrite <- mult_IZR; auto.
intros p H'; contradict H'; auto with zarith.
Qed.

Theorem powerRZ_R1 : forall n : Z, powerRZ 1 n = 1%R.
intros n; case n; simpl in |- *; auto.
intros p; elim (nat_of_P p); simpl in |- *; auto; intros n0 H'; rewrite H';
 ring.
intros p; elim (nat_of_P p); simpl in |- *.
exact Rinv_1.
intros n1 H'; rewrite Rinv_mult_distr; try rewrite Rinv_1; try rewrite H';
 auto with real.
Qed.

Theorem Rle_powerRZ :
 forall (e : R) (n m : Z),
 (1 <= e)%R -> (n <= m)%Z -> (powerRZ e n <= powerRZ e m)%R.
intros e n m H' H'0.
case H'; intros E1.
case (Zle_lt_or_eq _ _ H'0); intros E2.
apply Rlt_le; auto with real.
now apply Rlt_powerRZ.
rewrite <- E2; auto with real.
repeat rewrite <- E1; repeat rewrite powerRZ_R1; auto with real.
Qed.

Theorem Zlt_powerRZ :
 forall (e : R) (n m : Z),
 (1 <= e)%R -> (powerRZ e n < powerRZ e m)%R -> (n < m)%Z.
intros e n m H' H'0.
case (Zle_or_lt m n); auto; intros Z1.
contradict H'0.
apply Rle_not_lt.
apply Rle_powerRZ; auto.
Qed.

Theorem Zle_powerRZ :
 forall (e : R) (n m : Z),
 (1 < e)%R -> (powerRZ e n <= powerRZ e m)%R -> (n <= m)%Z.
intros e n m H' H'0.
case (Zle_or_lt n m); auto; intros Z1.
absurd (powerRZ e n <= powerRZ e m)%R; auto.
apply Rlt_not_le.
apply Rlt_powerRZ; auto.
Qed.

Theorem Rinv_powerRZ :
 forall (e : R) (n : Z), e <> 0%R -> (/ powerRZ e n)%R = powerRZ e (- n).
intros e n H.
apply Rmult_eq_reg_l with (powerRZ e n); auto with real zarith.
rewrite Rinv_r; auto with real zarith.
rewrite <- powerRZ_add; auto with real zarith.
ring_simplify (n + - n)%Z; simpl in |- *; auto.
now apply powerRZ_NOR.
now apply powerRZ_NOR.
Qed.

Section definitions.
Variable radix : Z.
Hypothesis radixMoreThanOne : (1 < radix)%Z.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ radixMoreThanOne).


Record float : Set := Float {Fnum : Z; Fexp : Z}.

Theorem floatEq :
 forall p q : float, Fnum p = Fnum q -> Fexp p = Fexp q -> p = q.
intros p q; case p; case q; simpl in |- *; intros;
 apply (f_equal2 (A1:=Z) (A2:=Z)); auto.
Qed.

Theorem floatDec : forall x y : float, {x = y} + {x <> y}.
intros x y; case x; case y; intros Fnum2 Fexp2 Fnum1 Fexp1.
case (Z.eq_dec Fnum1 Fnum2); intros H1.
case (Z.eq_dec Fexp1 Fexp2); intros H2.
left; apply floatEq; auto.
right; red in |- *; intros H'; contradict H2; inversion H'; auto.
right; red in |- *; intros H'; contradict H1; inversion H'; auto.
Qed.

Definition Fzero (x : Z) := Float 0 x.

Definition is_Fzero (x : float) := Fnum x = 0%Z.

Coercion IZR : Z >-> R.
Coercion INR : nat >-> R.
Coercion Z_of_nat : nat >-> Z.

Definition FtoR (x : float) := (Fnum x * powerRZ (IZR radix) (Fexp x))%R.

Local Coercion FtoR : float >-> R.

Theorem FzeroisReallyZero : forall z : Z, Fzero z = 0%R :>R.
intros z; unfold FtoR in |- *; simpl in |- *; auto with real.
Qed.

Theorem is_Fzero_rep1 : forall x : float, is_Fzero x -> x = 0%R :>R.
intros x H; unfold FtoR in |- *.
red in H; rewrite H; simpl in |- *; auto with real.
Qed.

Theorem LtFnumZERO : forall x : float, (0 < Fnum x)%Z -> (0 < x)%R.
intros x; case x; unfold FtoR in |- *; simpl in |- *.
intros Fnum1 Fexp1 H'; replace 0%R with (Fnum1 * 0)%R;
 [ apply Rmult_lt_compat_l | ring ]; auto with real zarith.
now apply Rlt_IZR.
now apply powerRZ_lt, Rlt_IZR.
Qed.

Theorem is_Fzero_rep2 : forall x : float, x = 0%R :>R -> is_Fzero x.
intros x H'.
case (Rmult_integral _ _ H'); simpl in |- *; auto.
case x; simpl in |- *.
intros Fnum1 Fexp1 H'0; red in |- *; simpl in |- *; auto with real zarith.
apply eq_IZR_R0; auto.
intros H'0; contradict H'0; apply powerRZ_NOR.
now apply Rgt_not_eq, Rlt_IZR.
Qed.

Theorem NisFzeroComp :
 forall x y : float, ~ is_Fzero x -> x = y :>R -> ~ is_Fzero y.
intros x y H' H'0; contradict H'.
apply is_Fzero_rep2; auto.
rewrite H'0.
apply is_Fzero_rep1; auto.
Qed.

Theorem Rlt_monotony_exp :
 forall (x y : R) (z : Z),
 (x < y)%R -> (x * powerRZ radix z < y * powerRZ radix z)%R.
intros x y z H'; apply Rmult_lt_compat_r; auto with real zarith.
now apply powerRZ_lt, Rlt_IZR.
Qed.

Theorem Rle_monotone_exp :
 forall (x y : R) (z : Z),
 (x <= y)%R -> (x * powerRZ radix z <= y * powerRZ radix z)%R.
intros x y z H'; apply Rmult_le_compat_r; auto with real zarith.
now apply powerRZ_le, Rlt_IZR.
Qed.

Theorem Rlt_monotony_contra_exp :
 forall (x y : R) (z : Z),
 (x * powerRZ radix z < y * powerRZ radix z)%R -> (x < y)%R.
intros x y z H'; apply Rmult_lt_reg_l with (r := powerRZ radix z);
 auto with real zarith.
now apply powerRZ_lt, Rlt_IZR.
repeat rewrite (Rmult_comm (powerRZ radix z)); auto.
Qed.

Theorem Rle_monotony_contra_exp :
 forall (x y : R) (z : Z),
 (x * powerRZ radix z <= y * powerRZ radix z)%R -> (x <= y)%R.
intros x y z H'; apply Rmult_le_reg_l with (r := powerRZ radix z);
 auto with real zarith.
now apply powerRZ_lt, Rlt_IZR.
repeat rewrite (Rmult_comm (powerRZ radix z)); auto.
Qed.

Theorem FtoREqInv2 :
 forall p q : float, p = q :>R -> Fexp p = Fexp q -> p = q.
intros p q H' H'0.
apply floatEq; auto.
apply eq_IZR; auto.
apply Rmult_eq_reg_l with (r := powerRZ radix (Fexp p));
 auto with real zarith.
repeat rewrite (Rmult_comm (powerRZ radix (Fexp p)));
 pattern (Fexp p) at 2 in |- *; rewrite H'0; auto with real zarith.
now apply Rgt_not_eq, powerRZ_lt, Rlt_IZR.
Qed.

Theorem Rlt_Float_Zlt :
 forall p q r : Z, (Float p r < Float q r)%R -> (p < q)%Z.
intros p q r H'.
apply lt_IZR.
apply Rlt_monotony_contra_exp with (z := r); auto with real.
Qed.


Theorem oneExp_le :
 forall x y : Z, (x <= y)%Z -> (Float 1%nat x <= Float 1%nat y)%R.
intros x y H'; unfold FtoR in |- *; simpl in |- *.
repeat rewrite Rmult_1_l; auto with real zarith.
apply Rle_powerRZ; try replace 1%R with (IZR 1); auto with real zarith zarith.
apply Rle_IZR; omega.
Qed.

Theorem oneExp_Zlt :
 forall x y : Z, (Float 1%nat x < Float 1%nat y)%R -> (x < y)%Z.
intros x y H'; case (Zle_or_lt y x); auto; intros ZH; contradict H'.
apply Rle_not_lt; apply oneExp_le; auto.
Qed.

Definition Fdigit (p : float) := digit radix (Fnum p).

Definition Fshift (n : nat) (x : float) :=
  Float (Fnum x * Zpower_nat radix n) (Fexp x - n).

Theorem sameExpEq : forall p q : float, p = q :>R -> Fexp p = Fexp q -> p = q.
intros p q; case p; case q; unfold FtoR in |- *; simpl in |- *.
intros Fnum1 Fexp1 Fnum2 Fexp2 H' H'0; rewrite H'0; rewrite H'0 in H'.
cut (Fnum1 = Fnum2).
intros H'1; rewrite <- H'1; auto.
apply eq_IZR; auto.
apply Rmult_eq_reg_l with (r := powerRZ radix Fexp1);
 repeat rewrite (Rmult_comm (powerRZ radix Fexp1));
 auto.
apply Rlt_dichotomy_converse; right; auto with real.
red in |- *; auto with real.
now apply powerRZ_lt, Rlt_IZR.
Qed.

Theorem FshiftFdigit :
 forall (n : nat) (x : float),
 ~ is_Fzero x -> Fdigit (Fshift n x) = Fdigit x + n.
intros n x; case x; unfold Fshift, Fdigit, is_Fzero in |- *; simpl in |- *.
intros p1 p2 H; apply digitAdd; auto.
Qed.

Theorem FshiftCorrect : forall (n : nat) (x : float), Fshift n x = x :>R.
intros n x; unfold FtoR in |- *; simpl in |- *.
rewrite mult_IZR.
rewrite Zpower_nat_Z_powerRZ; auto.
repeat rewrite Rmult_assoc.
rewrite <- powerRZ_add.
rewrite Zplus_minus; auto.
now apply Rgt_not_eq, Rlt_IZR.
Qed.

Theorem FshiftCorrectInv :
 forall x y : float,
 x = y :>R ->
 (Fexp x <= Fexp y)%Z -> Fshift (Z.abs_nat (Fexp y - Fexp x)) y = x.
intros x y H' H'0; try apply sameExpEq; auto.
apply trans_eq with (y := FtoR y); auto.
apply FshiftCorrect.
generalize H' H'0; case x; case y; simpl in |- *; clear H' H'0 x y.
intros Fnum1 Fexp1 Fnum2 Fexp2 H' H'0; rewrite inj_abs; auto with zarith.
Qed.

Theorem FshiftO : forall x : float, Fshift 0 x = x.
intros x; unfold Fshift in |- *; apply floatEq; simpl in |- *.
replace (Zpower_nat radix 0) with 1%Z; auto with zarith.
simpl in |- *; auto with zarith.
Qed.

Theorem FshiftCorrectSym :
 forall x y : float,
 x = y :>R -> exists n : nat, (exists m : nat, Fshift n x = Fshift m y).
intros x y H'.
case (Z_le_gt_dec (Fexp x) (Fexp y)); intros H'1.
exists 0; exists (Z.abs_nat (Fexp y - Fexp x)).
rewrite FshiftO.
apply sym_equal.
apply FshiftCorrectInv; auto.
exists (Z.abs_nat (Fexp x - Fexp y)); exists 0.
rewrite FshiftO.
apply FshiftCorrectInv; auto with zarith.
Qed.

Theorem FdigitEq :
 forall x y : float,
 ~ is_Fzero x -> x = y :>R -> Fdigit x = Fdigit y -> x = y.
intros x y H' H'0 H'1.
cut (~ is_Fzero y); [ intros NZy | idtac ].
2: red in |- *; intros H'2; case H'.
2: apply is_Fzero_rep2; rewrite H'0; apply is_Fzero_rep1; auto.
case (Zle_or_lt (Fexp x) (Fexp y)); intros Eq1.
case (Zle_lt_or_eq _ _ Eq1); clear Eq1; intros Eq1.
absurd
 (Fdigit (Fshift (Z.abs_nat (Fexp y - Fexp x)) y) =
  Fdigit y + Z.abs_nat (Fexp y - Fexp x)).
rewrite FshiftCorrectInv; auto.
rewrite <- H'1.
red in |- *; intros H'2.
absurd (0%Z = (Fexp y - Fexp x)%Z); auto with zarith arith.
apply Zlt_le_weak; auto.
apply FshiftFdigit; auto.
apply sameExpEq; auto.
absurd
 (Fdigit (Fshift (Z.abs_nat (Fexp x - Fexp y)) x) =
  Fdigit x + Z.abs_nat (Fexp x - Fexp y)).
rewrite FshiftCorrectInv; auto.
rewrite <- H'1.
red in |- *; intros H'2.
absurd (0%Z = (Fexp x - Fexp y)%Z); auto with zarith arith.
apply Zlt_le_weak; auto.
apply FshiftFdigit; auto.
Qed.
End definitions.


Section comparisons.
Variable radix : Z.
Hypothesis radixMoreThanOne : (1 < radix)%Z.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ radixMoreThanOne).

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.


Definition Fle (x y : float) := (x <= y)%R.

Lemma Fle_Zle :
 forall n1 n2 d : Z, (n1 <= n2)%Z -> Fle (Float n1 d) (Float n2 d).
intros; unfold Fle, FtoRradix, FtoR in |- *; simpl in |- *; auto.
case Zle_lt_or_eq with (1 := H); intros H1.
apply Rlt_le, Rmult_lt_compat_r.
now apply powerRZ_lt, Rlt_IZR.
apply Rlt_IZR; easy.
rewrite <- H1; auto with real.
Qed.


Theorem Rlt_Fexp_eq_Zlt :
 forall x y : float, (x < y)%R -> Fexp x = Fexp y -> (Fnum x < Fnum y)%Z.
intros x y H' H'0.
apply lt_IZR.
apply (Rlt_monotony_contra_exp radix) with (z := Fexp x);
 auto with real arith.
pattern (Fexp x) at 2 in |- *; rewrite H'0; auto.
Qed.

Theorem Rle_Fexp_eq_Zle :
 forall x y : float, (x <= y)%R -> Fexp x = Fexp y -> (Fnum x <= Fnum y)%Z.
intros x y H' H'0.
apply le_IZR.
apply (Rle_monotony_contra_exp radix) with (z := Fexp x);
 auto with real arith.
pattern (Fexp x) at 2 in |- *; rewrite H'0; auto.
Qed.

Theorem LtR0Fnum : forall p : float, (0 < p)%R -> (0 < Fnum p)%Z.
intros p H'.
apply lt_IZR.
apply (Rlt_monotony_contra_exp radix) with (z := Fexp p);
 auto with real arith.
simpl in |- *; rewrite Rmult_0_l; auto.
Qed.

Theorem LeR0Fnum : forall p : float, (0 <= p)%R -> (0 <= Fnum p)%Z.
intros p H'.
apply le_IZR.
apply (Rle_monotony_contra_exp radix) with (z := Fexp p);
 auto with real arith.
simpl in |- *; rewrite Rmult_0_l; auto.
Qed.

Theorem LeFnumZERO : forall x : float, (0 <= Fnum x)%Z -> (0 <= x)%R.
intros x H'; unfold FtoRradix, FtoR in |- *.
replace 0%R with (0%Z * 0%Z)%R; auto 6 with real zarith.
apply Rmult_le_compat; try apply Rle_IZR; auto with real zarith.
now apply powerRZ_le, Rlt_IZR.
Qed.

Theorem R0LtFnum : forall p : float, (p < 0)%R -> (Fnum p < 0)%Z.
intros p H'.
apply lt_IZR.
apply (Rlt_monotony_contra_exp radix) with (z := Fexp p);
 auto with real arith.
simpl in |- *; rewrite Rmult_0_l; auto.
Qed.

Theorem R0LeFnum : forall p : float, (p <= 0)%R -> (Fnum p <= 0)%Z.
intros p H'.
apply le_IZR.
apply (Rle_monotony_contra_exp radix) with (z := Fexp p);
 auto with real arith.
simpl in |- *; rewrite Rmult_0_l; auto.
Qed.

Theorem LeZEROFnum : forall x : float, (Fnum x <= 0)%Z -> (x <= 0)%R.
intros x H'; unfold FtoRradix, FtoR in |- *.
apply Ropp_le_cancel; rewrite Ropp_0; rewrite <- Ropp_mult_distr_l_reverse.
replace 0%R with (- 0%Z * 0)%R; try ring.
apply Rmult_le_compat; try auto with real.
rewrite <- 2!Ropp_Ropp_IZR; apply Rle_IZR; omega.
now apply powerRZ_le, Rlt_IZR.
Qed.
End comparisons.


Section operations.
Variable radix : Z.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.

Hypothesis radixNotZero : (0 < radix)%Z.

Definition Fplus (x y : float) :=
  Float
    (Fnum x * Zpower_nat radix (Z.abs_nat (Fexp x - Z.min (Fexp x) (Fexp y))) +
     Fnum y * Zpower_nat radix (Z.abs_nat (Fexp y - Z.min (Fexp x) (Fexp y))))
    (Z.min (Fexp x) (Fexp y)).

Theorem Fplus_correct : forall x y : float, Fplus x y = (x + y)%R :>R.
intros x y; unfold Fplus, Fshift, FtoRradix, FtoR in |- *; simpl in |- *.
rewrite plus_IZR.
rewrite Rmult_comm; rewrite Rmult_plus_distr_l; auto.
repeat rewrite mult_IZR.
repeat rewrite (Rmult_comm (Fnum x)); repeat rewrite (Rmult_comm (Fnum y)).
repeat rewrite Zpower_nat_Z_powerRZ; auto.
repeat rewrite <- Rmult_assoc.
repeat rewrite <- powerRZ_add; auto with real zarith arith.
2: now apply Rgt_not_eq, Rlt_IZR.
2: now apply Rgt_not_eq, Rlt_IZR.
repeat rewrite inj_abs; auto with real zarith.
repeat rewrite Zplus_minus; auto.
Qed.

Definition Fopp (x : float) := Float (- Fnum x) (Fexp x).

Theorem Fopp_correct : forall x : float, Fopp x = (- x)%R :>R.
unfold FtoRradix, FtoR, Fopp in |- *; simpl in |- *.
intros x.
rewrite Ropp_Ropp_IZR; auto with real.
Qed.

Theorem Fopp_Fopp : forall p : float, Fopp (Fopp p) = p.
intros p; case p; unfold Fopp in |- *; simpl in |- *; auto.
intros; rewrite Z.opp_involutive; auto.
Qed.

Theorem Fdigit_opp : forall x : float, Fdigit radix (Fopp x) = Fdigit radix x.
intros x; unfold Fopp, Fdigit in |- *; simpl in |- *.
rewrite <- (digit_abs radix (- Fnum x)).
rewrite <- (digit_abs radix (Fnum x)).
case (Fnum x); simpl in |- *; auto.
Qed.

Definition Fabs (x : float) := Float (Z.abs (Fnum x)) (Fexp x).

Theorem Fabs_correct1 :
 forall x : float, (0 <= FtoR radix x)%R -> Fabs x = x :>R.
intros x; case x; unfold FtoRradix, FtoR in |- *; simpl in |- *.
intros Fnum1 Fexp1 H'.
repeat rewrite <- (Rmult_comm (powerRZ radix Fexp1)); apply Rmult_eq_compat_l;
 auto.
cut (0 <= Fnum1)%Z.
unfold Z.abs, Z.le in |- *.
case Fnum1; simpl in |- *; auto.
intros p H'0; case H'0; auto.
apply Znot_gt_le; auto.
contradict H'.
apply Rgt_not_le; auto.
rewrite Rmult_comm.
replace 0%R with (powerRZ radix Fexp1 * 0)%R; auto with real.
red in |- *; apply Rmult_lt_compat_l; auto with real zarith.
apply powerRZ_lt, Rlt_IZR; easy.
apply Rlt_IZR; omega.
Qed.

Theorem Fabs_correct2 :
 forall x : float, (FtoR radix x <= 0)%R -> Fabs x = (- x)%R :>R.
intros x; case x; unfold FtoRradix, FtoR in |- *; simpl in |- *.
intros Fnum1 Fexp1 H'.
rewrite <- Ropp_mult_distr_l_reverse;
 repeat rewrite <- (Rmult_comm (powerRZ radix Fexp1));
 apply Rmult_eq_compat_l; auto.
cut (Fnum1 <= 0)%Z.
intros; rewrite Z.abs_neq; try assumption.
now rewrite opp_IZR.
apply Znot_gt_le.
contradict H'.
apply Rgt_not_le; auto.
rewrite Rmult_comm.
replace 0%R with (powerRZ radix Fexp1 * 0)%R; auto with real.
red in |- *; apply Rmult_lt_compat_l; auto with real arith.
replace 0%R with (IZR 0); auto with real zarith arith.
now apply powerRZ_lt, Rlt_IZR.
apply Rlt_IZR; omega.
Qed.

Theorem Fabs_correct : forall x : float, Fabs x = Rabs x :>R.
intros x; unfold Rabs in |- *.
case (Rcase_abs x); intros H1.
unfold FtoRradix in |- *; apply Fabs_correct2; auto with arith.
apply Rlt_le; auto.
unfold FtoRradix in |- *; apply Fabs_correct1; auto with arith.
apply Rge_le; auto.
Qed.

Theorem RleFexpFabs :
 forall p : float, p <> 0%R :>R -> (Float 1%nat (Fexp p) <= Fabs p)%R.
intros p H'.
unfold FtoRradix, FtoR, Fabs in |- *; simpl in |- *.
apply Rmult_le_compat_r; auto with real arith.
now apply powerRZ_le, Rlt_IZR.
rewrite Zabs_absolu.
replace 1%R with (INR 1); auto with real.
repeat rewrite <- INR_IZR_INZ; apply Rle_INR; auto.
cut (Z.abs_nat (Fnum p) <> 0); auto with zarith.
contradict H'.
unfold FtoRradix, FtoR in |- *; simpl in |- *.
replace (Fnum p) with 0%Z; try (simpl;ring).
generalize H'; case (Fnum p); simpl in |- *; auto with zarith arith;
 intros p0 H'3; contradict H'3; auto with zarith arith; apply convert_not_O.
Qed.

Theorem Fabs_Fzero : forall x : float, ~ is_Fzero x -> ~ is_Fzero (Fabs x).
intros x; case x; unfold is_Fzero in |- *; simpl in |- *.
intros n m; case n; simpl in |- *; auto with zarith; intros; red in |- *;
 discriminate.
Qed.

Theorem Fdigit_abs : forall x : float, Fdigit radix (Fabs x) = Fdigit radix x.
intros x; unfold Fabs, Fdigit in |- *; simpl in |- *.
case (Fnum x); auto.
Qed.

Definition Fminus (x y : float) := Fplus x (Fopp y).

Theorem Fminus_correct : forall x y : float, Fminus x y = (x - y)%R :>R.
intros x y; unfold Fminus in |- *.
rewrite Fplus_correct.
rewrite Fopp_correct; auto.
Qed.

Theorem Fopp_Fminus : forall p q : float, Fopp (Fminus p q) = Fminus q p.
intros p q; case p; case q; unfold Fopp, Fminus, Fplus in |- *; simpl in |- *;
 auto.
intros n1 e1 n2 e2; apply floatEq; simpl in |- *; repeat rewrite (Zmin_sym e1 e2);
 repeat rewrite Zopp_mult_distr_l_reverse; auto with zarith.
Qed.

Theorem Fopp_Fminus_dist :
 forall p q : float, Fopp (Fminus p q) = Fminus (Fopp p) (Fopp q).
intros p q; case p; case q; unfold Fopp, Fminus, Fplus in |- *; simpl in |- *;
 auto.
intros; apply floatEq; simpl in |- *; repeat rewrite (Zmin_sym Fexp0 Fexp);
 repeat rewrite Zopp_mult_distr_l_reverse; auto with zarith.
Qed.

Definition Fmult (x y : float) := Float (Fnum x * Fnum y) (Fexp x + Fexp y).

Definition Fmult_correct : forall x y : float, Fmult x y = (x * y)%R :>R.
intros x y; unfold FtoRradix, FtoR, Fmult in |- *; simpl in |- *; auto.
rewrite powerRZ_add; auto with real zarith.
repeat rewrite mult_IZR.
repeat rewrite Rmult_assoc.
repeat rewrite (Rmult_comm (Fnum y)).
repeat rewrite <- Rmult_assoc.
repeat rewrite Zmult_assoc_reverse; auto.
now apply Rgt_not_eq, Rlt_IZR.
Qed.

End operations.


Section Fbounded_Def.
Variable radix : Z.
Hypothesis radixMoreThanOne : (1 < radix)%Z.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ radixMoreThanOne).

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.

Coercion Z_of_N: N >-> Z.

Record Fbound : Set := Bound {vNum : positive; dExp : N}.

Definition Fbounded (b : Fbound) (d : float) :=
  (Z.abs (Fnum d) < Zpos (vNum b))%Z /\ (- dExp b <= Fexp d)%Z.

Theorem FzeroisZero : forall b : Fbound, Fzero (- dExp b) = 0%R :>R.
intros b; unfold FtoRradix, FtoR in |- *; simpl in |- *; auto with real.
Qed.

Theorem FboundedFzero : forall b : Fbound, Fbounded b (Fzero (- dExp b)).
intros b; repeat (split; simpl in |- *).
replace 0%Z with (- 0%nat)%Z; [ idtac | simpl in |- *; auto ].
apply Zeq_le; auto with arith.
Qed.

Theorem FboundedZeroSameExp :
 forall (b : Fbound) (p : float), Fbounded b p -> Fbounded b (Fzero (Fexp p)).
intros b p H'; (repeat split; simpl in |- *; auto with zarith).
apply H'.
Qed.

Theorem FBoundedScale :
 forall (b : Fbound) (p : float) (n : nat),
 Fbounded b p -> Fbounded b (Float (Fnum p) (Fexp p + n)).
intros b p n H'; repeat split; simpl in |- *; auto; try apply H'.
apply Z.le_trans with (Fexp p); try apply H'.
omega.
Qed.

Theorem FvalScale :
 forall (b : Fbound) (p : float) (n : nat),
 Float (Fnum p) (Fexp p + n) = (powerRZ radix n * p)%R :>R.
intros b p n; unfold FtoRradix, FtoR in |- *; simpl in |- *.
rewrite powerRZ_add; auto with real zarith.
ring.
now apply Rgt_not_eq, Rlt_IZR.
Qed.

Theorem oppBounded :
 forall (b : Fbound) (x : float), Fbounded b x -> Fbounded b (Fopp x).
intros b x H'; repeat split; simpl in |- *; auto with zarith.
replace (Z.abs (- Fnum x)) with (Z.abs (Fnum x)); try apply H'.
case (Fnum x); simpl in |- *; auto.
apply H'.
Qed.

Theorem oppBoundedInv :
 forall (b : Fbound) (x : float), Fbounded b (Fopp x) -> Fbounded b x.
intros b x H'; rewrite <- (Fopp_Fopp x).
apply oppBounded; auto.
Qed.

Theorem absFBounded :
 forall (b : Fbound) (f : float), Fbounded b f -> Fbounded b (Fabs f).
intros b f H'; repeat split; simpl in |- *; try apply H'.
replace (Z.abs (Z.abs (Fnum f))) with (Z.abs (Fnum f)); try apply H'.
case (Fnum f); auto.
Qed.

Theorem FboundedEqExp :
 forall (b : Fbound) (p q : float),
 Fbounded b p -> p = q :>R -> (Fexp p <= Fexp q)%R -> Fbounded b q.
intros b p q H' H'0 H'1; split.
apply Z.le_lt_trans with (Z.abs (Fnum p)); [ idtac | try apply H' ].
apply
 Z.le_trans with (Z.abs (Fnum (Fshift radix (Z.abs_nat (Fexp q - Fexp p)) q)));
 auto.
unfold Fshift in |- *; simpl in |- *; auto.
rewrite Zabs_Zmult.
pattern (Z.abs (Fnum q)) at 1 in |- *;
 replace (Z.abs (Fnum q)) with (Z.abs (Fnum q) * 1%nat)%Z;
 [ apply Zle_Zmult_comp_l | auto with zarith ]; auto with zarith.
rewrite Z.abs_eq; simpl in |- *; auto with zarith.
apply Zpower_NR1; omega.
apply Zpower_NR0; omega.
cut (Fexp p <= Fexp q)%Z; [ intros E2 | idtac ].
apply le_IZR; auto.
apply (Rle_monotony_contra_exp radix) with (z := Fexp p);
 auto with real arith.
pattern (Fexp p) at 2 in |- *;
 replace (Fexp p) with (Fexp (Fshift radix (Z.abs_nat (Fexp q - Fexp p)) q));
 auto.
rewrite <- (fun x => Rabs_pos_eq (powerRZ radix x)); auto with real zarith.
rewrite <- Rabs_Zabs.
rewrite <- Rabs_mult.
change
  (Rabs (FtoRradix (Fshift radix (Z.abs_nat (Fexp q - Fexp p)) q)) <=
   Z.abs (Fnum p) * powerRZ radix (Fexp p))%R in |- *.
unfold FtoRradix in |- *; rewrite FshiftCorrect; auto.
fold FtoRradix in |- *; rewrite <- H'0.
rewrite <- (Fabs_correct radix); auto with real zarith.
now apply powerRZ_le, Rlt_IZR.
unfold Fshift in |- *; simpl in |- *.
rewrite inj_abs; [ ring | auto with zarith ].
now apply le_IZR.
cut (Fexp p <= Fexp q)%Z; [ intros E2 | apply le_IZR ]; auto.
apply Z.le_trans with (Fexp p); try apply H'; now apply le_IZR.
Qed.

Theorem eqExpLess :
 forall (b : Fbound) (p q : float),
 Fbounded b p ->
 p = q :>R ->
 exists r : float, Fbounded b r /\ r = q :>R /\ (Fexp q <= Fexp r)%R.
intros b p q H' H'0.
case (Rle_or_lt (Fexp q) (Fexp p)); intros H'1.
exists p; repeat (split; auto).
exists q; split; [ idtac | split ]; auto with real.
apply FboundedEqExp with (p := p); auto.
apply Rlt_le; auto.
Qed.

Theorem FboundedShiftLess :
 forall (b : Fbound) (f : float) (n m : nat),
 m <= n -> Fbounded b (Fshift radix n f) -> Fbounded b (Fshift radix m f).
intros b f n m H' H'0; split; auto.
simpl in |- *; auto.
apply Z.le_lt_trans with (Z.abs (Fnum (Fshift radix n f))).
simpl in |- *; replace m with (m + 0); auto with arith.
replace n with (m + (n - m)); auto with arith.
repeat rewrite Zpower_nat_is_exp.
repeat rewrite Zabs_Zmult; auto.
apply Zle_Zmult_comp_l; auto with zarith.
apply Zle_Zmult_comp_l; auto with zarith.
repeat rewrite Z.abs_eq; auto with zarith.
apply Zpower_nat_monotone_le; omega.
apply Zpower_NR0; omega.
apply Zpower_NR0; omega.
apply H'0.
apply Z.le_trans with (Fexp (Fshift radix n f)); try apply H'0.
simpl in |- *; unfold Zminus in |- *; auto with zarith.
Qed.

Theorem eqExpMax :
 forall (b : Fbound) (p q : float),
 Fbounded b p ->
 Fbounded b q ->
 (Fabs p <= q)%R ->
 exists r : float, Fbounded b r /\ r = p :>R /\ (Fexp r <= Fexp q)%Z.
intros b p q H' H'0 H'1; case (Zle_or_lt (Fexp p) (Fexp q)); intros Rl0.
exists p; auto.
cut ((Fexp p - Z.abs_nat (Fexp p - Fexp q))%Z = Fexp q);
 [ intros Eq1 | idtac ].
exists (Fshift radix (Z.abs_nat (Fexp p - Fexp q)) p); split; split; auto.
apply Z.le_lt_trans with (Fnum q); try apply H'0.
replace (Z.abs (Fnum (Fshift radix (Z.abs_nat (Fexp p - Fexp q)) p))) with
 (Fnum (Fabs (Fshift radix (Z.abs_nat (Fexp p - Fexp q)) p)));
 auto.
apply (Rle_Fexp_eq_Zle radix); auto with arith.
rewrite Fabs_correct; auto with arith; rewrite FshiftCorrect; auto with arith;
 rewrite <- (Fabs_correct radix); auto with zarith.
rewrite <- (Z.abs_eq (Fnum q)); try apply H'0; auto with zarith.
apply (LeR0Fnum radix); auto.
apply Rle_trans with (2 := H'1); auto with real.
rewrite (Fabs_correct radix); auto with real zarith.
apply Rabs_pos.
unfold Fshift in |- *; simpl in |- *; rewrite Eq1; auto with zarith.
apply H'0.
unfold FtoRradix in |- *; apply FshiftCorrect; auto.
unfold Fshift in |- *; simpl in |- *.
rewrite Eq1; auto with zarith.
rewrite inj_abs; auto with zarith; ring.
Qed.

Theorem maxFbounded :
 forall (b : Fbound) (z : Z),
 (- dExp b <= z)%Z -> Fbounded b (Float (Z.pred (Zpos (vNum b))) z).
intros b z H; split; auto.
change (Z.abs (Z.pred (Zpos (vNum b))) < Zpos (vNum b))%Z in |- *.
rewrite Z.abs_eq; auto with zarith.
Qed.

Theorem maxMax :
 forall (b : Fbound) (p : float) (z : Z),
 Fbounded b p ->
 (Fexp p <= z)%Z -> (Fabs p < Float (Zpos (vNum b)) z)%R.
intros b p z H' H'0; unfold FtoRradix in |- *;
 rewrite <-
  (FshiftCorrect _ radixMoreThanOne (Z.abs_nat (z - Fexp p))
     (Float (Zpos (vNum b)) z)); unfold Fshift in |- *.
change
  (FtoR radix (Fabs p) <
   FtoR radix
     (Float (Zpos (vNum b) * Zpower_nat radix (Z.abs_nat (z - Fexp p)))
        (z - Z.abs_nat (z - Fexp p))))%R in |- *.
replace (z - Z.abs_nat (z - Fexp p))%Z with (Fexp p).
unfold Fabs, FtoR in |- *.
change
  (Z.abs (Fnum p) * powerRZ radix (Fexp p) <
   (Zpos (vNum b) * Zpower_nat radix (Z.abs_nat (z - Fexp p)))%Z *
   powerRZ radix (Fexp p))%R in |- *.
apply Rmult_lt_compat_r; auto with real zarith.
now apply powerRZ_lt, Rlt_IZR.
apply Rlt_le_trans with (IZR (Zpos (vNum b))).
apply Rlt_IZR, H'.
pattern (Zpos (vNum b)) at 1 in |- *;
 replace (Zpos (vNum b)) with (Zpos (vNum b) * 1)%Z; try ring.
apply Rle_IZR, Zmult_le_compat_l.
apply Zpower_NR1; omega.
auto with zarith.
rewrite inj_abs; auto with zarith; ring.
Qed.
End Fbounded_Def.


Section Fprop.
Variable radix : Z.
Hypothesis radixMoreThanOne : (1 < radix)%Z.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.

Variable b : Fbound.

Theorem SterbenzAux :
 forall x y : float,
 Fbounded b x ->
 Fbounded b y ->
 (y <= x)%R -> (x <= 2%nat * y)%R -> Fbounded b (Fminus radix x y).
intros x y H' H'0 H'1 H'2.
cut (0 <= Fminus radix x y)%R; [ intros Rle1 | idtac ].
cut (Fminus radix x y <= y)%R; [ intros Rle2 | idtac ].
case (Zle_or_lt (Fexp x) (Fexp y)); intros Zle1.
repeat split.
apply Z.le_lt_trans with (Z.abs (Fnum x)); try apply H'.
change (Fnum (Fabs (Fminus radix x y)) <= Fnum (Fabs x))%Z in |- *.
apply Rle_Fexp_eq_Zle with (radix := radix); auto with arith.
repeat rewrite Fabs_correct.
repeat rewrite Rabs_pos_eq; auto.
apply Rle_trans with (2 := H'1); auto.
apply Rle_trans with (2 := H'1); auto.
apply Rle_trans with (2 := Rle2); auto.
apply Z.lt_trans with (2 := radixMoreThanOne); auto with zarith.
apply Z.lt_trans with (2 := radixMoreThanOne); auto with zarith.
unfold Fminus in |- *; simpl in |- *; apply Zmin_le1; auto.
unfold Fminus in |- *; simpl in |- *; rewrite Zmin_le1; try easy.
apply H'.
repeat split.
apply Z.le_lt_trans with (Z.abs (Fnum y)); try apply H'0.
change (Fnum (Fabs (Fminus radix x y)) <= Fnum (Fabs y))%Z in |- *.
apply Rle_Fexp_eq_Zle with (radix := radix); auto with arith.
repeat rewrite Fabs_correct.
repeat rewrite Rabs_pos_eq; auto.
apply Rle_trans with (2 := Rle2); auto.
apply Z.lt_trans with (2 := radixMoreThanOne); auto with zarith.
apply Z.lt_trans with (2 := radixMoreThanOne); auto with zarith.
unfold Fminus in |- *; simpl in |- *; apply Zmin_le2; auto with zarith.
unfold Fminus in |- *; simpl in |- *; rewrite Zmin_le2; try apply H'0;
 auto with zarith.
rewrite (Fminus_correct radix); auto with arith; fold FtoRradix in |- *.
apply Rplus_le_reg_l with (r := FtoRradix y); auto.
replace (y + (x - y))%R with (FtoRradix x); [ idtac | ring ].
replace (y + y)%R with (2%nat * y)%R; [ auto | simpl in |- *; ring ].
apply Z.lt_trans with (2 := radixMoreThanOne); auto with zarith.
rewrite (Fminus_correct radix); auto with arith; fold FtoRradix in |- *.
apply Rplus_le_reg_l with (r := FtoRradix y); auto.
replace (y + (x - y))%R with (FtoRradix x); [ idtac | ring ].
replace (y + 0)%R with (FtoRradix y); [ auto | simpl in |- *; ring ].
apply Z.lt_trans with (2 := radixMoreThanOne); auto with zarith.
Qed.

Theorem Sterbenz :
 forall x y : float,
 Fbounded b x ->
 Fbounded b y ->
 (/ 2%nat * y <= x)%R -> (x <= 2%nat * y)%R -> Fbounded b (Fminus radix x y).
intros x y H' H'0 H'1 H'2.
cut (y <= 2%nat * x)%R; [ intros Le1 | idtac ].
case (Rle_or_lt x y); intros Le2; auto.
apply oppBoundedInv; auto.
rewrite Fopp_Fminus.
apply SterbenzAux; auto with real.
apply SterbenzAux; auto with real.
apply Rmult_le_reg_l with (r := (/ 2%nat)%R).
apply Rinv_0_lt_compat; auto with real.
rewrite <- Rmult_assoc; rewrite Rinv_l; auto with real.
rewrite Rmult_1_l; auto.
Qed.

End Fprop.


Fixpoint mZlist_aux (p : Z) (n : nat) {struct n} :
 list Z :=
  match n with
  | O => p :: nil
  | S n1 => p :: mZlist_aux (Z.succ p) n1
  end.

Theorem mZlist_aux_correct :
 forall (n : nat) (p q : Z),
 (p <= q)%Z -> (q <= p + Z_of_nat n)%Z -> In q (mZlist_aux p n).
intros n; elim n; clear n; auto.
intros p q; try rewrite <- Zplus_0_r_reverse.
intros H' H'0; simpl in |- *; left.
apply Zle_antisym; auto.
intros n H' p q H'0 H'1; case (Zle_lt_or_eq _ _ H'0); intros H'2.
simpl in |- *; right.
apply H'; auto with zarith.
simpl in |- *; auto.
Qed.

Theorem mZlist_aux_correct_rev1 :
 forall (n : nat) (p q : Z), In q (mZlist_aux p n) -> (p <= q)%Z.
intros n; elim n; clear n; simpl in |- *; auto.
intros p q H'; elim H'; auto with zarith.
intros n H' p q H'0; elim H'0; auto with zarith.
intros H'1; apply Zle_succ_le; auto with zarith.
Qed.

Theorem mZlist_aux_correct_rev2 :
 forall (n : nat) (p q : Z),
 In q (mZlist_aux p n) -> (q <= p + Z_of_nat n)%Z.
intros n; elim n; clear n; auto.
intros p q H'; elim H'; auto with zarith.
intros H'0; elim H'0.
intros n H' p q H'0; elim H'0; auto with zarith.
intros H'1; rewrite inj_S; rewrite <- Zplus_succ_comm; auto.
Qed.

Definition mZlist (p q : Z) : list Z :=
  match (q - p)%Z with
  | Z0 => p :: nil
  | Zpos d => mZlist_aux p (nat_of_P d)
  | Zneg _ => nil (A:=Z)
  end.

Theorem mZlist_correct :
 forall p q r : Z, (p <= r)%Z -> (r <= q)%Z -> In r (mZlist p q).
intros p q r H' H'0; unfold mZlist in |- *; CaseEq (q - p)%Z;
 auto with zarith.
intros H'1; rewrite (Zle_antisym r p); auto with datatypes.
auto with zarith.
intros p0 H'1; apply mZlist_aux_correct; auto.
rewrite inject_nat_convert with (1 := H'1); auto with zarith.
intros p0 H'1; absurd (p <= q)%Z; auto.
apply Zlt_not_le; auto.
apply Zlt_O_minus_lt; auto.
replace (p - q)%Z with (- (q - p))%Z; auto with zarith.
apply Z.le_trans with (m := r); auto.
Qed.

Theorem mZlist_correct_rev1 :
 forall p q r : Z, In r (mZlist p q) -> (p <= r)%Z.
intros p q r; unfold mZlist in |- *; CaseEq (q - p)%Z.
intros H' H'0; elim H'0; auto with zarith.
intros H'1; elim H'1.
intros p0 H' H'0.
apply mZlist_aux_correct_rev1 with (n := nat_of_P p0); auto.
intros p0 H' H'0; elim H'0.
Qed.

Theorem mZlist_correct_rev2 :
 forall p q r : Z, In r (mZlist p q) -> (r <= q)%Z.
intros p q r; unfold mZlist in |- *; CaseEq (q - p)%Z.
intros H' H'0; elim H'0; auto with zarith.
intros H'1; elim H'1.
intros p0 H' H'0.
rewrite <- (Zplus_minus p q).
rewrite <- inject_nat_convert with (1 := H').
apply mZlist_aux_correct_rev2; auto.
intros p0 H' H'0; elim H'0.
Qed.

Fixpoint mProd (A B C : Set) (l1 : list A) (l2 : list B) {struct l2} :
 list (A * B) :=
  match l2 with
  | nil => nil
  | b :: l2' => map (fun a : A => (a, b)) l1 ++ mProd A B C l1 l2'
  end.

Theorem mProd_correct :
 forall (A B C : Set) (l1 : list A) (l2 : list B) (a : A) (b : B),
 In a l1 -> In b l2 -> In (a, b) (mProd A B C l1 l2).
intros A B C l1 l2; elim l2; simpl in |- *; auto.
intros a l H' a0 b H'0 H'1; elim H'1;
 [ intros H'2; rewrite <- H'2; clear H'1 | intros H'2; clear H'1 ];
 auto with datatypes.
apply in_or_app; left; auto with datatypes.
generalize H'0; elim l1; simpl in |- *; auto with datatypes.
intros a1 l0 H'1 H'3; elim H'3; clear H'3; intros H'4;
 [ rewrite <- H'4 | idtac ]; auto with datatypes.
Qed.

Theorem mProd_correct_rev1 :
 forall (A B C : Set) (l1 : list A) (l2 : list B) (a : A) (b : B),
 In (a, b) (mProd A B C l1 l2) -> In a l1.
intros A B C l1 l2; elim l2; simpl in |- *; auto.
intros a H' H'0; elim H'0.
intros a l H' a0 b H'0.
case (in_app_or _ _ _ H'0); auto with datatypes.
elim l1; simpl in |- *; auto with datatypes.
intros a1 l0 H'1 H'2; elim H'2; clear H'2; intros H'3;
 [ inversion H'3 | idtac ]; auto with datatypes.
intros H'1; apply H' with (b := b); auto.
Qed.

Theorem mProd_correct_rev2 :
 forall (A B C : Set) (l1 : list A) (l2 : list B) (a : A) (b : B),
 In (a, b) (mProd A B C l1 l2) -> In b l2.
intros A B C l1 l2; elim l2; simpl in |- *; auto.
intros a l H' a0 b H'0.
case (in_app_or _ _ _ H'0); auto with datatypes.
elim l1; simpl in |- *; auto with datatypes.
intros H'1; elim H'1; auto.
intros a1 l0 H'1 H'2; elim H'2; clear H'2; intros H'3;
 [ inversion H'3 | idtac ]; auto with datatypes.
intros H'1; right; apply H' with (a := a0); auto.
Qed.

Theorem in_map_inv :
 forall (A B : Set) (f : A -> B) (l : list A) (x : A),
 (forall a b : A, f a = f b -> a = b) -> In (f x) (map f l) -> In x l.
intros A B f l; elim l; simpl in |- *; auto.
intros a l0 H' x H'0 H'1; elim H'1; clear H'1; intros H'2; auto.
Qed.

Section Fnormalized_Def.
Variable radix : Z.
Hypothesis radixMoreThanOne : (1 < radix)%Z.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ radixMoreThanOne).

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.
Variable b : Fbound.

Definition Fnormal (p : float) :=
  Fbounded b p /\ (Zpos (vNum b) <= Z.abs (radix * Fnum p))%Z.

Theorem FnormalBounded : forall p : float, Fnormal p -> Fbounded b p.
intros p H; case H; auto.
Qed.

Theorem FnormalNotZero : forall p : float, Fnormal p -> ~ is_Fzero p.
unfold is_Fzero in |- *; intros p H; red in |- *; intros H1.
case H; rewrite H1.
replace (Z.abs (radix * 0)) with 0%Z; auto with zarith.
Qed.

Theorem FnormalFop : forall p : float, Fnormal p -> Fnormal (Fopp p).
intros p H; split.
apply oppBounded; apply H.
replace (Z.abs (radix * Fnum (Fopp p))) with (Z.abs (radix * Fnum p));
 try apply H.
case p; simpl in |- *; auto with zarith.
Qed.

Theorem FnormalFabs : forall p : float, Fnormal p -> Fnormal (Fabs p).
intros p; case p; intros a e H; split.
simpl in |- *; case H; intros H1 H2; simpl in |- *; auto.
now apply absFBounded.
simpl; rewrite <- (Z.abs_eq radix); auto with zarith.
rewrite <- Zabs_Zmult.
rewrite (fun x => Z.abs_eq (Z.abs x)); auto with zarith.
apply H.
Qed.

Definition pPred x := Z.pred (Zpos x).

Theorem maxMax1 :
 forall (p : float) (z : Z),
 Fbounded b p -> (Fexp p <= z)%Z -> (Fabs p <= Float (pPred (vNum b)) z)%R.
intros p z H H0; unfold FtoRradix in |- *.
rewrite <-
 (FshiftCorrect _ radixMoreThanOne (Z.abs_nat (z - Fexp p))
    (Float (pPred (vNum b)) z)).
unfold FtoR, Fabs in |- *; simpl in |- *; auto with zarith.
rewrite mult_IZR; rewrite Zpower_nat_Z_powerRZ; auto with zarith.
repeat rewrite inj_abs; auto with zarith.
replace (z - (z - Fexp p))%Z with (Fexp p); [ idtac | ring ].
rewrite Rmult_assoc; rewrite <- powerRZ_add; auto with real zarith.
replace (z - Fexp p + Fexp p)%Z with z; [ idtac | ring ].
apply Rle_trans with (pPred (vNum b) * powerRZ radix (Fexp p))%R.
apply Rle_monotone_exp; auto with zarith; repeat rewrite Rmult_IZR;
 apply Rle_IZR; unfold pPred in |- *; apply Zle_Zpred;
 auto with real zarith.
apply H.
apply Rmult_le_compat_l; auto with real zarith.
apply Rle_IZR; unfold pPred in |- *;
 apply Zle_Zpred; auto with zarith.
apply Rle_powerRZ; auto with real zarith.
apply Rle_IZR; omega.
apply IZR_neq; omega.
Qed.

Definition Fsubnormal (p : float) :=
  Fbounded b p /\
  Fexp p = (- dExp b)%Z /\ (Z.abs (radix * Fnum p) < Zpos (vNum b))%Z.

Theorem FsubnormalFbounded : forall p : float, Fsubnormal p -> Fbounded b p.
intros p H; case H; auto.
Qed.

Theorem FsubnormalFexp :
 forall p : float, Fsubnormal p -> Fexp p = (- dExp b)%Z.
intros p H; case H; auto.
intros H1 H2; case H2; auto.
Qed.

Theorem FsubnormFopp : forall p : float, Fsubnormal p -> Fsubnormal (Fopp p).
intros p H'; repeat split; simpl in |- *; auto with zarith; try apply H'.
rewrite Zabs_Zopp; apply H'.
rewrite <- Zopp_mult_distr_r; rewrite Zabs_Zopp; apply H'.
Qed.

Theorem FsubnormFabs : forall p : float, Fsubnormal p -> Fsubnormal (Fabs p).
intros p; case p; intros a e H; split; try apply H.
apply absFBounded, H.
simpl in |- *; split; try apply H.
case H; intros H1 (H2, H3); auto.
rewrite <- (Z.abs_eq radix); auto with zarith.
rewrite <- Zabs_Zmult.
rewrite (fun x => Z.abs_eq (Z.abs x)); auto with zarith.
Qed.

Theorem FsubnormalUnique :
 forall p q : float, Fsubnormal p -> Fsubnormal q -> p = q :>R -> p = q.
intros p q H' H'0 H'1.
apply FtoREqInv2 with (radix := radix); auto.
generalize H' H'0; unfold Fsubnormal in |- *; auto with zarith.
Qed.

Theorem FsubnormalLt :
 forall p q : float,
 Fsubnormal p -> Fsubnormal q -> (p < q)%R -> (Fnum p < Fnum q)%Z.
intros p q H' H'0 H'1.
apply Rlt_Fexp_eq_Zlt with (radix := radix); auto with zarith.
apply trans_equal with (- dExp b)%Z.
case H'; auto.
intros H1 H2; case H2; auto.
apply sym_equal; case H'0; auto.
intros H1 H2; case H2; auto.
Qed.

Definition Fcanonic (a : float) := Fnormal a \/ Fsubnormal a.

Theorem FcanonicBound : forall p : float, Fcanonic p -> Fbounded b p.
intros p H; case H; intros T; apply T.
Qed.

Theorem FcanonicFopp : forall p : float, Fcanonic p -> Fcanonic (Fopp p).
intros p H'; case H'; intros H'1.
left; apply FnormalFop; auto.
right; apply FsubnormFopp; auto.
Qed.

Theorem FcanonicFabs : forall p : float, Fcanonic p -> Fcanonic (Fabs p).
intros p H'; case H'; clear H'.
intros H; left.
apply FnormalFabs; auto.
intros H; right.
apply FsubnormFabs; auto.
Qed.

Theorem MaxFloat :
 forall x : float,
 Fbounded b x -> (Rabs x < Float (Zpos (vNum b)) (Fexp x))%R.
intros.
replace (Rabs x) with (FtoR radix (Fabs x)).
unfold FtoRradix in |- *.
apply maxMax with (b := b); auto with *.
unfold FtoRradix in |- *.
apply Fabs_correct; auto with *.
Qed.
Variable precision : nat.
Hypothesis precisionNotZero : precision <> 0.
Hypothesis pGivesBound : Zpos (vNum b) = Zpower_nat radix precision.

Theorem FboundNext :
 forall p : float,
 Fbounded b p ->
 exists q : float, Fbounded b q /\ q = Float (Z.succ (Fnum p)) (Fexp p) :>R.
intros p H'.
case (Zle_lt_or_eq (Z.succ (Fnum p)) (Zpos (vNum b))).
case (Zle_or_lt 0 (Fnum p)); intros H1.
rewrite <- (Z.abs_eq (Fnum p)); auto with zarith.
generalize (proj1 H'); auto with zarith.
apply Z.le_trans with 0%Z; auto with zarith.
intros H'0; exists (Float (Z.succ (Fnum p)) (Fexp p)); split; auto.
repeat split; simpl in |- *; try apply H'; auto with zarith.
case (Zle_or_lt 0 (Fnum p)); intros H1; auto with zarith.
apply Z.lt_trans with (Z.abs (Fnum p)); try apply H'; auto with zarith.
repeat rewrite Zabs_eq_opp; auto with zarith.
intros H'0;
 exists (Float (Zpower_nat radix (pred precision)) (Z.succ (Fexp p)));
 split; auto.
repeat split; simpl in |- *; auto with zarith arith.
rewrite pGivesBound.
rewrite Z.abs_eq; auto with zarith.
apply Zpower_nat_monotone_lt; auto with zarith.
apply Zpower_NR0; auto with zarith.
generalize (proj2 H'); auto with zarith.
rewrite H'0; rewrite pGivesBound.
pattern precision at 2 in |- *; replace precision with (1 + pred precision).
rewrite Zpower_nat_is_exp.
rewrite Zpower_nat_1.
unfold FtoRradix, FtoR in |- *; simpl in |- *.
rewrite powerRZ_Zs.
rewrite mult_IZR; ring.
apply IZR_neq; auto with zarith.
generalize precisionNotZero;
  case precision; simpl in |- *; auto with zarith.
Qed.

Theorem digitPredVNumiSPrecision :
 digit radix (Z.pred (Zpos (vNum b))) = precision.
apply digitInv; auto.
rewrite pGivesBound.
rewrite Z.abs_eq; auto with zarith.
apply Zle_Zpred, Zpower_nat_monotone_lt; auto with zarith.
auto with zarith.
Qed.

Theorem digitVNumiSPrecision :
 digit radix (Zpos (vNum b)) = S precision.
apply digitInv; auto.
rewrite pGivesBound.
rewrite Z.abs_eq; auto with zarith.
rewrite Z.abs_eq; auto with zarith.
rewrite pGivesBound; auto with zarith.
apply Zpower_nat_monotone_lt; auto with zarith.
Qed.

Theorem vNumPrecision :
 forall n : Z,
 digit radix n <= precision -> (Z.abs n < Zpos (vNum b))%Z.
intros n H'.
rewrite <- (Z.abs_eq (Zpos (vNum b))); auto with zarith.
apply digit_anti_monotone_lt with (n := radix); auto.
rewrite digitVNumiSPrecision; auto with arith.
Qed.

Theorem pGivesDigit :
 forall p : float, Fbounded b p -> Fdigit radix p <= precision.
intros p H; unfold Fdigit in |- *.
rewrite <- digitPredVNumiSPrecision.
apply digit_monotone; auto with zarith.
rewrite (fun x => Z.abs_eq (Z.pred x)).
generalize (proj1 H); auto with zarith.
apply Zle_Zpred; auto with zarith.
Qed.

Theorem digitGivesBoundedNum :
 forall p : float,
 Fdigit radix p <= precision -> (Z.abs (Fnum p) < Zpos (vNum b))%Z.
intros p H; apply vNumPrecision; auto.
Qed.

Theorem FboundedMboundPos :
 forall z m : Z,
 (0 <= m)%Z ->
 (m <= Zpower_nat radix precision)%Z ->
 (- dExp b <= z)%Z ->
 exists c : float, Fbounded b c /\ c = (m * powerRZ radix z)%R :>R.
intros z m H' H'0 H'1; case (Zle_lt_or_eq _ _ H'0); intros H'2.
exists (Float m z); split; auto with zarith.
repeat split; simpl in |- *; auto with zarith.
case (FboundNext (Float (Z.pred (Zpos (vNum b))) z)).
split; auto with zarith.
rewrite Z.abs_eq.
apply Z.le_lt_trans with (Z.pred (Z.pos (vNum b))); auto with zarith.
apply Zle_Zpred; auto with zarith.
intros f' (H1, H2); exists f'; split; auto.
rewrite H2; rewrite pGivesBound.
unfold FtoRradix, FtoR in |- *; simpl in |- *; auto.
rewrite <- Zsucc_pred; rewrite <- H'2; auto; ring.
Qed.

Theorem FboundedMbound :
 forall z m : Z,
 (Z.abs m <= Zpower_nat radix precision)%Z ->
 (- dExp b <= z)%Z ->
 exists c : float, Fbounded b c /\ c = (m * powerRZ radix z)%R :>R.
intros z m H H0.
case (Zle_or_lt 0 m); intros H1.
case (FboundedMboundPos z (Z.abs m)); auto; try rewrite Z.abs_eq; auto.
intros f (H2, H3); exists f; split; auto.
case (FboundedMboundPos z (Z.abs m)); auto; try rewrite Zabs_eq_opp;
 auto with zarith.
intros f (H2, H3); exists (Fopp f); split.
apply oppBounded; easy.
rewrite (Fopp_correct radix); auto with arith; fold FtoRradix in |- *;
 rewrite H3.
rewrite Ropp_Ropp_IZR; ring.
Qed.

Theorem FnormalPrecision :
 forall p : float, Fnormal p -> Fdigit radix p = precision.
intros p H; apply le_antisym.
apply pGivesDigit; apply H.
apply le_S_n.
rewrite <- digitVNumiSPrecision.
unfold Fdigit in |- *.
replace (S (digit radix (Fnum p))) with (digit radix (Fnum p) + 1).
rewrite <- digitAdd; auto with zarith.
apply digit_monotone; try auto.
rewrite (fun x => Z.abs_eq (Zpos x)); auto with zarith.
rewrite Zmult_comm; rewrite Zpower_nat_1; auto with zarith.
apply H.
destruct H as (H1,H2).
intros H3; contradict H2; rewrite H3.
replace (Z.abs (radix * 0)) with 0%Z; auto with zarith.
rewrite plus_comm; simpl in |- *; auto.
Qed.

Theorem FnormalUnique :
 forall p q : float, Fnormal p -> Fnormal q -> p = q :>R -> p = q.
intros p q H' H'0 H'1.
apply (FdigitEq radix); auto.
apply FnormalNotZero; auto.
apply trans_equal with (y := precision).
now apply FnormalPrecision.
now apply sym_equal, FnormalPrecision.
Qed.

Theorem FnormalLtPos :
 forall p q : float,
 Fnormal p ->
 Fnormal q ->
 (0 <= p)%R ->
 (p < q)%R -> (Fexp p < Fexp q)%Z \/ Fexp p = Fexp q /\ (Fnum p < Fnum q)%Z.
intros p q H' H'0 H'1 H'2.
case (Zle_or_lt (Fexp q) (Fexp p)); auto.
intros H'3; right.
case (Zle_lt_or_eq _ _ H'3); intros H'4.
2: split; auto.
2: apply Rlt_Fexp_eq_Zlt with (radix := radix); auto with zarith.
absurd (Fnum (Fshift radix (Z.abs_nat (Fexp p - Fexp q)) p) < Fnum q)%Z; auto.
2: apply Rlt_Fexp_eq_Zlt with (radix := radix); auto with zarith.
2: unfold FtoRradix in |- *; rewrite FshiftCorrect; auto.
2: unfold Fshift in |- *; simpl in |- *; auto with zarith.
red in |- *; intros H'5.
absurd
 (Fdigit radix (Fshift radix (Z.abs_nat (Fexp p - Fexp q)) p) <=
  Fdigit radix q); auto with arith.
rewrite FshiftFdigit; auto with arith.
replace (Fdigit radix p) with precision.
replace (Fdigit radix q) with precision; auto with zarith.
now apply sym_equal, FnormalPrecision.
now apply sym_equal, FnormalPrecision.
apply FnormalNotZero; auto with arith.
unfold Fdigit in |- *; apply digit_monotone; auto with arith.
repeat rewrite Z.abs_eq; auto with zarith.
apply LeR0Fnum with (radix := radix); auto with zarith.
apply Rle_trans with (r2 := FtoRradix p); auto with real.
apply LeR0Fnum with (radix := radix); auto with zarith.
unfold FtoRradix in |- *; rewrite FshiftCorrect; auto.
Qed.

Definition nNormMin := Zpower_nat radix (pred precision).

Theorem nNormPos : (0 < nNormMin)%Z.
unfold nNormMin in |- *; auto with zarith.
apply Zpower_nat_less; omega.
Qed.

Theorem digitnNormMin : digit radix nNormMin = precision.
unfold nNormMin, Fdigit in |- *; simpl in |- *; apply digitInv; try assumption.
rewrite Z.abs_eq; try apply Zpower_NR0; auto with zarith.
rewrite Z.abs_eq; try apply Zpower_NR0; auto with zarith.
apply Zpower_nat_monotone_lt; auto with zarith.
Qed.

Theorem nNrMMimLevNum : (nNormMin <= Zpos (vNum b))%Z.
rewrite pGivesBound.
unfold nNormMin in |- *; simpl in |- *; auto with zarith arith.
apply Zpower_nat_monotone_le; auto with zarith.
Qed.

Definition firstNormalPos := Float nNormMin (- dExp b).

Theorem firstNormalPosNormal : Fnormal firstNormalPos.
repeat split; unfold firstNormalPos in |- *; simpl in |- *; auto with zarith.
rewrite pGivesBound.
rewrite Z.abs_eq; auto with zarith.
unfold nNormMin in |- *; simpl in |- *.
apply Zpower_nat_monotone_lt; auto with zarith.
apply Zlt_le_weak;apply nNormPos.
rewrite pGivesBound.
replace precision with (pred precision + 1).
rewrite Zpower_nat_is_exp; auto with zarith.
rewrite Zpower_nat_1; auto with zarith.
rewrite (fun x => Zmult_comm x radix); unfold nNormMin in |- *;
 auto with zarith.
lia.
Qed.

Theorem pNormal_absolu_min :
 forall p : float, Fnormal p -> (nNormMin <= Z.abs (Fnum p))%Z.
intros p H; apply Zmult_le_reg_r with (p := radix); auto with zarith.
unfold nNormMin in |- *.
pattern radix at 2 in |- *; rewrite <- (Zpower_nat_1 radix).
rewrite <- Zpower_nat_is_exp; auto with zarith.
replace (pred precision + 1) with precision.
rewrite <- pGivesBound.
rewrite <- (Z.abs_eq radix); auto with zarith.
rewrite <- Zabs_Zmult; rewrite Zmult_comm; apply H.
omega.
Qed.

Theorem maxMaxBis :
 forall (p : float) (z : Z),
 Fbounded b p -> (Fexp p < z)%Z -> (Fabs p < Float nNormMin z)%R.
intros p z H' H'0;
 apply
  Rlt_le_trans with (FtoR radix (Float (Zpos (vNum b)) (Z.pred z))).
unfold FtoRradix in |- *; apply maxMax; auto with zarith;
 unfold Z.pred in |- *; auto with zarith.
unfold FtoRradix, FtoR, nNormMin in |- *; simpl in |- *.
pattern z at 2 in |- *; replace z with (Z.succ (Z.pred z));
 [ rewrite powerRZ_Zs; auto with real zarith
 | unfold Z.succ, Z.pred in |- *; ring ].
rewrite <- Rmult_assoc.
apply Rmult_le_compat_r; auto with real arith.
apply powerRZ_le, Rlt_IZR; omega.
pattern radix at 2 in |- *; rewrite <- (Zpower_nat_1 radix).
rewrite <- mult_IZR.
rewrite <- Zpower_nat_is_exp.
replace (pred precision + 1) with precision.
rewrite pGivesBound.
apply Rle_refl.
auto with zarith.
apply IZR_neq; omega.
Qed.

Theorem FnormalLtFirstNormalPos :
 forall p : float, Fnormal p -> (0 <= p)%R -> (firstNormalPos <= p)%R.
intros p H' H'0.
case (Rle_or_lt firstNormalPos p); intros Lt0; auto with real.
case (FnormalLtPos p firstNormalPos); auto.
apply firstNormalPosNormal.
intros H'1; contradict H'1; unfold firstNormalPos in |- *; simpl in |- *.
apply Zle_not_lt; apply H'.
intros H'1; elim H'1; intros H'2 H'3; contradict H'3.
unfold firstNormalPos in |- *; simpl in |- *.
apply Zle_not_lt.
rewrite <- (Z.abs_eq (Fnum p)); auto with zarith.
apply pNormal_absolu_min; auto.
apply LeR0Fnum with (radix := radix); auto with arith.
Qed.

Theorem FsubnormalDigit :
 forall p : float, Fsubnormal p -> Fdigit radix p < precision.
intros p H; unfold Fdigit in |- *.
case (Z.eq_dec (Fnum p) 0); intros Z1.
rewrite Z1; simpl in |- *; auto with zarith.
apply lt_S_n; apply le_lt_n_Sm.
rewrite <- digitPredVNumiSPrecision.
replace (S (digit radix (Fnum p))) with (digit radix (Fnum p) + 1).
rewrite <- digitAdd; auto with zarith.
apply digit_monotone; try assumption.
rewrite (fun x => Z.abs_eq (Z.pred x)); auto with zarith.
rewrite Zmult_comm; rewrite Zpower_nat_1.
generalize (proj2 (proj2 H)); omega.
rewrite plus_comm; simpl in |- *; auto.
Qed.

Theorem pSubnormal_absolu_min :
 forall p : float, Fsubnormal p -> (Z.abs (Fnum p) < nNormMin)%Z.
intros p H'; apply Zlt_mult_simpl_l with (c := radix); auto with zarith.
replace (radix * Z.abs (Fnum p))%Z with (Z.abs (radix * Fnum p)).
replace (radix * nNormMin)%Z with (Zpos (vNum b)); try apply H'.
rewrite pGivesBound.
replace precision with (1 + pred precision).
rewrite Zpower_nat_is_exp; auto with zarith; rewrite Zpower_nat_1; auto.
generalize precisionNotZero; case precision; simpl in |- *; auto.
intros H; contradict H; auto.
rewrite Zabs_Zmult; rewrite (Z.abs_eq radix); auto with zarith.
Qed.

Theorem FsubnormalLtFirstNormalPos :
 forall p : float, Fsubnormal p -> (0 <= p)%R -> (p < firstNormalPos)%R.
intros p H' H'0; unfold FtoRradix, FtoR, firstNormalPos in |- *;
 simpl in |- *.
replace (Fexp p) with (- dExp b)%Z.
2: apply sym_equal; case H'; intros H1 H2; case H2; auto.
apply Rmult_lt_compat_r; auto with real arith.
apply powerRZ_lt, Rlt_IZR; easy.
apply Rlt_IZR.
rewrite <- (Z.abs_eq (Fnum p)).
2: apply LeR0Fnum with (radix := radix); auto with zarith.
apply pSubnormal_absolu_min; auto.
Qed.

Theorem FsubnormalnormalLtPos :
 forall p q : float,
 Fsubnormal p -> Fnormal q -> (0 <= p)%R -> (0 <= q)%R -> (p < q)%R.
intros p q H' H'0 H'1 H'2.
apply Rlt_le_trans with (r2 := FtoRradix firstNormalPos).
apply FsubnormalLtFirstNormalPos; auto.
apply FnormalLtFirstNormalPos; auto.
Qed.

Theorem FsubnormalnormalLtNeg :
 forall p q : float,
 Fsubnormal p -> Fnormal q -> (p <= 0)%R -> (q <= 0)%R -> (q < p)%R.
intros p q H' H'0 H'1 H'2.
rewrite <- (Ropp_involutive p); rewrite <- (Ropp_involutive q).
apply Ropp_gt_lt_contravar; red in |- *.
unfold FtoRradix in |- *; repeat rewrite <- Fopp_correct.
apply FsubnormalnormalLtPos; auto.
apply FsubnormFopp; auto.
apply FnormalFop; auto.
unfold FtoRradix in |- *; rewrite Fopp_correct; replace 0%R with (-0)%R;
 auto with real.
unfold FtoRradix in |- *; rewrite Fopp_correct; replace 0%R with (-0)%R;
 auto with real.
Qed.

Definition Fnormalize (p : float) :=
  match Z_zerop (Fnum p) with
  | left _ => Float 0 (- dExp b)
  | right _ =>
      Fshift radix
        (min (precision - Fdigit radix p) (Z.abs_nat (dExp b + Fexp p))) p
  end.

Theorem FnormalizeCorrect : forall p : float, Fnormalize p = p :>R.
intros p; unfold Fnormalize in |- *.
case (Z_zerop (Fnum p)).
case p; intros Fnum1 Fexp1 H'; unfold FtoRradix, FtoR in |- *; rewrite H';
 simpl in |- *; auto with real.
apply trans_eq with 0%R; auto with real.
intros H'; unfold FtoRradix in |- *; apply FshiftCorrect; auto.
Qed.

Theorem Fnormalize_Fopp :
 forall p : float, Fnormalize (Fopp p) = Fopp (Fnormalize p).
intros p; case p; unfold Fnormalize in |- *; simpl in |- *.
intros Fnum1 Fexp1; case (Z_zerop Fnum1); intros H'.
rewrite H'; simpl in |- *; auto.
case (Z_zerop (- Fnum1)); intros H'0; simpl in |- *; auto.
case H'; replace Fnum1 with (- - Fnum1)%Z; auto with zarith.
unfold Fopp, Fshift, Fdigit in |- *; simpl in |- *.
replace (digit radix (- Fnum1)) with (digit radix Fnum1).
apply floatEq; simpl in |- *; auto with zarith.
case Fnum1; simpl in |- *; auto.
Qed.

Theorem FnormalizeBounded :
 forall p : float, Fbounded b p -> Fbounded b (Fnormalize p).
intros p H'; red in |- *; split.
unfold Fnormalize in |- *; case (Z_zerop (Fnum p)); auto.
intros H'0; simpl in |- *; auto with zarith.
intros H'0.
apply digitGivesBoundedNum; auto.
rewrite FshiftFdigit; auto.
apply le_trans with (m := Fdigit radix p + (precision - Fdigit radix p));
 auto with arith.
rewrite <- le_plus_minus; auto.
apply pGivesDigit; auto.
unfold Fnormalize in |- *; case (Z_zerop (Fnum p)); auto.
simpl in |- *; auto with zarith.
generalize H'; case p; unfold Fbounded, Fnormal, Fdigit in |- *;
 simpl in |- *.
intros Fnum1 Fexp1 H'0 H'1.
apply Z.le_trans with (m := (Fexp1 - Z.abs_nat (dExp b + Fexp1))%Z).
rewrite inj_abs; auto with zarith.
unfold Zminus in |- *; apply Zplus_le_compat_l; auto.
apply Zle_Zopp; auto.
apply inj_le; auto with arith.
Qed.

Theorem FnormalizeCanonic :
 forall p : float, Fbounded b p -> Fcanonic (Fnormalize p).
intros p H'.
generalize (FnormalizeBounded p H').
unfold Fnormalize in |- *; case (Z_zerop (Fnum p)); auto.
intros H'0; right; repeat split; simpl in |- *; auto with zarith.
intros H'1.
case (min_or (precision - Fdigit radix p) (Z.abs_nat (dExp b + Fexp p)));
 intros Min; case Min; clear Min; intros MinR MinL.
intros H'2; left; split; auto.
rewrite MinR; unfold Fshift in |- *; simpl in |- *.
apply
 Z.le_trans
  with
    (Z.abs
       (radix *
        (Zpower_nat radix (pred (Fdigit radix p)) *
         Zpower_nat radix (precision - Fdigit radix p)))).
pattern radix at 1 in |- *; rewrite <- (Zpower_nat_1 radix).
repeat rewrite <- Zpower_nat_is_exp; auto with zarith.
replace (1 + (pred (Fdigit radix p) + (precision - Fdigit radix p))) with
 precision; auto.
rewrite pGivesBound; auto with real.
rewrite Z.abs_eq; auto with zarith.
cut (Fdigit radix p <= precision).
2: now apply pGivesDigit.
unfold Fdigit in |- *.
generalize (digitNotZero _ radixMoreThanOne _ H'1);
 case (digit radix (Fnum p)); simpl in |- *; auto.
intros tmp; contradict tmp; auto with arith.
intros n H H0; change (precision = S n + (precision - S n)) in |- *.
apply le_plus_minus; auto.
repeat rewrite Zabs_Zmult.
apply Zle_Zmult_comp_l.
apply Zle_ZERO_Zabs.
apply Zle_Zmult_comp_r.
apply Zle_ZERO_Zabs.
rewrite (fun x => Z.abs_eq (Zpower_nat radix x)); auto with zarith.
unfold Fdigit in |- *; apply digitLess; auto.
apply Zpower_NR0; omega.
intros H'0; right; split; auto; split.
rewrite MinR; clear MinR; auto.
cut (- dExp b <= Fexp p)%Z; [ idtac | apply H' ].
case p; simpl in |- *.
intros Fnum1 Fexp1 H'2; rewrite inj_abs; auto with zarith.
rewrite MinR.
rewrite <- (fun x => Z.abs_eq (Zpos x)).
unfold Fshift in |- *; simpl in |- *.
apply
 Z.lt_le_trans
  with
    (Z.abs
       (radix *
        (Zpower_nat radix (Fdigit radix p) *
         Zpower_nat radix (Z.abs_nat (dExp b + Fexp p))))).
repeat rewrite Zabs_Zmult.
apply Zmult_gt_0_lt_compat_l.
apply Z.lt_gt; rewrite Z.abs_eq; auto with zarith.
apply Zmult_gt_0_lt_compat_r.
apply Z.lt_gt; rewrite Z.abs_eq; auto with zarith.
now apply Zpower_nat_less.
apply Zpower_NR0; omega.
rewrite (fun x => Z.abs_eq (Zpower_nat radix x)); auto with zarith.
unfold Fdigit in |- *; apply digitMore; auto.
apply Zpower_NR0; omega.
pattern radix at 1 in |- *; rewrite <- (Zpower_nat_1 radix).
repeat rewrite <- Zpower_nat_is_exp; auto with zarith.
rewrite Z.abs_eq, pGivesBound.
apply Zpower_nat_monotone_le; omega.
apply Zpower_NR0; omega.
auto with zarith.
Qed.

Theorem NormalAndSubNormalNotEq :
 forall p q : float, Fnormal p -> Fsubnormal q -> p <> q :>R.
intros p q H' H'0; red in |- *; intros H'1.
case (Rtotal_order 0 p); intros H'2.
absurd (q < p)%R.
rewrite <- H'1; auto with real.
apply FsubnormalnormalLtPos; auto with real.
rewrite <- H'1; auto with real.
absurd (p < q)%R.
rewrite <- H'1; auto with real.
apply FsubnormalnormalLtNeg; auto with real.
rewrite <- H'1; auto with real.
elim H'2; intros H'3; try rewrite <- H'3; auto with real.
elim H'2; intros H'3; try rewrite <- H'3; auto with real.
Qed.

Theorem FcanonicUnique :
 forall p q : float, Fcanonic p -> Fcanonic q -> p = q :>R -> p = q.
intros p q H' H'0 H'1; case H'; case H'0; intros H'2 H'3.
apply FnormalUnique; auto.
contradict H'1; apply NormalAndSubNormalNotEq; auto.
absurd (q = p :>R); auto; apply NormalAndSubNormalNotEq; auto.
apply FsubnormalUnique; auto.
Qed.

Theorem FcanonicLeastExp :
 forall x y : float,
 x = y :>R -> Fbounded b x -> Fcanonic y -> (Fexp y <= Fexp x)%Z.
intros x y H H0 H1.
cut (Fcanonic (Fnormalize x)); [ intros | apply FnormalizeCanonic; auto ].
replace y with (Fnormalize x);
 [ simpl in |- * | apply FcanonicUnique; auto with real ].
unfold Fnormalize in |- *.
case (Z_zerop (Fnum x)); simpl in |- *; intros Z1; try apply H0.
apply Zplus_le_reg_l with (- Fexp x)%Z.
replace (- Fexp x + Fexp x)%Z with (- (0))%Z; try ring.
replace
 (- Fexp x +
  (Fexp x - min (precision - Fdigit radix x) (Z.abs_nat (dExp b + Fexp x))))%Z
 with (- min (precision - Fdigit radix x) (Z.abs_nat (dExp b + Fexp x)))%Z;
 try ring.
apply Zle_Zopp; auto with arith zarith.
rewrite <- H.
apply FnormalizeCorrect.
Qed.

Theorem FcanonicLtPos :
 forall p q : float,
 Fcanonic p ->
 Fcanonic q ->
 (0 <= p)%R ->
 (p < q)%R -> (Fexp p < Fexp q)%Z \/ Fexp p = Fexp q /\ (Fnum p < Fnum q)%Z.
intros p q H' H'0 H'1 H'2; case H'; case H'0.
intros H'3 H'4; apply FnormalLtPos; auto.
intros H'3 H'4; absurd (p < q)%R; auto.
apply Rlt_asym.
apply FsubnormalnormalLtPos; auto.
apply Rle_trans with (r2 := FtoRradix p); auto with real.
intros H'3 H'4; case (Z.eq_dec (Fexp q) (- dExp b)); intros H'5.
right; split.
rewrite H'5; case H'4; intros H1 H2; case H2; auto.
apply Rlt_Fexp_eq_Zlt with (radix := radix); auto with zarith.
rewrite H'5; case H'4; intros H1 H2; case H2; auto.
left.
replace (Fexp p) with (- dExp b)%Z;
 [ idtac | apply sym_equal, H'4 ].
case (Zle_lt_or_eq (- dExp b) (Fexp q)); try apply H'3; auto with zarith.
intros H'3 H'4; right; split.
apply trans_equal with (- dExp b)%Z; try apply H'4.
apply sym_equal, H'3.
apply FsubnormalLt; auto.
Qed.

Theorem Fcanonic_Rle_Zle :
 forall x y : float,
 Fcanonic x -> Fcanonic y -> (Rabs x <= Rabs y)%R -> (Fexp x <= Fexp y)%Z.
intros x y H H0 H1.
cut (forall z : float, Fexp z = Fexp (Fabs z) :>Z);
 [ intros E | intros; unfold Fabs in |- *; simpl in |- *; auto with zarith ].
rewrite (E x); rewrite (E y).
cut (Fcanonic (Fabs x)); [ intros D | apply FcanonicFabs; auto ].
cut (Fcanonic (Fabs y)); [ intros G | apply FcanonicFabs; auto ].
case H1; intros Z2.
case (FcanonicLtPos (Fabs x) (Fabs y)); auto with zarith.
rewrite (Fabs_correct radix); auto with real zarith.
apply Rabs_pos.
repeat rewrite (Fabs_correct radix); auto with real zarith.
rewrite (FcanonicUnique (Fabs x) (Fabs y)); auto with zarith.
repeat rewrite (Fabs_correct radix); auto with real zarith.
Qed.

Theorem FcanonicLtNeg :
 forall p q : float,
 Fcanonic p ->
 Fcanonic q ->
 (q <= 0)%R ->
 (p < q)%R -> (Fexp q < Fexp p)%Z \/ Fexp p = Fexp q /\ (Fnum p < Fnum q)%Z.
intros p q H' H'0 H'1 H'2.
cut
 ((Fexp (Fopp q) < Fexp (Fopp p))%Z \/
  Fexp (Fopp q) = Fexp (Fopp p) /\ (Fnum (Fopp q) < Fnum (Fopp p))%Z).
simpl in |- *.
intros H'3; elim H'3; clear H'3; intros H'3;
 [ idtac | elim H'3; clear H'3; intros H'3 H'4 ]; auto;
 right; split; auto with zarith.
apply FcanonicLtPos; try apply FcanonicFopp; auto; unfold FtoRradix in |- *;
 repeat rewrite Fopp_correct; replace 0%R with (-0)%R;
 auto with real.
Qed.

Theorem FcanonicFnormalizeEq :
 forall p : float, Fcanonic p -> Fnormalize p = p.
intros p H'.
apply FcanonicUnique; auto.
apply FnormalizeCanonic; auto.
apply FcanonicBound with (1 := H'); auto.
apply FnormalizeCorrect; auto.
Qed.

Theorem FcanonicPosFexpRlt :
 forall x y : float,
 (0 <= x)%R ->
 (0 <= y)%R -> Fcanonic x -> Fcanonic y -> (Fexp x < Fexp y)%Z -> (x < y)%R.
intros x y H' H'0 H'1 H'2 H'3.
case (Rle_or_lt y x); auto.
intros H'4; case H'4; clear H'4; intros H'4.
case FcanonicLtPos with (p := y) (q := x); auto.
intros H'5; contradict H'3; auto with zarith.
intros H'5; elim H'5; intros H'6 H'7; clear H'5; contradict H'3; rewrite H'6;
 auto with zarith.
contradict H'3.
rewrite FcanonicUnique with (p := x) (q := y); auto with zarith.
Qed.

Theorem FcanonicNegFexpRlt :
 forall x y : float,
 (x <= 0)%R ->
 (y <= 0)%R -> Fcanonic x -> Fcanonic y -> (Fexp x < Fexp y)%Z -> (y < x)%R.
intros x y H' H'0 H'1 H'2 H'3.
case (Rle_or_lt x y); auto.
intros H'4; case H'4; clear H'4; intros H'4.
case FcanonicLtNeg with (p := x) (q := y); auto.
intros H'5; contradict H'3; auto with zarith.
intros H'5; elim H'5; intros H'6 H'7; clear H'5; contradict H'3; rewrite H'6;
 auto with zarith.
contradict H'3.
rewrite FcanonicUnique with (p := x) (q := y); auto with zarith.
Qed.

Theorem vNumbMoreThanOne : (1 < Zpos (vNum b))%Z.
replace 1%Z with (Z_of_nat 1); [ idtac | simpl in |- *; auto ].
rewrite <- (Zpower_nat_O radix); rewrite pGivesBound; auto with zarith.
apply Zpower_nat_monotone_lt; omega.
Qed.

Theorem PosNormMin : Zpos (vNum b) = (radix * nNormMin)%Z.
pattern radix at 1 in |- *; rewrite <- (Zpower_nat_1 radix);
 unfold nNormMin in |- *.
rewrite pGivesBound; rewrite <- Zpower_nat_is_exp.
generalize precisionNotZero; case precision; auto with zarith.
Qed.

Theorem FnormalPpred :
 forall x : Z, (- dExp b <= x)%Z -> Fnormal (Float (pPred (vNum b)) x).
intros x H;
 (cut (0 <= pPred (vNum b))%Z;
   [ intros Z1 | unfold pPred in |- *; auto with zarith ]).
repeat split; simpl in |- *; auto with zarith.
rewrite (Z.abs_eq (pPred (vNum b))).
unfold pPred in |- *; auto with zarith.
unfold pPred in |- *; rewrite pGivesBound; auto with zarith.
rewrite Zabs_Zmult; repeat rewrite Z.abs_eq; auto with zarith.
apply Z.le_trans with ((1 + 1) * pPred (vNum b))%Z; auto with zarith.
replace ((1 + 1) * pPred (vNum b))%Z with (pPred (vNum b) + pPred (vNum b))%Z;
 auto with zarith.
replace (Zpos (vNum b)) with (1 + Z.pred (Zpos (vNum b)))%Z;
 unfold pPred in |- *; auto with zarith.
apply Zplus_le_compat_r; apply Zle_Zpred.
apply vNumbMoreThanOne.
Qed.

Theorem FcanonicPpred :
 forall x : Z,
 (- dExp b <= x)%Z -> Fcanonic (Float (pPred (vNum b)) x).
intros x H; left; apply FnormalPpred; auto.
Qed.

Theorem FnormalNnormMin :
 forall x : Z, (- dExp b <= x)%Z -> Fnormal (Float nNormMin x).
intros x H; (cut (0 < nNormMin)%Z; [ intros Z1 | apply nNormPos ]).
repeat split; simpl in |- *; auto with zarith.
rewrite Z.abs_eq; auto with zarith.
rewrite PosNormMin.
pattern nNormMin at 1 in |- *; replace nNormMin with (1 * nNormMin)%Z;
 auto with zarith.
apply Zmult_gt_0_lt_compat_r; auto with zarith.
rewrite PosNormMin; auto with zarith.
Qed.

Theorem FcanonicNnormMin :
 forall x : Z, (- dExp b <= x)%Z -> Fcanonic (Float nNormMin x).
intros x H; left; apply FnormalNnormMin; auto.
Qed.

End Fnormalized_Def.
Section suc.
Variable b : Fbound.
Variable radix : Z.
Variable precision : nat.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.

Hypothesis radixMoreThanOne : (1 < radix)%Z.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ radixMoreThanOne).

Hypothesis precisionNotZero : precision <> 0.
Hypothesis pGivesBound : Zpos (vNum b) = Zpower_nat radix precision.

Definition FSucc (x : float) :=
  match Z_eq_bool (Fnum x) (pPred (vNum b)) with
  | true => Float (nNormMin radix precision) (Z.succ (Fexp x))
  | false =>
      match Z_eq_bool (Fnum x) (- nNormMin radix precision) with
      | true =>
          match Z_eq_bool (Fexp x) (- dExp b) with
          | true => Float (Z.succ (Fnum x)) (Fexp x)
          | false => Float (- pPred (vNum b)) (Z.pred (Fexp x))
          end
      | false => Float (Z.succ (Fnum x)) (Fexp x)
      end
  end.

Theorem FSuccSimpl1 :
 forall x : float,
 Fnum x = pPred (vNum b) ->
 FSucc x = Float (nNormMin radix precision) (Z.succ (Fexp x)).
intros x H'; unfold FSucc in |- *.
generalize (Z_eq_bool_correct (Fnum x) (pPred (vNum b)));
 case (Z_eq_bool (Fnum x) (pPred (vNum b))); auto.
intros H'0; contradict H'0; auto.
Qed.

Theorem FSuccSimpl2 :
 forall x : float,
 Fnum x = (- nNormMin radix precision)%Z ->
 Fexp x <> (- dExp b)%Z ->
 FSucc x = Float (- pPred (vNum b)) (Z.pred (Fexp x)).
intros x H' H'0; unfold FSucc in |- *.
generalize (Z_eq_bool_correct (Fnum x) (pPred (vNum b)));
 case (Z_eq_bool (Fnum x) (pPred (vNum b))); auto.
intros H'1; absurd (0%nat <= pPred (vNum b))%Z; auto with zarith arith.
rewrite <- H'1; rewrite H'.
unfold nNormMin in |- *; simpl in |- *; auto with zarith.
replace 0%Z with (- (0))%Z; auto with zarith.
now apply Zlt_not_le, Zlt_Zopp, Zpower_nat_less.
unfold pPred in |- *; apply Zle_Zpred; auto with zarith.
intros H'1;
 generalize (Z_eq_bool_correct (Fnum x) (- nNormMin radix precision));
 case (Z_eq_bool (Fnum x) (- nNormMin radix precision)).
intros H'2; generalize (Z_eq_bool_correct (Fexp x) (- dExp b));
 case (Z_eq_bool (Fexp x) (- dExp b)); auto.
intros H'3; contradict H'0; auto.
intros H'2; contradict H'2; auto.
Qed.

Theorem FSuccSimpl3 :
 FSucc (Float (- nNormMin radix precision) (- dExp b)) =
 Float (Z.succ (- nNormMin radix precision)) (- dExp b).
unfold FSucc in |- *; simpl in |- *.
generalize (Z_eq_bool_correct (- nNormMin radix precision) (pPred (vNum b)));
 case (Z_eq_bool (- nNormMin radix precision) (pPred (vNum b)));
 auto.
intros H'1; absurd (0%nat <= pPred (vNum b))%Z; auto with zarith arith.
rewrite <- H'1.
unfold nNormMin in |- *; simpl in |- *; auto with zarith.
replace 0%Z with (- (0))%Z; auto with zarith.
now apply Zlt_not_le, Zlt_Zopp, Zpower_nat_less.
unfold pPred in |- *; apply Zle_Zpred; auto with zarith.
intros H';
 generalize
  (Z_eq_bool_correct (- nNormMin radix precision)
     (- nNormMin radix precision));
 case (Z_eq_bool (- nNormMin radix precision) (- nNormMin radix precision)).
intros H'0; generalize (Z_eq_bool_correct (- dExp b) (- dExp b));
 case (Z_eq_bool (- dExp b) (- dExp b)); auto.
intros H'1; contradict H'1; auto.
intros H'1; contradict H'1; auto.
Qed.

Theorem FSuccSimpl4 :
 forall x : float,
 Fnum x <> pPred (vNum b) ->
 Fnum x <> (- nNormMin radix precision)%Z ->
 FSucc x = Float (Z.succ (Fnum x)) (Fexp x).
intros x H' H'0; unfold FSucc in |- *.
generalize (Z_eq_bool_correct (Fnum x) (pPred (vNum b)));
 case (Z_eq_bool (Fnum x) (pPred (vNum b))); auto.
intros H'1; contradict H'; auto.
intros H'1;
 generalize (Z_eq_bool_correct (Fnum x) (- nNormMin radix precision));
 case (Z_eq_bool (Fnum x) (- nNormMin radix precision));
 auto.
intros H'2; contradict H'0; auto.
Qed.

Theorem FSuccDiff1 :
 forall x : float,
 Fnum x <> (- nNormMin radix precision)%Z ->
 Fminus radix (FSucc x) x = Float 1%nat (Fexp x) :>R.
intros x H'.
generalize (Z_eq_bool_correct (Fnum x) (pPred (vNum b)));
 case (Z_eq_bool (Fnum x) (pPred (vNum b))); intros H'1.
rewrite FSuccSimpl1; auto.
unfold FtoRradix, FtoR, Fminus, Fopp, Fplus in |- *; simpl in |- *; auto.
repeat rewrite Zmin_le2; auto with zarith.
rewrite <- Zminus_succ_l; repeat rewrite <- Zminus_diag_reverse.
rewrite absolu_Zs; auto with zarith; simpl in |- *.
rewrite H'1; unfold pPred in |- *; rewrite pGivesBound;
 unfold nNormMin in |- *.
replace (Zpower_nat radix (pred precision) * (radix * 1))%Z with
 (Zpower_nat radix precision). f_equal. unfold Z.pred.
rewrite Z.opp_add_distr. rewrite Z.mul_1_r. rewrite Z.add_assoc. now rewrite Z.add_opp_diag_r.
rewrite Z.mul_1_r.
pattern precision at 1 in |- *; replace precision with (pred precision + 1).
rewrite Zpower_nat_is_exp; rewrite Zpower_nat_1; auto.
generalize precisionNotZero; case precision; simpl in |- *;
 auto with zarith arith.
rewrite FSuccSimpl4; auto.
unfold FtoRradix, FtoR, Fminus, Fopp, Fplus in |- *; simpl in |- *; auto.
repeat rewrite Zmin_n_n; repeat rewrite <- Zminus_diag_reverse; simpl in |- *.
repeat rewrite Zmult_1_r.
replace (Z.succ (Fnum x) + - Fnum x)%Z with (Z_of_nat 1).
simpl in |- *; auto.
simpl in |- *; unfold Z.succ in |- *; ring.
Qed.

Theorem FSuccDiff2 :
 forall x : float,
 Fnum x = (- nNormMin radix precision)%Z ->
 Fexp x = (- dExp b)%Z -> Fminus radix (FSucc x) x = Float 1%nat (Fexp x) :>R.
intros x H' H'0; replace x with (Float (Fnum x) (Fexp x)).
rewrite H'; rewrite H'0; rewrite FSuccSimpl3; auto.
unfold FtoRradix, FtoR, Fminus, Fopp, Fplus in |- *; simpl in |- *; auto.
repeat rewrite Zmin_n_n; repeat rewrite <- Zminus_diag_reverse;
 auto with zarith.
simpl in |- *; repeat rewrite Zmult_1_r.
rewrite Zplus_succ_l; rewrite Zplus_opp_r; simpl in |- *; auto.
case x; simpl in |- *; auto.
Qed.

Theorem FSuccDiff3 :
 forall x : float,
 Fnum x = (- nNormMin radix precision)%Z ->
 Fexp x <> (- dExp b)%Z ->
 Fminus radix (FSucc x) x = Float 1%nat (Z.pred (Fexp x)) :>R.
intros x H' H'1; rewrite FSuccSimpl2; auto.
unfold FtoRradix, FtoR, Fminus, Fopp, Fplus in |- *; simpl in |- *; auto.
repeat rewrite Zmin_le1; auto with zarith.
rewrite <- Zminus_diag_reverse; rewrite <- Zminus_n_predm;
 repeat rewrite <- Zminus_diag_reverse.
rewrite absolu_Zs; auto with zarith; simpl in |- *.
rewrite H'; unfold pPred in |- *; rewrite pGivesBound;
 unfold nNormMin in |- *.
rewrite Z.opp_involutive; repeat rewrite Zmult_1_r.
replace (Zpower_nat radix (pred precision) * radix)%Z with
 (Zpower_nat radix precision).
unfold Z.pred in |- *; simpl in |- *;
 repeat rewrite plus_IZR || rewrite Ropp_Ropp_IZR. f_equal.
simpl in |- *; ring.
pattern precision at 1 in |- *; replace precision with (pred precision + 1).
rewrite Zpower_nat_is_exp; rewrite Zpower_nat_1; auto.
generalize precisionNotZero; case precision; simpl in |- *;
 auto with zarith arith.
Qed.

Theorem ZltNormMinVnum : (nNormMin radix precision < Zpos (vNum b))%Z.
unfold nNormMin in |- *; rewrite pGivesBound; auto with zarith.
apply Zpower_nat_monotone_lt; omega.
Qed.

Theorem FSuccNormPos :
 forall a : float,
 (0 <= a)%R -> Fnormal radix b a -> Fnormal radix b (FSucc a).
intros a H' H'0; unfold FSucc in |- *.
cut (Fbounded b a);
 [ intros B0 | apply FnormalBounded with (1 := H'0); auto ].
generalize (Z_eq_bool_correct (Fnum a) (pPred (vNum b)));
 case (Z_eq_bool (Fnum a) (pPred (vNum b))); auto.
intros H'3; repeat split; simpl in |- *; auto.
rewrite Z.abs_eq; try apply ZltNormMinVnum.
apply Zlt_le_weak, nNormPos; auto with zarith.
apply Z.le_trans with (m := Fexp a); try apply B0; auto with zarith.
rewrite pGivesBound; rewrite Z.abs_eq; auto with zarith.
pattern precision at 1 in |- *; replace precision with (1 + pred precision).
rewrite Zpower_nat_is_exp; rewrite Zpower_nat_1; unfold nNormMin in |- *;
 auto with zarith.
generalize precisionNotZero; case precision; auto with zarith.
apply Zle_mult_gen; simpl in |- *; auto with zarith.
apply Zlt_le_weak, nNormPos; auto with zarith.
intros H'3;
 generalize (Z_eq_bool_correct (Fnum a) (- nNormMin radix precision));
 case (Z_eq_bool (Fnum a) (- nNormMin radix precision)).
intros H'4; absurd (0 <= Fnum a)%Z; auto.
2: apply LeR0Fnum with (radix := radix); auto with zarith.
rewrite H'4; auto.
apply Zlt_not_le.
replace 0%Z with (- 0%nat)%Z by easy.
apply Zlt_Zopp, nNormPos; auto with zarith.
intros H'4; repeat split; simpl in |- *; auto with zarith.
apply Z.le_lt_trans with (Z.succ (Z.abs (Fnum a))); auto with zarith.
case (Zlt_next (Z.abs (Fnum a)) (Zpos (vNum b)));
 auto with zarith arith; try apply B0.
intros H1; contradict H'3.
unfold pPred in |- *; rewrite H1; rewrite Z.abs_eq; auto with zarith.
apply LeR0Fnum with (radix := radix); auto with zarith.
apply B0.
apply Z.le_trans with (Z.abs (radix * Fnum a)); try apply H'0.
case H'0; auto.
repeat rewrite Zabs_Zmult.
cut (0 <= Fnum a)%Z; [ intros Z1 | apply LeR0Fnum with (radix := radix) ];
 auto.
rewrite (Z.abs_eq (Fnum a)); auto.
rewrite (Z.abs_eq (Z.succ (Fnum a))); auto with zarith.
Qed.

Theorem FSuccSubnormNotNearNormMin :
 forall a : float,
 Fsubnormal radix b a ->
 Fnum a <> Z.pred (nNormMin radix precision) -> Fsubnormal radix b (FSucc a).
intros a H' H'0.
cut (Fbounded b a);
 [ intros B0 | apply FsubnormalFbounded with (1 := H'); auto ].
unfold FSucc in |- *.
generalize (Z_eq_bool_correct (Fnum a) (pPred (vNum b)));
 case (Z_eq_bool (Fnum a) (pPred (vNum b))); auto.
intros H'2; absurd (Fdigit radix a < precision).
2: apply FsubnormalDigit with (b := b); auto.
unfold Fdigit in |- *; rewrite H'2.
unfold pPred in |- *;
 rewrite
  (digitPredVNumiSPrecision radix) with (b := b) (precision := precision);
 auto with arith.
intros H'3;
 generalize (Z_eq_bool_correct (Fnum a) (- nNormMin radix precision));
 case (Z_eq_bool (Fnum a) (- nNormMin radix precision)).
intros H'2; absurd (Fdigit radix a < precision).
unfold Fdigit in |- *; rewrite H'2.
replace (digit radix (- nNormMin radix precision)) with
 (digit radix (nNormMin radix precision)).
rewrite digitnNormMin; auto with arith.
case (nNormMin radix precision); simpl in |- *; auto.
apply FsubnormalDigit with (b := b); auto.
intros H'4; repeat split; simpl in |- *; auto with zarith arith; try apply H'.
apply Z.le_lt_trans with (m := Z.succ (Z.abs (Fnum a))).
apply Zabs_Zs.
apply Z.lt_le_trans with (m := Z.succ (nNormMin radix precision)).
apply Zsucc_lt_compat; apply pSubnormal_absolu_min with (3 := pGivesBound);
 auto with zarith arith.
case H'; intros H1 (H2, H3).
apply Zlt_le_succ,ZltNormMinVnum.
rewrite Zabs_Zmult.
rewrite (Z.abs_eq radix); auto with zarith.
apply Z.lt_le_trans with (m := (radix * nNormMin radix precision)%Z).
apply Zmult_gt_0_lt_compat_l; try apply Z.lt_gt; auto with zarith.
apply Zlt_Zabs_Zpred; auto with zarith arith.
apply pSubnormal_absolu_min with (3 := pGivesBound); auto.
pattern radix at 1 in |- *; rewrite <- (Zpower_nat_1 radix);
 unfold nNormMin in |- *; rewrite <- Zpower_nat_is_exp.
rewrite pGivesBound.
generalize precisionNotZero; case precision; simpl in |- *; auto with zarith.
Qed.

Theorem FSuccSubnormNearNormMin :
 forall a : float,
 Fsubnormal radix b a ->
 Fnum a = Z.pred (nNormMin radix precision) -> Fnormal radix b (FSucc a).
intros a H' H'0.
cut (Fbounded b a); [ intros Fb0 | apply FsubnormalFbounded with (1 := H') ].
unfold FSucc in |- *.
generalize (Z_eq_bool_correct (Fnum a) (pPred (vNum b)));
 case (Z_eq_bool (Fnum a) (pPred (vNum b))); auto.
intros H'1; absurd (nNormMin radix precision < Zpos (vNum b))%Z.
2: apply ZltNormMinVnum.
apply Zle_not_lt.
apply Zle_n_Zpred; unfold pPred in H'1; rewrite <- H'1; rewrite H'0;
 auto with zarith.
intros H'3;
 generalize (Z_eq_bool_correct (Fnum a) (- nNormMin radix precision));
 case (Z_eq_bool (Fnum a) (- nNormMin radix precision)).
intros H'1;
 absurd (- nNormMin radix precision < Z.pred (nNormMin radix precision))%Z.
rewrite <- H'1; rewrite <- H'0; auto with zarith.
unfold nNormMin in |- *; apply Z.lt_le_trans with (m := (- (0))%Z);
 auto with zarith.
intros H'4; repeat split; simpl in |- *; auto with zarith arith; try apply Fb0.
rewrite H'0.
rewrite <- Zsucc_pred.
rewrite Z.abs_eq; auto with zarith.
apply ZltNormMinVnum.
apply Zlt_le_weak, nNormPos; easy.
rewrite H'0.
rewrite <- Zsucc_pred.
pattern radix at 1 in |- *; rewrite <- (Zpower_nat_1 radix);
 unfold nNormMin in |- *; rewrite <- Zpower_nat_is_exp.
rewrite pGivesBound.
rewrite Z.abs_eq.
generalize precisionNotZero; case precision; simpl in |- *; auto with zarith.
apply Zpower_NR0; auto with zarith.
Qed.

Theorem FBoundedSuc : forall f : float, Fbounded b f -> Fbounded b (FSucc f).
intros f H'; unfold FSucc in |- *.
generalize (Z_eq_bool_correct (Fnum f) (pPred (vNum b)));
 case (Z_eq_bool (Fnum f) (pPred (vNum b))); intros H'1.
repeat split; simpl in |- *.
rewrite Z.abs_eq.
apply ZltNormMinVnum.
apply Zlt_le_weak, nNormPos; easy.
apply Z.le_trans with (Fexp f); auto with zarith; apply H'.
generalize (Z_eq_bool_correct (Fnum f) (- nNormMin radix precision));
 case (Z_eq_bool (Fnum f) (- nNormMin radix precision));
 intros H'2.
generalize (Z_eq_bool_correct (Fexp f) (- dExp b));
 case (Z_eq_bool (Fexp f) (- dExp b)); intros H'3.
repeat split; simpl in |- *; auto with zarith arith.
apply Zlt_Zabs_Zpred; auto with zarith arith.
apply H'.
repeat split; simpl in |- *; auto with zarith arith.
rewrite Zabs_Zopp.
rewrite Z.abs_eq; unfold pPred in |- *; auto with zarith.
generalize (proj2 H'); omega.
repeat split; simpl in |- *; try apply H'.
apply Zlt_Zabs_Zpred; try apply H'; auto with zarith.
Qed.

Theorem FSuccSubnormal :
 forall a : float, Fsubnormal radix b a -> Fcanonic radix b (FSucc a).
intros a H'.
generalize (Z_eq_bool_correct (Fnum a) (Z.pred (nNormMin radix precision)));
 case (Z_eq_bool (Fnum a) (Z.pred (nNormMin radix precision)));
 intros H'1.
left; apply FSuccSubnormNearNormMin; auto.
right; apply FSuccSubnormNotNearNormMin; auto.
Qed.

Theorem FSuccPosNotMax :
 forall a : float,
 (0 <= a)%R -> Fcanonic radix b a -> Fcanonic radix b (FSucc a).
intros a H' H'0; case H'0; intros H'2.
left; apply FSuccNormPos; auto.
apply FSuccSubnormal; auto.
Qed.

Theorem FSuccNormNegNotNormMin :
 forall a : float,
 (a <= 0)%R ->
 Fnormal radix b a ->
 a <> Float (- nNormMin radix precision) (- dExp b) ->
 Fnormal radix b (FSucc a).
intros a H' H'0 H'1; cut (Fbounded b a);
 [ intros Fb0 | apply FnormalBounded with (1 := H'0) ].
cut (Fnum a <= 0)%Z; [ intros Z0 | apply R0LeFnum with (radix := radix) ];
 auto with zarith.
case (Zle_lt_or_eq _ _ Z0); intros Z1.
2: absurd (is_Fzero a).
2: apply FnormalNotZero with (1 := H'0); auto.
unfold FSucc in |- *.
generalize (Z_eq_bool_correct (Fnum a) (pPred (vNum b)));
 case (Z_eq_bool (Fnum a) (pPred (vNum b))); auto.
intros H'2; absurd (0 < Fnum a)%Z; auto with zarith arith.
rewrite H'2; unfold pPred in |- *; apply Zlt_succ_pred; simpl in |- *;
 apply (vNumbMoreThanOne radix) with (precision := precision);
 auto with zarith arith.
intros H'3;
 generalize (Z_eq_bool_correct (Fnum a) (- nNormMin radix precision));
 case (Z_eq_bool (Fnum a) (- nNormMin radix precision));
 auto.
intros H'2; generalize (Z_eq_bool_correct (Fexp a) (- dExp b));
 case (Z_eq_bool (Fexp a) (- dExp b)).
intros H'4; contradict H'1; auto.
apply floatEq; auto.
intros H'4; repeat split; simpl in |- *; auto with zarith.
rewrite Zabs_Zopp.
unfold pPred in |- *; rewrite Z.abs_eq; auto with zarith.
apply Zle_Zpred; auto with zarith.
case (Zle_next (- dExp b) (Fexp a)); try apply Fb0; auto with zarith.
rewrite <- Zopp_mult_distr_r; rewrite Zabs_Zopp.
rewrite Zabs_Zmult.
repeat rewrite Z.abs_eq; auto with zarith.
pattern (Zpos (vNum b)) at 1 in |- *;
 rewrite (PosNormMin radix) with (precision := precision);
 auto with zarith.
apply Zle_Zmult_comp_l; auto with zarith.
unfold pPred in |- *; apply Zle_Zpred; auto with zarith.
apply ZltNormMinVnum.
apply Zle_Zpred; auto with zarith.
intros H'2; repeat split; simpl in |- *; auto with zarith arith.
apply Z.lt_trans with (Z.abs (Fnum a)); try apply Fb0.
repeat rewrite Zabs_eq_opp; auto with zarith.
apply Fb0.
rewrite Zabs_Zmult.
rewrite (Z.abs_eq radix);
 [ idtac | apply Z.le_trans with 1%Z; auto with zarith ].
repeat rewrite Zabs_eq_opp; auto with zarith.
pattern (Zpos (vNum b)) at 1 in |- *;
 rewrite (PosNormMin radix) with (precision := precision);
 auto with zarith.
apply Zle_Zmult_comp_l; auto with zarith.
replace (- Z.succ (Fnum a))%Z with (Z.pred (- Fnum a)).
case (Zle_lt_or_eq (nNormMin radix precision) (- Fnum a)); auto with zarith.
rewrite <- Zabs_eq_opp; auto with zarith.
apply pNormal_absolu_min with (b := b); auto.
apply Zpred_Zopp_Zs; auto.
apply is_Fzero_rep2 with radix; try easy.
unfold FtoR; rewrite Z1; simpl; ring.
Qed.

Theorem FSuccNormNegNormMin :
 Fsubnormal radix b (FSucc (Float (- nNormMin radix precision) (- dExp b))).
unfold FSucc in |- *; simpl in |- *.
generalize (Z_eq_bool_correct (- nNormMin radix precision) (pPred (vNum b)));
 case (Z_eq_bool (- nNormMin radix precision) (pPred (vNum b)));
 intros H'; auto.
absurd (0%nat < pPred (vNum b))%Z; auto.
rewrite <- H'; auto with zarith.
replace (Z_of_nat 0) with (- (0))%Z; [ idtac | simpl in |- *; auto ].
apply Zle_not_lt; apply Zle_Zopp, Zlt_le_weak.
apply nNormPos; auto with zarith.
unfold pPred in |- *; apply Zlt_succ_pred; simpl in |- *;
 auto with zarith.
apply (vNumbMoreThanOne radix) with (precision := precision);
 auto with zarith.
generalize
 (Z_eq_bool_correct (- nNormMin radix precision) (- nNormMin radix precision));
 case (Z_eq_bool (- nNormMin radix precision) (- nNormMin radix precision));
 intros H'0.
2: contradict H'0; auto.
generalize (Z_eq_bool_correct (- dExp b) (- dExp b));
 case (Z_eq_bool (- dExp b) (- dExp b)); intros H'1.
2: contradict H'1; auto.
repeat split; simpl in |- *; auto with zarith.
apply Z.le_lt_trans with (m := nNormMin radix precision);
 auto with zarith.
rewrite <- Zopp_Zpred_Zs; rewrite Zabs_Zopp; rewrite Z.abs_eq;
 auto with zarith.
apply Zle_Zpred; simpl in |- *; auto with zarith.
apply nNormPos; auto with zarith.
apply ZltNormMinVnum.
rewrite Zabs_Zmult; rewrite (Z.abs_eq radix); auto with zarith.
rewrite (PosNormMin radix) with (precision := precision); auto with zarith.
apply Zmult_gt_0_lt_compat_l; auto with zarith.
rewrite <- Zopp_Zpred_Zs; rewrite Zabs_Zopp.
rewrite Z.abs_eq; auto with zarith.
apply Zle_Zpred; simpl in |- *; auto with zarith.
apply nNormPos; auto with zarith.
Qed.

Theorem FSuccNegCanonic :
 forall a : float,
 (a <= 0)%R -> Fcanonic radix b a -> Fcanonic radix b (FSucc a).
intros a H' H'0; case H'0; intros H'1.
case (floatDec a (Float (- nNormMin radix precision) (- dExp b))); intros H'2.
rewrite H'2; right; apply FSuccNormNegNormMin; auto.
left; apply FSuccNormNegNotNormMin; auto.
apply FSuccSubnormal; auto.
Qed.

Theorem FSuccCanonic :
 forall a : float, Fcanonic radix b a -> Fcanonic radix b (FSucc a).
intros a H'.
case (Rle_or_lt 0 a); intros H'3.
apply FSuccPosNotMax; auto.
apply FSuccNegCanonic; auto with real.
Qed.

Theorem FSuccLt : forall a : float, (a < FSucc a)%R.
intros a; unfold FSucc in |- *.
generalize (Z_eq_bool_correct (Fnum a) (pPred (vNum b)));
 case (Z_eq_bool (Fnum a) (pPred (vNum b))); auto.
intros H'; unfold FtoRradix, FtoR in |- *; simpl in |- *; rewrite H'.
unfold pPred in |- *;
 rewrite (PosNormMin radix) with (precision := precision);
 auto with zarith; unfold nNormMin in |- *.
rewrite powerRZ_Zs; auto with real zarith.
2: apply Rgt_not_eq, Rlt_IZR; omega.
repeat rewrite <- Rmult_assoc.
apply Rlt_monotony_exp; auto with zarith.
rewrite Zmult_comm.
rewrite <- mult_IZR.
apply Rlt_IZR; auto with zarith.
intros H';
 generalize (Z_eq_bool_correct (Fnum a) (- nNormMin radix precision));
 case (Z_eq_bool (Fnum a) (- nNormMin radix precision)).
intros H'0; generalize (Z_eq_bool_correct (Fexp a) (- dExp b));
 case (Z_eq_bool (Fexp a) (- dExp b)).
intros H'1; unfold FtoRradix, FtoR in |- *; simpl in |- *.
apply Rlt_monotony_exp; auto with real zarith.
apply Rlt_IZR; auto with zarith.
intros H'1; unfold FtoRradix, FtoR in |- *; simpl in |- *; rewrite H'0.
pattern (Fexp a) at 1 in |- *; replace (Fexp a) with (Z.succ (Z.pred (Fexp a))).
rewrite powerRZ_Zs.
repeat rewrite <- Rmult_assoc.
apply Rlt_monotony_exp; auto with real zarith.
rewrite <- mult_IZR.
apply Rlt_IZR; auto with zarith.
rewrite <- Zopp_mult_distr_l.
apply Zlt_Zopp.
rewrite Zmult_comm.
unfold pPred in |- *;
 rewrite (PosNormMin radix) with (precision := precision);
 auto with zarith.
apply Rgt_not_eq, Rlt_IZR; omega.
apply sym_equal; apply Zsucc_pred.
intros H'1; unfold FtoRradix, FtoR in |- *; simpl in |- *;
 auto with real zarith.
apply Rmult_lt_compat_r.
apply powerRZ_lt, Rlt_IZR; auto with zarith.
apply Rlt_IZR; auto with zarith.
Qed.

Theorem FSuccPropPos :
 forall x y : float,
 (0 <= x)%R ->
 Fcanonic radix b x -> Fcanonic radix b y -> (x < y)%R -> (FSucc x <= y)%R.
intros x y H' H'0 H'1 H'2.
cut (Fbounded b x); [ intros Fb0 | apply FcanonicBound with (1 := H'0) ].
cut (Fbounded b y); [ intros Fb1 | apply FcanonicBound with (1 := H'1) ].
case FcanonicLtPos with (p := x) (q := y) (3 := pGivesBound); auto.
case (Z.eq_dec (Fnum x) (pPred (vNum b))); intros H'4.
rewrite FSuccSimpl1; auto.
intros H'5; case (Zlt_next _ _ H'5); intros H'6.
replace y with (Float (Fnum y) (Fexp y)).
rewrite H'6.
generalize Fle_Zle; unfold Fle, FtoRradix in |- *; intros H'7; apply H'7;
 clear H'7; auto with arith.
rewrite <- (Z.abs_eq (Fnum y)); auto with zarith.
apply pNormal_absolu_min with (b := b); auto.
case H'1; auto with zarith.
intros H'7; contradict H'5; apply Zle_not_lt.
replace (Fexp y) with (- dExp b)%Z; try apply Fb0.
case H'7; intros H'8 (H'9, H'10); auto.
apply LeR0Fnum with (radix := radix); auto with zarith.
apply Rle_trans with (r2 := FtoR radix x); auto with real.
case y; auto.
apply Rlt_le; auto.
unfold FtoRradix in |- *; apply FcanonicPosFexpRlt with (3 := pGivesBound);
 auto.
apply LeFnumZERO with (radix := radix); auto with zarith.
simpl in |- *; auto with zarith.
apply Zlt_le_weak; apply nNormPos; try easy.
apply Rle_trans with (r2 := FtoR radix x); auto with real.
rewrite <- FSuccSimpl1; auto.
apply FSuccCanonic; auto.
intros H'5; apply Rlt_le.
unfold FtoRradix in |- *; apply FcanonicPosFexpRlt with (3 := pGivesBound);
 auto.
apply Rle_trans with (r2 := FtoR radix x); auto.
apply Rlt_le; auto.
apply FSuccLt; auto.
apply Rle_trans with (r2 := FtoR radix x); auto with real.
apply FSuccCanonic; auto.
rewrite FSuccSimpl4; auto.
apply sym_not_equal; apply Zlt_not_eq.
apply Z.lt_le_trans with (m := 0%Z); auto with zarith.
replace 0%Z with (- 0%nat)%Z; auto with zarith.
apply Zlt_Zopp.
apply nNormPos; auto.
apply LeR0Fnum with (radix := radix); auto with zarith.
intros H'4; elim H'4; intros H'5 H'6; clear H'4.
generalize (Z_eq_bool_correct (Fnum x) (Zpos (vNum b)));
 case (Z_eq_bool (Fnum x) (Zpos (vNum b)));
 intros H'4.
contradict H'6; auto.
apply Zle_not_lt; apply Zlt_le_weak.
rewrite H'4; auto with zarith.
rewrite <- (Z.abs_eq (Fnum y)); auto with zarith; try apply Fb1.
apply LeR0Fnum with (radix := radix); auto with zarith.
apply Rle_trans with (FtoRradix x); auto with real.
case (Zlt_next _ _ H'6); intros H'7.
rewrite FSuccSimpl4; auto.
rewrite <- H'7; rewrite H'5; unfold FtoRradix, FtoR in |- *; simpl in |- *;
 auto with real.
apply Zlt_not_eq.
unfold pPred in |- *; apply Zlt_succ_pred; rewrite <- H'7; auto with zarith.
rewrite <- (Z.abs_eq (Fnum y)); auto with zarith; try apply Fb1.
apply LeR0Fnum with (radix := radix); auto with zarith.
apply Rle_trans with (FtoRradix x); auto with real.
apply Zlt_not_eq_rev.
apply Z.lt_le_trans with (m := 0%Z); auto with zarith.
replace 0%Z with (- 0%nat)%Z; auto with zarith.
apply Zlt_Zopp.
apply nNormPos; auto.
apply LeR0Fnum with (radix := radix); auto with zarith.
rewrite FSuccSimpl4; auto.
replace y with (Float (Fnum y) (Fexp y)).
rewrite H'5.
unfold FtoRradix, FtoR in |- *; simpl in |- *; auto with real.
apply Rmult_le_compat_r.
apply powerRZ_le, Rlt_IZR; auto with zarith.
apply Rle_IZR; auto with zarith.
case y; simpl in |- *; auto.
contradict H'7; auto.
apply Zle_not_lt; apply Zlt_le_weak.
rewrite H'7.
unfold pPred in |- *; rewrite <- Zsucc_pred.
rewrite <- (Z.abs_eq (Fnum y)); try apply Fb1.
apply LeR0Fnum with (radix := radix); auto with zarith.
apply Rle_trans with (FtoRradix x); auto with real.
apply Zlt_not_eq_rev.
apply Z.lt_le_trans with (m := 0%Z); auto with zarith.
replace 0%Z with (- 0%nat)%Z; auto with zarith.
apply Zlt_Zopp.
apply nNormPos; auto.
apply LeR0Fnum with (radix := radix); auto with zarith.
Qed.

Theorem R0RltRleSucc : forall x : float, (x < 0)%R -> (FSucc x <= 0)%R.
intros a H'; unfold FSucc in |- *.
generalize (Z_eq_bool_correct (Fnum a) (pPred (vNum b)));
 case (Z_eq_bool (Fnum a) (pPred (vNum b))); auto.
intros H'0; absurd (Fnum a < 0)%Z; auto.
rewrite H'0; auto with zarith arith.
apply Zle_not_lt; unfold pPred in |- *; apply Zle_Zpred; auto with zarith.
apply R0LtFnum with (radix := radix); auto with zarith.
generalize (Z_eq_bool_correct (Fnum a) (- nNormMin radix precision));
 case (Z_eq_bool (Fnum a) (- nNormMin radix precision));
 intros H'1.
generalize (Z_eq_bool_correct (Fexp a) (- dExp b));
 case (Z_eq_bool (Fexp a) (- dExp b)); intros H'2.
intros H'0.
apply LeZEROFnum with (radix := radix); simpl in |- *; auto with zarith.
apply Zlt_le_succ.
apply R0LtFnum with (radix := radix); auto with zarith.
intros H'0.
apply LeZEROFnum with (radix := radix); simpl in |- *; auto with zarith.
replace 0%Z with (- (0))%Z; [ apply Zle_Zopp | simpl in |- *; auto ].
unfold pPred in |- *; apply Zle_Zpred; apply Z.lt_trans with 1%Z;
 auto with zarith;
 apply (vNumbMoreThanOne radix) with (precision := precision);
 auto with zarith.
intros H'0.
apply LeZEROFnum with (radix := radix); simpl in |- *; auto with zarith.
apply Zlt_le_succ.
apply R0LtFnum with (radix := radix); auto with zarith.
Qed.

Theorem FSuccPropNeg :
 forall x y : float,
 (x < 0)%R ->
 Fcanonic radix b x -> Fcanonic radix b y -> (x < y)%R -> (FSucc x <= y)%R.
intros x y H' H'0 H'1 H'2.
cut (Fbounded b x); [ intros Fb0 | apply FcanonicBound with (1 := H'0) ].
cut (Fbounded b y); [ intros Fb1 | apply FcanonicBound with (1 := H'1) ].
case (Rle_or_lt 0 y); intros Rle0.
apply Rle_trans with (r2 := 0%R); auto.
apply R0RltRleSucc; auto.
cut (Fnum x <> pPred (vNum b)); [ intros N0 | idtac ]; auto with zarith.
generalize (Z_eq_bool_correct (Fnum x) (- nNormMin radix precision));
 case (Z_eq_bool (Fnum x) (- nNormMin radix precision));
 intros H'4.
generalize (Z_eq_bool_correct (Fexp x) (- dExp b));
 case (Z_eq_bool (Fexp x) (- dExp b)); intros H'5.
replace x with (Float (Fnum x) (Fexp x)).
rewrite H'4; rewrite H'5; rewrite FSuccSimpl3; auto.
case FcanonicLtNeg with (p := x) (q := y) (3 := pGivesBound); auto with real.
intros H'6; contradict H'6; rewrite H'5; apply Zle_not_lt; apply Fb1.
intros H'6; elim H'6; intros H'7 H'8; clear H'6;
 replace y with (Float (Fnum y) (Fexp y)).
rewrite <- H'7; rewrite H'5.
generalize Fle_Zle; unfold Fle, FtoRradix in |- *; intros H'9; apply H'9;
 clear H'9; auto with arith.
rewrite <- H'4; auto with zarith.
case y; auto.
case x; auto.
rewrite FSuccSimpl2; auto.
case FcanonicLtNeg with (p := x) (q := y) (3 := pGivesBound); auto with real.
intros H'6; replace y with (Float (Fnum y) (Fexp y)).
case (Zlt_next _ _ H'6); intros H'7.
rewrite H'7.
rewrite <- Zpred_succ.
unfold FtoRradix, FtoR in |- *; simpl in |- *.
apply Rle_monotone_exp; auto with zarith.
rewrite <- (Z.opp_involutive (Fnum y)); apply Rle_IZR; apply Zle_Zopp.
unfold pPred in |- *; apply Zle_Zpred; rewrite <- Zabs_eq_opp;
 auto with zarith; try apply Fb1.
apply Zlt_le_weak; apply R0LtFnum with (radix := radix); auto with zarith.
apply Rlt_le; auto with real.
unfold FtoRradix in |- *; apply FcanonicNegFexpRlt with (3 := pGivesBound);
 auto.
apply Rlt_le; auto.
rewrite <- FSuccSimpl2; auto.
apply R0RltRleSucc; auto.
rewrite <- FSuccSimpl2; auto.
apply FSuccCanonic; auto.
simpl in |- *; auto.
apply Zsucc_lt_reg; auto.
rewrite <- Zsucc_pred; auto with zarith.
case y; auto.
intros H'6; elim H'6; intros H'7 H'8; clear H'6; apply Rlt_le.
contradict H'8; rewrite H'4.
apply Zle_not_lt.
replace (Fnum y) with (- Z.abs (Fnum y))%Z.
apply Zle_Zopp.
apply pNormal_absolu_min with (3 := pGivesBound); auto.
case H'1; auto.
intros H'6; contradict H'5; rewrite H'7; apply H'6.
rewrite Zabs_eq_opp.
ring.
apply R0LeFnum with (radix := radix); auto with zarith.
apply Rlt_le; auto.
rewrite FSuccSimpl4; auto.
case FcanonicLtNeg with (p := x) (q := y) (3 := pGivesBound); auto.
apply Rlt_le; auto with real.
intros H'5; apply Rlt_le; auto.
unfold FtoRradix in |- *; apply FcanonicNegFexpRlt with (3 := pGivesBound);
 auto.
apply Rlt_le; auto.
rewrite <- FSuccSimpl4; auto.
apply R0RltRleSucc; auto.
rewrite <- FSuccSimpl4; auto.
apply FSuccCanonic; auto.
intros H'5; elim H'5; intros H'6 H'7; clear H'5.
replace y with (Float (Fnum y) (Fexp y)).
rewrite H'6.
generalize Fle_Zle; unfold Fle, FtoRradix in |- *; intros H'8; apply H'8;
 clear H'8; auto with zarith arith.
case y; auto.
apply Zlt_not_eq.
apply Z.lt_trans with 0%Z; auto with zarith.
apply R0LtFnum with (radix := radix); auto with zarith.
unfold pPred in |- *; apply Zlt_succ_pred.
replace (Z.succ 0) with (Z_of_nat 1);
 [ apply (vNumbMoreThanOne radix) with (precision := precision)
 | simpl in |- * ]; auto with zarith.
Qed.

Theorem FSuccProp :
 forall x y : float,
 Fcanonic radix b x -> Fcanonic radix b y -> (x < y)%R -> (FSucc x <= y)%R.
intros x y H' H'0 H'1; case (Rle_or_lt 0 x); intros H'2.
apply FSuccPropPos; auto.
apply FSuccPropNeg; auto.
Qed.

Theorem FSuccZleEq :
 forall p q : float,
 (p <= q)%R -> (q < FSucc p)%R -> (Fexp p <= Fexp q)%Z -> p = q :>R.
intros p q H'.
generalize (Z_eq_bool_correct (Fnum p) (pPred (vNum b)));
 case (Z_eq_bool (Fnum p) (pPred (vNum b))); intros H'0.
rewrite FSuccSimpl1; simpl in |- *; auto with arith.
intros H'1 H'2.
replace p with (Fshift radix (Z.abs_nat (Fexp q - Fexp p)) q).
unfold FtoRradix in |- *; rewrite FshiftCorrect; auto with real.
cut (Fexp (Fshift radix (Z.abs_nat (Fexp q - Fexp p)) q) = Fexp p);
 [ intros Eq0 | idtac ].
apply floatEq; auto.
apply sym_equal; apply Zeq_Zs; auto.
apply Rle_Fexp_eq_Zle with (radix := radix); auto with arith.
rewrite FshiftCorrect; auto.
replace (Z.succ (Fnum p)) with (Fnum (Fshift radix 1 (FSucc p))); auto.
apply Rlt_Fexp_eq_Zlt with (radix := radix); auto with arith.
repeat rewrite FshiftCorrect; auto.
rewrite FSuccSimpl1; simpl in |- *; auto with arith.
unfold Fshift in |- *; simpl in |- *.
rewrite FSuccSimpl1; simpl in |- *; auto with arith.
rewrite inj_abs; auto with zarith.
unfold Fshift in |- *; simpl in |- *.
rewrite FSuccSimpl1; simpl in |- *; auto with arith.
rewrite Z.mul_1_r.
rewrite H'0.
unfold pPred in |- *; rewrite <- Zsucc_pred.
rewrite (PosNormMin radix) with (precision := precision); auto with zarith;
 apply Zmult_comm.
unfold Fshift in |- *; simpl in |- *.
rewrite inj_abs; auto with zarith.
generalize (Z_eq_bool_correct (Fnum p) (- nNormMin radix precision));
 case (Z_eq_bool (Fnum p) (- nNormMin radix precision));
 intros H'1.
generalize (Z_eq_bool_correct (Fexp p) (- dExp b));
 case (Z_eq_bool (Fexp p) (- dExp b)); intros H'2.
pattern p at 1 in |- *; replace p with (Float (Fnum p) (Fexp p)).
rewrite H'1; rewrite H'2.
rewrite FSuccSimpl3; auto with arith.
intros H'3 H'4.
replace p with (Fshift radix (Z.abs_nat (Fexp q - Fexp p)) q).
unfold FtoRradix in |- *; rewrite FshiftCorrect; auto with real.
cut (Fexp (Fshift radix (Z.abs_nat (Fexp q - Fexp p)) q) = Fexp p);
 [ intros Eq0 | idtac ].
apply floatEq; auto.
apply sym_equal; apply Zeq_Zs; auto.
apply Rle_Fexp_eq_Zle with (radix := radix); auto with arith.
rewrite FshiftCorrect; auto.
replace (Z.succ (Fnum p)) with (Fnum (FSucc p)); auto.
pattern p at 2 in |- *; replace p with (Float (Fnum p) (Fexp p)).
rewrite H'1; rewrite H'2.
rewrite FSuccSimpl3; auto with arith.
rewrite <- H'2.
apply Rlt_Fexp_eq_Zlt with (radix := radix); auto with arith.
rewrite FshiftCorrect; auto.
rewrite H'2; auto.
case p; simpl in |- *; auto.
pattern p at 1 in |- *; replace p with (Float (Fnum p) (Fexp p)).
rewrite H'1; rewrite H'2.
rewrite FSuccSimpl3; auto with arith.
case p; simpl in |- *; auto.
unfold Fshift in |- *; simpl in |- *.
rewrite inj_abs; auto with zarith.
case p; simpl in |- *; auto.
rewrite FSuccSimpl2; auto with arith.
intros H'3 H'4.
unfold FtoRradix in |- *; rewrite <- FshiftCorrect with (n := 1) (x := p);
 auto.
replace (Fshift radix 1 p) with
 (Fshift radix (S (Z.abs_nat (Fexp q - Fexp p))) q).
repeat rewrite FshiftCorrect; auto with real.
cut
 (Fexp (Fshift radix (S (Z.abs_nat (Fexp q - Fexp p))) q) =
  Fexp (Fshift radix 1 p)); [ intros Eq0 | idtac ].
apply floatEq; auto.
apply sym_equal; apply Zeq_Zs; auto.
apply Rle_Fexp_eq_Zle with (radix := radix); auto with arith.
repeat rewrite FshiftCorrect; auto.
replace (Z.succ (Fnum (Fshift radix 1 p))) with (Fnum (FSucc p)); auto.
apply Rlt_Fexp_eq_Zlt with (radix := radix); auto with arith.
repeat rewrite FshiftCorrect; auto.
rewrite FSuccSimpl2; auto with arith.
rewrite FSuccSimpl2; auto with arith.
rewrite FSuccSimpl2; auto with arith.
unfold Fshift in |- *; simpl in |- *.
rewrite Z.mul_1_r; auto.
unfold pPred in |- *;
 rewrite (PosNormMin radix) with (precision := precision);
 auto with zarith; rewrite H'1.
rewrite Zopp_mult_distr_l_reverse.
rewrite (Zmult_comm radix).
apply Zopp_Zpred_Zs.
unfold Fshift in |- *; simpl in |- *.
replace (Zpos (P_of_succ_nat (Z.abs_nat (Fexp q - Fexp p))))
 with (Z.succ (Fexp q - Fexp p)).
unfold Z.succ, Z.pred in |- *; ring.
rewrite <- (inj_abs (Fexp q - Fexp p)); auto with zarith.
rewrite FSuccSimpl4; auto.
intros H'2 H'3.
replace p with (Fshift radix (Z.abs_nat (Fexp q - Fexp p)) q).
unfold FtoRradix in |- *; rewrite FshiftCorrect; auto with real.
cut (Fexp (Fshift radix (Z.abs_nat (Fexp q - Fexp p)) q) = Fexp p);
 [ intros Eq0 | idtac ].
apply floatEq; auto.
apply sym_equal; apply Zeq_Zs; auto.
apply Rle_Fexp_eq_Zle with (radix := radix); auto with arith.
rewrite FshiftCorrect; auto.
replace (Z.succ (Fnum p)) with (Fnum (FSucc p)); auto.
rewrite FSuccSimpl4; auto.
apply Rlt_Fexp_eq_Zlt with (radix := radix); auto with arith.
repeat rewrite FshiftCorrect; auto.
rewrite FSuccSimpl4; auto.
unfold Fshift in |- *; simpl in |- *.
rewrite inj_abs; auto with zarith.
Qed.

Definition FNSucc x := FSucc (Fnormalize radix b precision x).

Theorem FNSuccCanonic :
 forall a : float, Fbounded b a -> Fcanonic radix b (FNSucc a).
intros a H'; unfold FNSucc in |- *.
apply FSuccCanonic; auto with zarith.
apply FnormalizeCanonic; easy.
Qed.

Theorem FNSuccLt : forall a : float, (a < FNSucc a)%R.
intros a; unfold FNSucc in |- *.
unfold FtoRradix in |- *;
 rewrite <- (FnormalizeCorrect _ radixMoreThanOne b precision a).
apply FSuccLt; auto.
Qed.

Theorem FNSuccProp :
 forall x y : float,
 Fbounded b x -> Fbounded b y -> (x < y)%R -> (FNSucc x <= y)%R.
intros x y H' H'0 H'1; unfold FNSucc in |- *.
replace (FtoRradix y) with (FtoRradix (Fnormalize radix b precision y)).
2: unfold FtoRradix in |- *; apply FnormalizeCorrect; auto.
apply FSuccProp; auto with zarith.
apply FnormalizeCanonic; easy.
apply FnormalizeCanonic; easy.
unfold FtoRradix in |- *; repeat rewrite FnormalizeCorrect; auto.
Qed.

End suc.

Section suc1.
Variable b : Fbound.
Variable radix : Z.
Variable precision : nat.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.

Hypothesis radixMoreThanOne : (1 < radix)%Z.
Hypothesis precisionNotZero : 1 < precision.
Hypothesis pGivesBound : Zpos (vNum b) = Zpower_nat radix precision.

Theorem nNormMimLtvNum : (nNormMin radix precision < pPred (vNum b))%Z.
unfold pPred in |- *;
 rewrite PosNormMin with (radix := radix) (precision := precision);
 auto with zarith.
apply Z.lt_le_trans with (Z.pred (2 * nNormMin radix precision)).
replace (Z.pred (2 * nNormMin radix precision)) with
 (Z.pred (nNormMin radix precision) + nNormMin radix precision)%Z;
 [ idtac | unfold Z.pred in |- *; ring ].
pattern (nNormMin radix precision) at 1 in |- *;
 replace (nNormMin radix precision) with (0 + nNormMin radix precision)%Z;
 [ idtac | ring ].
apply Zplus_lt_compat_r; auto.
apply Zlt_succ_pred.
replace (Z.succ 0) with (Z_of_nat 1); [ idtac | simpl in |- *; auto ].
rewrite <- (Zpower_nat_O radix); unfold nNormMin in |- *.
apply Zpower_nat_monotone_lt. assumption. now apply lt_pred.
apply Zle_Zpred_Zpred. apply Zle_Zmult_comp_r; auto with zarith.
apply Z.lt_le_incl; apply nNormPos; auto with zarith.
Qed.

End suc1.

Section pred.
Variable b : Fbound.
Variable radix : Z.
Variable precision : nat.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.

Hypothesis radixMoreThanOne : (1 < radix)%Z.
Hypothesis precisionNotZero : precision <> 0.
Hypothesis pGivesBound : Zpos (vNum b) = Zpower_nat radix precision.

Definition FPred (x : float) :=
  match Z_eq_bool (Fnum x) (- pPred (vNum b)) with
  | true => Float (- nNormMin radix precision) (Z.succ (Fexp x))
  | false =>
      match Z_eq_bool (Fnum x) (nNormMin radix precision) with
      | true =>
          match Z_eq_bool (Fexp x) (- dExp b) with
          | true => Float (Z.pred (Fnum x)) (Fexp x)
          | false => Float (pPred (vNum b)) (Z.pred (Fexp x))
          end
      | false => Float (Z.pred (Fnum x)) (Fexp x)
      end
  end.

Theorem FPredSimpl1 :
 forall x : float,
 Fnum x = (- pPred (vNum b))%Z ->
 FPred x = Float (- nNormMin radix precision) (Z.succ (Fexp x)).
intros x H'; unfold FPred in |- *.
generalize (Z_eq_bool_correct (Fnum x) (- pPred (vNum b)));
 case (Z_eq_bool (Fnum x) (- pPred (vNum b))); auto.
intros H'0; contradict H'0; auto.
Qed.

Theorem FPredSimpl2 :
 forall x : float,
 Fnum x = nNormMin radix precision ->
 Fexp x <> (- dExp b)%Z -> FPred x = Float (pPred (vNum b)) (Z.pred (Fexp x)).
intros x H' H'0; unfold FPred in |- *.
generalize (Z_eq_bool_correct (Fnum x) (- pPred (vNum b)));
 case (Z_eq_bool (Fnum x) (- pPred (vNum b))); auto.
intros H'1; absurd (0%nat < Fnum x)%Z; auto with zarith arith.
apply Zle_not_lt; rewrite H'1; replace (Z_of_nat 0) with (- (0))%Z;
 [ apply Zle_Zopp | simpl in |- *; auto ].
unfold pPred in |- *; apply Zle_Zpred; red in |- *; simpl in |- *; auto.
rewrite H'.
apply nNormPos; auto with zarith.
intros H'1;
 generalize (Z_eq_bool_correct (Fnum x) (nNormMin radix precision));
 case (Z_eq_bool (Fnum x) (nNormMin radix precision)).
intros H'2; generalize (Z_eq_bool_correct (Fexp x) (- dExp b));
 case (Z_eq_bool (Fexp x) (- dExp b)); auto.
intros H'3; contradict H'0; auto.
intros H'2; contradict H'2; auto.
Qed.

Theorem FPredSimpl3 :
 FPred (Float (nNormMin radix precision) (- dExp b)) =
 Float (Z.pred (nNormMin radix precision)) (- dExp b).
unfold FPred in |- *; simpl in |- *.
generalize (Z_eq_bool_correct (nNormMin radix precision) (- pPred (vNum b)));
 case (Z_eq_bool (nNormMin radix precision) (- pPred (vNum b)));
 auto.
intros H'0; absurd (0 < pPred (vNum b))%Z; auto with zarith arith.
rewrite <- (Z.opp_involutive (pPred (vNum b))); rewrite <- H'0.
apply Zle_not_lt; replace 0%Z with (- (0))%Z;
 [ apply Zle_Zopp | simpl in |- *; auto ].
apply Zlt_le_weak; apply nNormPos; auto with zarith.
unfold pPred in |- *; apply Zlt_succ_pred; simpl in |- *;
 auto with zarith.
simpl in |- *; apply vNumbMoreThanOne with (3 := pGivesBound); auto.
intros H';
 generalize
  (Z_eq_bool_correct (nNormMin radix precision) (nNormMin radix precision));
 case (Z_eq_bool (nNormMin radix precision) (nNormMin radix precision)).
intros H'0; generalize (Z_eq_bool_correct (- dExp b) (- dExp b));
 case (Z_eq_bool (- dExp b) (- dExp b)); auto.
intros H'1; contradict H'1; auto.
intros H'1; contradict H'1; auto.
Qed.

Theorem FPredSimpl4 :
 forall x : float,
 Fnum x <> (- pPred (vNum b))%Z ->
 Fnum x <> nNormMin radix precision ->
 FPred x = Float (Z.pred (Fnum x)) (Fexp x).
intros x H' H'0; unfold FPred in |- *.
generalize (Z_eq_bool_correct (Fnum x) (- pPred (vNum b)));
 case (Z_eq_bool (Fnum x) (- pPred (vNum b))); auto.
intros H'1; contradict H'; auto.
intros H'1;
 generalize (Z_eq_bool_correct (Fnum x) (nNormMin radix precision));
 case (Z_eq_bool (Fnum x) (nNormMin radix precision));
 auto.
intros H'2; contradict H'0; auto.
Qed.

Theorem FPredFopFSucc :
 forall x : float, FPred x = Fopp (FSucc b radix precision (Fopp x)).
intros x.
generalize (Z_eq_bool_correct (Fnum x) (- pPred (vNum b)));
 case (Z_eq_bool (Fnum x) (- pPred (vNum b))); intros H'1.
rewrite FPredSimpl1; auto; rewrite FSuccSimpl1; auto.
unfold Fopp in |- *; simpl in |- *; rewrite H'1; auto with zarith.
generalize (Z_eq_bool_correct (Fnum x) (nNormMin radix precision));
 case (Z_eq_bool (Fnum x) (nNormMin radix precision));
 intros H'2.
generalize (Z_eq_bool_correct (Fexp x) (- dExp b));
 case (Z_eq_bool (Fexp x) (- dExp b)); intros H'3.
replace x with (Float (Fnum x) (Fexp x)).
rewrite H'2; rewrite H'3; rewrite FPredSimpl3; unfold Fopp in |- *;
 simpl in |- *; rewrite FSuccSimpl3; simpl in |- *;
 auto.
rewrite <- Zopp_Zpred_Zs; rewrite Z.opp_involutive; auto.
case x; simpl in |- *; auto.
rewrite FPredSimpl2; auto; rewrite FSuccSimpl2; unfold Fopp in |- *;
 simpl in |- *; try rewrite Z.opp_involutive;
 auto.
rewrite H'2; auto.
rewrite FPredSimpl4; auto; rewrite FSuccSimpl4; auto.
unfold Fopp in |- *; simpl in |- *; rewrite <- Zopp_Zpred_Zs;
 rewrite Z.opp_involutive; auto.
unfold Fopp in |- *; simpl in |- *; contradict H'1; rewrite <- H'1;
 rewrite Z.opp_involutive; auto.
unfold Fopp in |- *; simpl in |- *; contradict H'2; auto with zarith.
Qed.

Theorem FPredDiff1 :
 forall x : float,
 Fnum x <> nNormMin radix precision ->
 Fminus radix x (FPred x) = Float 1%nat (Fexp x) :>R.
intros x H'; rewrite (FPredFopFSucc x).
pattern x at 1 in |- *; rewrite <- (Fopp_Fopp x).
rewrite <- Fopp_Fminus_dist.
rewrite Fopp_Fminus.
unfold FtoRradix in |- *; rewrite FSuccDiff1; auto.
replace (Fnum (Fopp x)) with (- Fnum x)%Z.
contradict H'; rewrite <- (Z.opp_involutive (Fnum x)); rewrite H';
 auto with zarith.
case x; simpl in |- *; auto.
Qed.

Theorem FPredDiff2 :
 forall x : float,
 Fnum x = nNormMin radix precision ->
 Fexp x = (- dExp b)%Z -> Fminus radix x (FPred x) = Float 1%nat (Fexp x) :>R.
intros x H' H'0; rewrite (FPredFopFSucc x).
pattern x at 1 in |- *; rewrite <- (Fopp_Fopp x).
rewrite <- Fopp_Fminus_dist.
rewrite Fopp_Fminus.
unfold FtoRradix in |- *; rewrite FSuccDiff2; auto.
rewrite <- H'; case x; auto.
Qed.

Theorem FPredDiff3 :
 forall x : float,
 Fnum x = nNormMin radix precision ->
 Fexp x <> (- dExp b)%Z ->
 Fminus radix x (FPred x) = Float 1%nat (Z.pred (Fexp x)) :>R.
intros x H' H'0; rewrite (FPredFopFSucc x).
pattern x at 1 in |- *; rewrite <- (Fopp_Fopp x).
rewrite <- Fopp_Fminus_dist.
rewrite Fopp_Fminus.
unfold FtoRradix in |- *; rewrite FSuccDiff3; auto.
rewrite <- H'; case x; auto.
Qed.

Theorem FBoundedPred : forall f : float, Fbounded b f -> Fbounded b (FPred f).
intros f H'; rewrite (FPredFopFSucc f); auto with zarith.
apply oppBounded, FBoundedSuc; try easy.
now apply oppBounded.
Qed.

Theorem FPredCanonic :
 forall a : float, Fcanonic radix b a -> Fcanonic radix b (FPred a).
intros a H'.
rewrite FPredFopFSucc.
apply FcanonicFopp.
apply FSuccCanonic; try easy.
now apply FcanonicFopp.
Qed.

Theorem FPredLt : forall a : float, (FPred a < a)%R.
intros a; rewrite FPredFopFSucc.
pattern a at 2 in |- *; rewrite <- (Fopp_Fopp a).
unfold FtoRradix in |- *; repeat rewrite Fopp_correct.
apply Ropp_lt_contravar.
rewrite <- Fopp_correct; auto with zarith.
apply FSuccLt; try easy.
Qed.

Theorem R0RltRlePred : forall x : float, (0 < x)%R -> (0 <= FPred x)%R.
intros x H'; rewrite FPredFopFSucc.
unfold FtoRradix in |- *; repeat rewrite Fopp_correct.
replace 0%R with (-0)%R; auto with real.
apply Ropp_le_contravar.
apply R0RltRleSucc; auto.
unfold FtoRradix in |- *; repeat rewrite Fopp_correct.
replace 0%R with (-0)%R; auto with real.
Qed.

Theorem FPredProp :
 forall x y : float,
 Fcanonic radix b x -> Fcanonic radix b y -> (x < y)%R -> (x <= FPred y)%R.
intros x y H' H'0 H'1; rewrite FPredFopFSucc.
rewrite <- (Fopp_Fopp x).
unfold FtoRradix in |- *; rewrite Fopp_correct with (x := Fopp x).
rewrite Fopp_correct with (x := FSucc b radix precision (Fopp y));
 auto with real.
apply Ropp_le_contravar.
apply FSuccProp; auto with zarith.
now apply FcanonicFopp.
now apply FcanonicFopp.
repeat rewrite Fopp_correct; auto with real.
Qed.

Definition FNPred (x : float) := FPred (Fnormalize radix b precision x).

Theorem FNPredFopFNSucc :
 forall x : float, FNPred x = Fopp (FNSucc b radix precision (Fopp x)).
intros x; unfold FNPred, FNSucc in |- *; auto.
rewrite Fnormalize_Fopp; auto.
apply FPredFopFSucc; auto.
Qed.

Theorem FNPredCanonic :
 forall a : float, Fbounded b a -> Fcanonic radix b (FNPred a).
intros a H'; unfold FNPred in |- *.
apply FPredCanonic; auto with zarith.
apply FnormalizeCanonic; easy.
Qed.

Theorem FNPredLt : forall a : float, (FNPred a < a)%R.
intros a; unfold FNPred in |- *.
unfold FtoRradix in |- *;
 rewrite <- (FnormalizeCorrect _ radixMoreThanOne b precision a).
apply FPredLt; auto.
Qed.

Theorem FPredSuc :
 forall x : float,
 Fcanonic radix b x -> FPred (FSucc b radix precision x) = x.
intros x H; unfold FPred, FSucc in |- *.
cut (Fbounded b x); [ intros Fb0 | apply FcanonicBound with (1 := H) ].
generalize (Z_eq_bool_correct (Fnum x) (pPred (vNum b)));
 case (Z_eq_bool (Fnum x) (pPred (vNum b))); simpl in |- *.
generalize (Z_eq_bool_correct (nNormMin radix precision) (- pPred (vNum b)));
 case (Z_eq_bool (nNormMin radix precision) (- pPred (vNum b)));
 simpl in |- *.
intros H'; contradict H'; apply sym_not_equal; apply Zlt_not_eq; auto.
apply Z.lt_le_trans with (- 0%nat)%Z.
apply Zlt_Zopp; unfold pPred in |- *; apply Zlt_succ_pred; simpl in |- *;
 apply vNumbMoreThanOne with (3 := pGivesBound); auto.
simpl in |- *; apply Zlt_le_weak; apply nNormPos; auto.
generalize
 (Z_eq_bool_correct (nNormMin radix precision) (nNormMin radix precision));
 case (Z_eq_bool (nNormMin radix precision) (nNormMin radix precision));
 simpl in |- *.
generalize (Z_eq_bool_correct (Z.succ (Fexp x)) (- dExp b));
 case (Z_eq_bool (Z.succ (Fexp x)) (- dExp b)); simpl in |- *.
intros H' H'0 H'1 H'2; absurd (- dExp b <= Fexp x)%Z; try apply Fb0.
rewrite <- H'; auto with zarith.
replace (Z.pred (Z.succ (Fexp x))) with (Fexp x);
 [ idtac | unfold Z.succ, Z.pred in |- *; ring ]; auto.
intros H' H'0 H'1 H'2; rewrite <- H'2; auto.
apply floatEq; auto.
intros H'; case H'; auto.
generalize (Z_eq_bool_correct (Fnum x) (- nNormMin radix precision));
 case (Z_eq_bool (Fnum x) (- nNormMin radix precision));
 simpl in |- *.
generalize (Z_eq_bool_correct (Fexp x) (- dExp b));
 case (Z_eq_bool (Fexp x) (- dExp b)); simpl in |- *.
generalize (Z_eq_bool_correct (Z.succ (Fnum x)) (- pPred (vNum b)));
 case (Z_eq_bool (Z.succ (Fnum x)) (- pPred (vNum b)));
 simpl in |- *.
intros H0 H1 H2; absurd (Z.succ (Fnum x) <= Fnum x)%Z; auto with zarith.
rewrite H0; rewrite H2; (apply Zle_Zopp; auto with zarith).
unfold pPred in |- *; apply Zle_Zpred; apply ZltNormMinVnum; auto with zarith.
generalize (Z_eq_bool_correct (Z.succ (Fnum x)) (nNormMin radix precision));
 case (Z_eq_bool (Z.succ (Fnum x)) (nNormMin radix precision));
 simpl in |- *.
intros H' H'0 H'1 H'2; contradict H'2.
rewrite <- H'; auto with zarith.
replace (Z.pred (Z.succ (Fnum x))) with (Fnum x);
 [ idtac | unfold Z.succ, Z.pred in |- *; ring ]; auto.
intros H' H'0 H'1 H'2 H'3; apply floatEq; auto.
generalize (Z_eq_bool_correct (- pPred (vNum b)) (- pPred (vNum b)));
 case (Z_eq_bool (- pPred (vNum b)) (- pPred (vNum b)));
 auto.
intros H' H'0 H'1 H'2; rewrite <- H'1.
replace (Z.succ (Z.pred (Fexp x))) with (Fexp x);
 [ idtac | unfold Z.succ, Z.pred in |- *; ring ]; auto.
apply floatEq; auto.
intros H'; case H'; auto.
generalize (Z_eq_bool_correct (Z.succ (Fnum x)) (- pPred (vNum b)));
 case (Z_eq_bool (Z.succ (Fnum x)) (- pPred (vNum b)));
 simpl in |- *.
intros H'; absurd (- pPred (vNum b) <= Fnum x)%Z; auto with zarith.
apply Zle_Zabs_inv1; try apply Fb0.
unfold pPred in |- *; apply Zle_Zpred; try apply Fb0.
generalize (Z_eq_bool_correct (Z.succ (Fnum x)) (nNormMin radix precision));
 case (Z_eq_bool (Z.succ (Fnum x)) (nNormMin radix precision));
 simpl in |- *.
generalize (Z_eq_bool_correct (Fexp x) (- dExp b));
 case (Z_eq_bool (Fexp x) (- dExp b)); simpl in |- *.
intros H' H'0 H'1 H'2 H'3.
replace (Z.pred (Z.succ (Fnum x))) with (Fnum x);
 [ idtac | unfold Z.succ, Z.pred in |- *; ring ]; auto.
apply floatEq; auto.
intros H' H'0 H'1 H'2 H'3; case H.
intros H'4; absurd (nNormMin radix precision <= Z.abs (Fnum x))%Z.
replace (Fnum x) with (Z.pred (Z.succ (Fnum x)));
 [ idtac | unfold Z.succ, Z.pred in |- *; ring ]; auto.
rewrite H'0.
apply Zlt_not_le; rewrite Z.abs_eq; auto with zarith.
apply Zle_Zpred; apply nNormPos; auto with zarith.
apply pNormal_absolu_min with (b := b); auto.
intros H'4; contradict H'; apply FsubnormalFexp with (1 := H'4).
intros H' H'0 H'1 H'2; apply floatEq; simpl in |- *; auto.
unfold Z.pred, Z.succ in |- *; ring.
Qed.

Theorem FSucPred :
 forall x : float,
 Fcanonic radix b x -> FSucc b radix precision (FPred x) = x.
intros x H; unfold FPred, FSucc in |- *.
cut (Fbounded b x); [ intros Fb0 | apply FcanonicBound with (1 := H) ].
generalize (Z_eq_bool_correct (Fnum x) (- pPred (vNum b)));
 case (Z_eq_bool (Fnum x) (- pPred (vNum b))); simpl in |- *.
generalize (Z_eq_bool_correct (- nNormMin radix precision) (pPred (vNum b)));
 case (Z_eq_bool (- nNormMin radix precision) (pPred (vNum b)));
 simpl in |- *.
intros H'; contradict H'; apply Zlt_not_eq; auto.
rewrite <- (Z.opp_involutive (pPred (vNum b))); apply Zlt_Zopp.
apply Z.lt_le_trans with (- 0%nat)%Z.
apply Zlt_Zopp; unfold pPred in |- *; apply Zlt_succ_pred; simpl in |- *.
apply (vNumbMoreThanOne radix) with (precision := precision); auto.
simpl in |- *; apply Zlt_le_weak; apply nNormPos; auto with zarith arith.
generalize
 (Z_eq_bool_correct (- nNormMin radix precision) (- nNormMin radix precision));
 case (Z_eq_bool (- nNormMin radix precision) (- nNormMin radix precision));
 simpl in |- *.
generalize (Z_eq_bool_correct (Z.succ (Fexp x)) (- dExp b));
 case (Z_eq_bool (Z.succ (Fexp x)) (- dExp b)); simpl in |- *.
intros H' H'0 H'1 H'2; absurd (- dExp b <= Fexp x)%Z; try apply Fb0.
rewrite <- H'; auto with zarith.
intros H' H'0 H'1 H'2; rewrite <- H'2; apply floatEq; simpl in |- *; auto;
 unfold Z.succ, Z.pred in |- *; ring.
intros H'; case H'; auto.
generalize (Z_eq_bool_correct (Fnum x) (nNormMin radix precision));
 case (Z_eq_bool (Fnum x) (nNormMin radix precision));
 simpl in |- *.
generalize (Z_eq_bool_correct (Fexp x) (- dExp b));
 case (Z_eq_bool (Fexp x) (- dExp b)); simpl in |- *.
generalize (Z_eq_bool_correct (Z.pred (Fnum x)) (pPred (vNum b)));
 case (Z_eq_bool (Z.pred (Fnum x)) (pPred (vNum b)));
 simpl in |- *.
intros H' H'0 H'1 H'2; absurd (nNormMin radix precision <= pPred (vNum b))%Z.
rewrite <- H'; rewrite H'1; auto with zarith.
apply Zle_Zpred, ZltNormMinVnum; easy.
generalize (Z_eq_bool_correct (Z.pred (Fnum x)) (- nNormMin radix precision));
 case (Z_eq_bool (Z.pred (Fnum x)) (- nNormMin radix precision));
 simpl in |- *.
intros H' H'0 H'1 H'2 H'3;
 absurd (Z.pred (nNormMin radix precision) = (- nNormMin radix precision)%Z);
 auto with zarith.
intros H' H'0 H'1 H'2 H'3; apply floatEq; simpl in |- *; auto;
 unfold Z.pred, Z.succ in |- *; ring.
generalize (Z_eq_bool_correct (pPred (vNum b)) (pPred (vNum b)));
 case (Z_eq_bool (pPred (vNum b)) (pPred (vNum b)));
 auto.
intros H' H'0 H'1 H'2; rewrite <- H'1; apply floatEq; simpl in |- *; auto;
 unfold Z.pred, Z.succ in |- *; ring.
intros H'; case H'; auto.
generalize (Z_eq_bool_correct (Z.pred (Fnum x)) (pPred (vNum b)));
 case (Z_eq_bool (Z.pred (Fnum x)) (pPred (vNum b)));
 simpl in |- *.
intros H'; absurd (Fnum x <= pPred (vNum b))%Z; auto with zarith.
apply Zle_Zabs_inv2; unfold pPred in |- *; apply Zle_Zpred, Fb0.
generalize (Z_eq_bool_correct (Z.pred (Fnum x)) (- nNormMin radix precision));
 case (Z_eq_bool (Z.pred (Fnum x)) (- nNormMin radix precision));
 simpl in |- *.
generalize (Z_eq_bool_correct (Fexp x) (- dExp b));
 case (Z_eq_bool (Fexp x) (- dExp b)); simpl in |- *.
intros H' H'0 H'1 H'2 H'3; apply floatEq; simpl in |- *; auto;
 unfold Z.succ, Z.pred in |- *; ring.
intros H' H'0 H'1 H'2 H'3; case H; intros C0.
absurd (nNormMin radix precision <= Z.abs (Fnum x))%Z.
replace (Fnum x) with (Z.succ (Z.pred (Fnum x)));
 [ idtac | unfold Z.succ, Z.pred in |- *; ring ].
rewrite H'0.
rewrite <- Zopp_Zpred_Zs; rewrite Zabs_Zopp.
rewrite Z.abs_eq; auto with zarith.
apply Zle_Zpred; simpl in |- *; apply nNormPos; auto with zarith.
apply pNormal_absolu_min with (b := b); auto.
contradict H'; apply FsubnormalFexp with (1 := C0).
intros H' H'0 H'1 H'2; apply floatEq; simpl in |- *; auto.
unfold Z.pred, Z.succ in |- *; ring.
Qed.

End pred.
Section FMinMax.
Variable b : Fbound.
Variable radix : Z.
Variable precision : nat.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.
Hypothesis radixMoreThanOne : (1 < radix)%Z.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ radixMoreThanOne).

Hypothesis precisionNotZero : precision <> 0.
Hypothesis pGivesBound : Zpos (vNum b) = Zpower_nat radix precision.

Definition boundNat (n : nat) := Float 1%nat (digit radix n).

Theorem boundNatCorrect : forall n : nat, (n < boundNat n)%R.
intros n; unfold FtoRradix, FtoR, boundNat in |- *; simpl in |- *.
rewrite Rmult_1_l.
rewrite <- Zpower_nat_Z_powerRZ; auto with real zarith.
rewrite INR_IZR_INZ; auto with real zarith.
apply Rle_lt_trans with (Z.abs n); [rewrite (Z.abs_eq (Z_of_nat n))|idtac];auto with real zarith.
apply Rlt_IZR, digitMore; easy.
Qed.


Definition boundR (r : R) := boundNat (Z.abs_nat (up (Rabs r))).

Theorem boundRCorrect1 : forall r : R, (r < boundR r)%R.
intros r; case (Rle_or_lt r 0); intros H'.
apply Rle_lt_trans with (1 := H').
unfold boundR, boundNat, FtoRradix, FtoR in |- *; simpl in |- *;
 auto with real.
rewrite Rmult_1_l; auto with real zarith.
apply powerRZ_lt, Rlt_IZR; omega.
apply Rlt_trans with (2 := boundNatCorrect (Z.abs_nat (up (Rabs r)))).
replace (Rabs r) with r; auto with real.
apply Rlt_le_trans with (r2 := IZR (up r)); auto with real zarith.
case (archimed r); auto.
rewrite INR_IZR_INZ; auto with real zarith.
apply Rle_IZR, Zle_abs.
unfold Rabs in |- *; case (Rcase_abs r); auto with real.
intros H'0; contradict H'0; auto with real.
Qed.

Theorem boundRrOpp : forall r : R, boundR r = boundR (- r).
intros R; unfold boundR in |- *.
rewrite Rabs_Ropp; auto.
Qed.

Theorem boundRCorrect2 : forall r : R, (Fopp (boundR r) < r)%R.
intros r; case (Rle_or_lt r 0); intros H'.
rewrite boundRrOpp.
pattern r at 2 in |- *; rewrite <- (Ropp_involutive r).
unfold FtoRradix in |- *; rewrite Fopp_correct.
apply Ropp_lt_contravar; apply boundRCorrect1; auto.
apply Rle_lt_trans with 0%R; auto.
replace 0%R with (-0)%R; auto with real.
unfold FtoRradix in |- *; rewrite Fopp_correct.
apply Ropp_le_contravar.
unfold boundR, boundNat, FtoRradix, FtoR in |- *; simpl in |- *;
 auto with real zarith.
rewrite Rmult_1_l; apply Rlt_le; auto with real zarith arith.
apply powerRZ_lt, Rlt_IZR; omega.
Qed.

Definition mBFloat (p : R) :=
  map (fun p : Z * Z => Float (fst p) (snd p))
    (mProd Z Z (Z * Z)
       (mZlist (- pPred (vNum b)) (pPred (vNum b)))
       (mZlist (- dExp b) (Fexp (boundR p)))).

Theorem mBFadic_correct1 :
 forall (r : R) (q : float),
 ~ is_Fzero q ->
 (Fopp (boundR r) < q)%R ->
 (q < boundR r)%R -> Fbounded b q -> In q (mBFloat r).
intros r q.
case (Zle_or_lt (Fexp (boundR r)) (Fexp q)); intros H'.
intros H'0 H'1 H'2 H'3; case H'0.
apply is_Fzero_rep2 with (radix := radix); auto.
rewrite <-
 FshiftCorrect with (n := Z.abs_nat (Fexp q - Fexp (boundR r))) (x := q);
 auto with arith.
apply is_Fzero_rep1 with (radix := radix).
unfold is_Fzero in |- *.
cut (forall p : Z, (- 1%nat < p)%Z -> (p < 1%nat)%Z -> p = 0%Z);
 [ intros tmp; apply tmp | idtac ].
replace (- 1%nat)%Z with (Fnum (Fopp (boundR r))).
apply Rlt_Fexp_eq_Zlt with (radix := radix); auto with real zarith.
rewrite FshiftCorrect; auto.
unfold Fshift in |- *; simpl in |- *.
rewrite (fun x y => inj_abs (x - y)); auto with zarith.
apply Zle_Zminus_ZERO; assumption.
now simpl.
replace (Z_of_nat 1) with (Fnum (boundR r)).
apply Rlt_Fexp_eq_Zlt with (radix := radix); auto with zarith.
rewrite FshiftCorrect; auto.
unfold Fshift in |- *; simpl in |- *.
rewrite inj_abs; auto with zarith.
apply Zle_Zminus_ZERO; assumption.
now simpl.
intros p0; case p0; simpl in |- *; auto with zarith.
intros H'0 H'1 H'2 H'3; unfold mBFloat in |- *.
replace q with
 ((fun p : Z * Z => Float (fst p) (snd p)) (Fnum q, Fexp q)).
apply in_map with (f := fun p : Z * Z => Float (fst p) (snd p));
 auto.
apply mProd_correct; auto.
apply mZlist_correct; auto.
apply Zle_Zabs_inv1; apply Zle_Zpred, H'3.
apply Zle_Zabs_inv2; apply Zle_Zpred, H'3.
apply mZlist_correct; try apply H'3.
auto with zarith.
case q; simpl in |- *; auto with zarith.
Qed.

Theorem mBFadic_correct3 : forall r : R, In (Fopp (boundR r)) (mBFloat r).
intros r; unfold mBFloat in |- *.
replace (Fopp (boundR r)) with
 ((fun p : Z * Z => Float (fst p) (snd p))
    (Fnum (Fopp (boundR r)), Fexp (Fopp (boundR r)))).
apply in_map with (f := fun p : Z * Z => Float (fst p) (snd p));
 auto.
apply mProd_correct; auto.
apply mZlist_correct; auto.
unfold boundR, boundNat in |- *; simpl in |- *; auto with zarith.
replace (-1)%Z with (- Z_of_nat 1)%Z; auto with zarith.
apply Zle_Zopp.
unfold pPred in |- *; apply Zle_Zpred; simpl in |- *.
apply (vNumbMoreThanOne radix) with (precision := precision);
 auto with zarith.
unfold pPred in |- *; apply Zle_Zpred; simpl in |- *.
red in |- *; simpl in |- *; auto.
apply mZlist_correct; auto.
unfold boundR, boundNat in |- *; simpl in |- *; auto with zarith.
case (dExp b); try intros; try rewrite N2Z.inj_pos; auto with zarith arith.
auto with zarith.
Qed.

Theorem mBFadic_correct4 :
 forall r : R, In (Float 0%nat (- dExp b)) (mBFloat r).
intros p; unfold mBFloat in |- *.
replace (Float 0%nat (- dExp b)) with
 ((fun p : Z * Z => Float (fst p) (snd p))
    (Fnum (Float 0%nat (- dExp b)), Fexp (Float 0%nat (- dExp b)))).
apply in_map with (f := fun p : Z * Z => Float (fst p) (snd p));
 auto.
apply mProd_correct; auto.
apply mZlist_correct; auto.
simpl in |- *; auto with zarith.
replace 0%Z with (- (0))%Z; [ idtac | simpl in |- *; auto ].
apply Zle_Zopp; unfold pPred in |- *; apply Zle_Zpred.
red in |- *; simpl in |- *; auto with zarith.
simpl in |- *; auto with zarith.
unfold pPred in |- *; apply Zle_Zpred.
red in |- *; simpl in |- *; auto with zarith.
apply mZlist_correct; auto.
simpl in |- *; auto with zarith.
unfold boundR, boundNat in |- *; simpl in |- *; auto with zarith.
case (dExp b); auto with zarith.
Qed.

Theorem mBPadic_Fbounded :
 forall (p : float) (r : R), In p (mBFloat r) -> Fbounded b p.
intros p r H'; red in |- *; repeat (split; auto).
apply Zpred_Zle_Zabs_intro.
apply mZlist_correct_rev1 with (q := Z.pred (Zpos (vNum b)));
 auto with real.
apply
 mProd_correct_rev1
  with
    (l2 := mZlist (- dExp b) (Fexp (boundR r)))
    (C := (Z * Z)%type)
    (b := Fexp p); auto.
apply
 in_map_inv with (f := fun p : Z * Z => Float (fst p) (snd p));
 auto.
intros a1 b1; case a1; case b1; simpl in |- *.
intros z z0 z1 z2 H'0; inversion H'0; auto.
generalize H'; case p; auto.
apply mZlist_correct_rev2 with (p := (- Z.pred (Zpos (vNum b)))%Z);
 auto.
apply
 mProd_correct_rev1
  with
    (l2 := mZlist (- dExp b) (Fexp (boundR r)))
    (C := (Z * Z)%type)
    (b := Fexp p); auto.
apply
 in_map_inv with (f := fun p : Z * Z => Float (fst p) (snd p));
 auto.
intros a1 b1; case a1; case b1; simpl in |- *.
intros z z0 z1 z2 H'0; inversion H'0; auto.
generalize H'; case p; auto.
apply mZlist_correct_rev1 with (q := Fexp (boundR r)); auto.
apply
 mProd_correct_rev2
  with
    (l1 := mZlist (- pPred (vNum b)) (pPred (vNum b)))
    (C := (Z * Z)%type)
    (a := Fnum p); auto.
apply
 in_map_inv with (f := fun p : Z * Z => Float (fst p) (snd p));
 auto.
intros a1 b1; case a1; case b1; simpl in |- *.
intros z z0 z1 z2 H'0; inversion H'0; auto.
generalize H'; case p; auto.
Qed.


Definition ProjectorP (P : R -> float -> Prop) :=
  forall p q : float, Fbounded b p -> P p q -> p = q :>R.

Definition MonotoneP (P : R -> float -> Prop) :=
  forall (p q : R) (p' q' : float),
  (p < q)%R -> P p p' -> P q q' -> (p' <= q')%R.

Definition isMin (r : R) (min : float) :=
  Fbounded b min /\
  (min <= r)%R /\
  (forall f : float, Fbounded b f -> (f <= r)%R -> (f <= min)%R).

Theorem isMin_inv1 : forall (p : float) (r : R), isMin r p -> (p <= r)%R.
intros p r H; case H; intros H1 H2; case H2; auto.
Qed.

Theorem ProjectMin : ProjectorP isMin.
red in |- *.
intros p q H' H'0; apply Rle_antisym.
elim H'0; intros H'1 H'2; elim H'2; intros H'3 H'4; apply H'4; clear H'2;
 auto with real.
apply isMin_inv1 with (1 := H'0); auto.
Qed.

Theorem MonotoneMin : MonotoneP isMin.
red in |- *.
intros p q p' q' H' H'0 H'1.
elim H'1; intros H'2 H'3; elim H'3; intros H'4 H'5; apply H'5; clear H'3 H'1;
 auto.
case H'0; auto.
apply Rle_trans with p; auto.
apply isMin_inv1 with (1 := H'0); auto.
apply Rlt_le; auto.
Qed.

Definition isMax (r : R) (max : float) :=
  Fbounded b max /\
  (r <= max)%R /\
  (forall f : float, Fbounded b f -> (r <= f)%R -> (max <= f)%R).

Theorem isMax_inv1 : forall (p : float) (r : R), isMax r p -> (r <= p)%R.
intros p r H; case H; intros H1 H2; case H2; auto.
Qed.

Theorem ProjectMax : ProjectorP isMax.
red in |- *.
intros p q H' H'0; apply Rle_antisym.
apply isMax_inv1 with (1 := H'0); auto.
elim H'0; intros H'1 H'2; elim H'2; intros H'3 H'4; apply H'4; clear H'2;
 auto with real.
Qed.

Theorem MonotoneMax : MonotoneP isMax.
red in |- *.
intros p q p' q' H' H'0 H'1.
elim H'0; intros H'2 H'3; elim H'3; intros H'4 H'5; apply H'5; clear H'3 H'0.
case H'1; auto.
apply Rle_trans with q; auto.
apply Rlt_le; auto.
apply isMax_inv1 with (1 := H'1); auto.
Qed.

Theorem MinEq :
 forall (p q : float) (r : R), isMin r p -> isMin r q -> p = q :>R.
intros p q r H' H'0; apply Rle_antisym.
elim H'0; intros H'1 H'2; elim H'2; intros H'3 H'4; apply H'4; clear H'2 H'0;
 auto.
case H'; auto.
apply isMin_inv1 with (1 := H'); auto.
elim H'; intros H'1 H'2; elim H'2; intros H'3 H'4; apply H'4; clear H'2 H';
 auto.
case H'0; auto.
apply isMin_inv1 with (1 := H'0); auto.
Qed.

Theorem MaxEq :
 forall (p q : float) (r : R), isMax r p -> isMax r q -> p = q :>R.
intros p q r H' H'0; apply Rle_antisym.
elim H'; intros H'1 H'2; elim H'2; intros H'3 H'4; apply H'4; clear H'2 H';
 auto.
case H'0; auto.
apply isMax_inv1 with (1 := H'0); auto.
elim H'0; intros H'1 H'2; elim H'2; intros H'3 H'4; apply H'4; clear H'2 H'0;
 auto.
case H'; auto.
apply isMax_inv1 with (1 := H'); auto.
Qed.

Theorem MinOppMax :
 forall (p : float) (r : R), isMin r p -> isMax (- r) (Fopp p).
intros p r H'; split.
apply oppBounded; case H'; auto.
split.
unfold FtoRradix in |- *; rewrite Fopp_correct.
apply Ropp_le_contravar; apply isMin_inv1 with (1 := H'); auto.
intros f H'0 H'1.
rewrite <- (Fopp_Fopp f).
unfold FtoRradix in |- *; rewrite Fopp_correct; rewrite Fopp_correct.
apply Ropp_le_contravar.
elim H'.
intros H'2 H'3; elim H'3; intros H'4 H'5; apply H'5; clear H'3.
apply oppBounded; case H'; auto.
rewrite <- (Ropp_involutive r).
unfold FtoRradix in |- *; rewrite Fopp_correct; auto with real.
Qed.

Theorem MaxOppMin :
 forall (p : float) (r : R), isMax r p -> isMin (- r) (Fopp p).
intros p r H'; split.
apply oppBounded; case H'; auto.
split.
unfold FtoRradix in |- *; rewrite Fopp_correct.
apply Ropp_le_contravar; apply isMax_inv1 with (1 := H'); auto.
intros f H'0 H'1.
rewrite <- (Fopp_Fopp f).
unfold FtoRradix in |- *; repeat rewrite Fopp_correct.
apply Ropp_le_contravar.
rewrite <- (Fopp_correct radix f).
elim H'.
intros H'2 H'3; elim H'3; intros H'4 H'5; apply H'5; clear H'3.
apply oppBounded; auto.
rewrite <- (Ropp_involutive r).
unfold FtoRradix in |- *; rewrite Fopp_correct; auto with real.
Qed.

Theorem MinMax :
 forall (p : float) (r : R),
 isMin r p -> r <> p :>R -> isMax r (FNSucc b radix precision p).
intros p r H' H'0.
split.
apply FcanonicBound with (radix := radix); auto with zarith.
apply FNSuccCanonic; auto.
inversion H'; auto.
split.
case (Rle_or_lt (FNSucc b radix precision p) r); intros H'2; auto.
absurd (FNSucc b radix precision p <= p)%R.
apply Rlt_not_le.
unfold FtoRradix in |- *; apply FNSuccLt; auto.
inversion H'; auto.
elim H0; intros H'1 H'3; apply H'3; auto.
apply FcanonicBound with (radix := radix); auto with zarith.
apply FNSuccCanonic; auto.
apply Rlt_le; auto.
intros f H'2 H'3.
replace (FtoRradix f) with (FtoRradix (Fnormalize radix b precision f)).
unfold FtoRradix in |- *; apply FNSuccProp; auto.
inversion H'; auto.
apply FcanonicBound with (radix := radix); auto with zarith.
apply FnormalizeCanonic; easy.
apply Rlt_le_trans with r; auto.
case (Rle_or_lt r p); auto.
intros H'4; contradict H'0.
apply Rle_antisym; auto; apply isMin_inv1 with (1 := H'); auto.
rewrite FnormalizeCorrect; auto.
unfold FtoRradix in |- *; apply FnormalizeCorrect; auto.
Qed.

Theorem MinExList :
 forall (r : R) (L : list float),
 (forall f : float, In f L -> (r < f)%R) \/
 (exists min : float,
    In min L /\
    (min <= r)%R /\ (forall f : float, In f L -> (f <= r)%R -> (f <= min)%R)).
intros r L; elim L; simpl in |- *; auto.
left; intros f H'; elim H'.
intros a l H'.
elim H';
 [ intros H'0; clear H'
 | intros H'0; elim H'0; intros min E; elim E; intros H'1 H'2; elim H'2;
    intros H'3 H'4; try exact H'4; clear H'2 E H'0 H' ].
case (Rle_or_lt a r); intros H'1.
right; exists a; repeat split; auto.
intros f H'; elim H';
 [ intros H'2; rewrite <- H'2; clear H' | intros H'2; clear H' ];
 auto with real.
intros H'; contradict H'; auto with real.
apply Rlt_not_le; auto with real.
left; intros f H'; elim H';
 [ intros H'2; rewrite <- H'2; clear H' | intros H'2; clear H' ];
 auto.
case (Rle_or_lt a min); intros H'5.
right; exists min; repeat split; auto.
intros f H'; elim H';
 [ intros H'0; rewrite <- H'0; clear H' | intros H'0; clear H' ];
 auto.
case (Rle_or_lt a r); intros H'6.
right; exists a; repeat split; auto.
intros f H'; elim H';
 [ intros H'0; rewrite <- H'0; clear H' | intros H'0; clear H' ];
 auto with real.
intros H'; apply Rle_trans with (FtoRradix min); auto with real.
right; exists min; split; auto; split; auto.
intros f H'; elim H';
 [ intros H'0; elim H'0; clear H' | intros H'0; clear H' ];
 auto.
intros H'; contradict H'6; auto with real.
apply Rle_not_lt; auto.
Qed.

Theorem MinEx : forall r : R, exists min : float, isMin r min.
intros r.
case (MinExList r (mBFloat r)).
intros H'0; absurd (Fopp (boundR r) <= r)%R; auto.
apply Rlt_not_le.
apply H'0.
apply mBFadic_correct3; auto.
apply Rlt_le.
apply boundRCorrect2; auto.
intros H'0; elim H'0; intros min E; elim E; intros H'1 H'2; elim H'2;
 intros H'3 H'4; clear H'2 E H'0.
exists min; split; auto.
apply mBPadic_Fbounded with (r := r); auto.
split; auto.
intros f H'0 H'2.
case (Req_dec f 0); intros H'6.
replace (FtoRradix f) with (FtoRradix (Float 0%nat (- dExp b))).
apply H'4; auto.
apply mBFadic_correct4; auto.
replace (FtoRradix (Float 0%nat (- dExp b))) with (FtoRradix f); auto.
rewrite H'6.
unfold FtoRradix, FtoR in |- *; simpl in |- *; auto with real.
rewrite H'6.
unfold FtoRradix, FtoR in |- *; simpl in |- *; auto with real.
case (Rle_or_lt f (Fopp (boundR r))); intros H'5.
apply Rle_trans with (FtoRradix (Fopp (boundR r))); auto.
apply H'4; auto.
apply mBFadic_correct3; auto.
apply Rlt_le.
apply boundRCorrect2; auto.
case (Rle_or_lt (boundR r) f); intros H'7.
contradict H'2; apply Rlt_not_le.
apply Rlt_le_trans with (FtoRradix (boundR r)); auto.
apply boundRCorrect1; auto.
apply H'4; auto.
apply mBFadic_correct1; auto.
contradict H'6; unfold FtoRradix in |- *; apply is_Fzero_rep1; auto.
Qed.

Theorem MaxEx : forall r : R, exists max : float, isMax r max.
intros r; case (MinEx r).
intros x H'.
case (Req_dec x r); intros H'1.
exists x.
rewrite <- H'1.
red in |- *; split; [ case H' | split ]; auto with real.
exists (FNSucc b radix precision x).
apply MinMax; auto.
Qed.

Theorem FminRep :
 forall p q : float,
 isMin p q -> exists m : Z, q = Float m (Fexp p) :>R.
intros p q H'.
replace (FtoRradix q) with (FtoRradix (Fnormalize radix b precision q)).
2: unfold FtoRradix in |- *; apply FnormalizeCorrect; auto.
case (Zle_or_lt (Fexp (Fnormalize radix b precision q)) (Fexp p)); intros H'1.
exists (Fnum p).
unfold FtoRradix in |- *; apply FSuccZleEq with (3 := pGivesBound); auto.
replace (Float (Fnum p) (Fexp p)) with p; [ idtac | case p ]; auto.
replace (FtoR radix (Fnormalize radix b precision q)) with (FtoR radix q);
 [ idtac | rewrite FnormalizeCorrect ]; auto.
apply isMin_inv1 with (1 := H'); auto.
replace (FSucc b radix precision (Fnormalize radix b precision q)) with
 (FNSucc b radix precision q); [ idtac | case p ];
 auto.
replace (Float (Fnum p) (Fexp p)) with p; [ idtac | case p ]; auto.
case (Req_dec p q); intros Eq0.
unfold FtoRradix in Eq0; rewrite Eq0.
apply FNSuccLt; auto.
case (MinMax q p); auto.
intros H'2 H'3; elim H'3; intros H'4 H'5; clear H'3.
case H'4; auto.
intros H'0; absurd (p <= q)%R; rewrite H'0; auto.
apply Rlt_not_le; auto.
unfold FtoRradix in |- *; apply FNSuccLt; auto.
inversion H'.
elim H0; intros H'3 H'6; apply H'6; clear H0; auto.
rewrite <- H'0; auto with real.
exists
 (Fnum
    (Fshift radix (Z.abs_nat (Fexp (Fnormalize radix b precision q) - Fexp p))
       (Fnormalize radix b precision q))).
pattern (Fexp p) at 2 in |- *;
 replace (Fexp p) with
  (Fexp
     (Fshift radix
        (Z.abs_nat (Fexp (Fnormalize radix b precision q) - Fexp p))
        (Fnormalize radix b precision q))).
unfold FtoRradix in |- *;
 rewrite <-
  FshiftCorrect
                with
                (n :=
                  Z.abs_nat (Fexp (Fnormalize radix b precision q) - Fexp p))
               (x := Fnormalize radix b precision q).
case
 (Fshift radix (Z.abs_nat (Fexp (Fnormalize radix b precision q) - Fexp p))
    (Fnormalize radix b precision q)); auto.
auto with arith.
simpl in |- *; rewrite inj_abs; auto with zarith.
Qed.

Theorem MaxMin :
 forall (p : float) (r : R),
 isMax r p -> r <> p :>R -> isMin r (FNPred b radix precision p).
intros p r H' H'0.
rewrite <- (Fopp_Fopp (FNPred b radix precision p)).
rewrite <- (Ropp_involutive r).
apply MaxOppMin.
rewrite FNPredFopFNSucc; auto.
rewrite Fopp_Fopp; auto.
apply MinMax; auto.
apply MaxOppMin; auto.
contradict H'0.
rewrite <- (Ropp_involutive r); rewrite H'0; auto; unfold FtoRradix in |- *;
 rewrite Fopp_correct; auto; apply Ropp_involutive.
Qed.

Theorem FmaxRep :
 forall p q : float,
 isMax p q -> exists m : Z, q = Float m (Fexp p) :>R.
intros p q H'; case (FminRep (Fopp p) (Fopp q)).
unfold FtoRradix in |- *; rewrite Fopp_correct.
apply MaxOppMin; auto.
intros x H'0.
exists (- x)%Z.
rewrite <- (Ropp_involutive (FtoRradix q)).
unfold FtoRradix in |- *; rewrite <- Fopp_correct.
unfold FtoRradix in H'0; rewrite H'0.
unfold FtoR in |- *; simpl in |- *; auto with real.
rewrite Ropp_Ropp_IZR; rewrite Ropp_mult_distr_l_reverse; auto.
Qed.

End FMinMax.


Section FOdd.
Variable b : Fbound.
Variable radix : Z.
Variable precision : nat.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.

Hypothesis radixMoreThanOne : (1 < radix)%Z.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ radixMoreThanOne).

Hypothesis precisionGreaterThanOne : 1 < precision.
Hypothesis pGivesBound : Zpos (vNum b) = Zpower_nat radix precision.

Definition Even (z : Z) : Prop := exists z1 : _, z = (2 * z1)%Z.

Definition Odd (z : Z) : Prop := exists z1 : _, z = (2 * z1 + 1)%Z.

Theorem OddSEven : forall n : Z, Odd n -> Even (Z.succ n).
intros n H'; case H'; intros m H'1; exists (Z.succ m).
rewrite H'1; unfold Z.succ in |- *; ring.
Qed.

Theorem EvenSOdd : forall n : Z, Even n -> Odd (Z.succ n).
intros n H'; case H'; intros m H'1; exists m.
rewrite H'1; unfold Z.succ in |- *; ring.
Qed.

Theorem OddSEvenInv : forall n : Z, Odd (Z.succ n) -> Even n.
intros n H'; case H'; intros m H'1; exists m.
apply Z.succ_inj; rewrite H'1; (unfold Z.succ in |- *; ring).
Qed.

Theorem EvenSOddInv : forall n : Z, Even (Z.succ n) -> Odd n.
intros n H'; case H'; intros m H'1; exists (Z.pred m).
apply Z.succ_inj; rewrite H'1; (unfold Z.succ, Z.pred in |- *; ring).
Qed.

Theorem EvenO : Even 0.
exists 0%Z; simpl in |- *; auto.
Qed.

Theorem Odd1 : Odd 1.
exists 0%Z; simpl in |- *; auto.
Qed.

Theorem OddOpp : forall z : Z, Odd z -> Odd (- z).
intros z H; case H; intros z1 H1; exists (- Z.succ z1)%Z; rewrite H1.
unfold Z.succ in |- *; ring.
Qed.

Theorem EvenOpp : forall z : Z, Even z -> Even (- z).
intros z H; case H; intros z1 H1; exists (- z1)%Z; rewrite H1; ring.
Qed.

Theorem OddEvenDec : forall n : Z, {Odd n} + {Even n}.
intros z; case z; simpl in |- *; auto with zarith.
right; apply EvenO.
intros p; case p; simpl in |- *; auto with zarith.
intros p1; left; exists (Zpos p1); rewrite Zplus_comm;
 simpl in |- *; auto.
intros p1; right; exists (Zpos p1); simpl in |- *; auto.
left; apply Odd1.
change
  (forall p : positive,
   {Odd (- Zpos p)} + {Even (- Zpos p)})
 in |- *.
intros p; case p; auto with zarith.
intros p1; left; apply OddOpp; exists (Zpos p1);
 rewrite Zplus_comm; simpl in |- *; auto.
intros p1; right; apply EvenOpp; exists (Zpos p1); simpl in |- *; auto.
left; apply OddOpp, Odd1.
Qed.

Theorem OddNEven : forall n : Z, Odd n -> ~ Even n.
intros n H1; red in |- *; intros H2; case H1; case H2; intros z1 Hz1 z2 Hz2.
absurd (n = n); auto.
pattern n at 1 in |- *; rewrite Hz1; rewrite Hz2;
 repeat rewrite (fun x => Zplus_comm x 1).
case z1; case z2; simpl in |- *;
 try (intros; red in |- *; intros; discriminate).
intros p p0; case p; simpl in |- *;
 try (intros; red in |- *; intros; discriminate).
Qed.

Theorem EvenNOdd : forall n : Z, Even n -> ~ Odd n.
intros n H1; red in |- *; intros H2; case H1; case H2; intros z1 Hz1 z2 Hz2.
absurd (n = n); auto.
pattern n at 1 in |- *; rewrite Hz1; rewrite Hz2;
 repeat rewrite (fun x => Zplus_comm x 1).
case z1; case z2; simpl in |- *;
 try (intros; red in |- *; intros; discriminate).
intros p p0; case p0; simpl in |- *;
 try (intros; red in |- *; intros; discriminate).
Qed.

Theorem EvenPlus1 : forall n m : Z, Even n -> Even m -> Even (n + m).
intros n m H H0; case H; case H0; intros z1 Hz1 z2 Hz2.
exists (z2 + z1)%Z; try rewrite Hz1; try rewrite Hz2; ring.
Qed.

Theorem OddPlus2 : forall n m : Z, Even n -> Odd m -> Odd (n + m).
intros n m H H0; case H; case H0; intros z1 Hz1 z2 Hz2.
exists (z2 + z1)%Z; try rewrite Hz1; try rewrite Hz2; ring.
Qed.

Theorem EvenMult1 : forall n m : Z, Even n -> Even (n * m).
intros n m H; case H; intros z1 Hz1; exists (z1 * m)%Z; rewrite Hz1; ring.
Qed.

Theorem EvenMult2 : forall n m : Z, Even m -> Even (n * m).
intros n m H; case H; intros z1 Hz1; exists (z1 * n)%Z; rewrite Hz1; ring.
Qed.

Theorem OddMult : forall n m : Z, Odd n -> Odd m -> Odd (n * m).
intros n m H1 H2; case H1; case H2; intros z1 Hz1 z2 Hz2;
 exists (2 * z1 * z2 + z1 + z2)%Z; rewrite Hz1; rewrite Hz2;
 ring.
Qed.

Theorem EvenMultInv : forall n m : Z, Even (n * m) -> Odd n -> Even m.
intros n m H H0; case (OddEvenDec m); auto; intros Z1.
contradict H; auto with zarith.
apply OddNEven, OddMult; easy.
Qed.

Theorem EvenExp :
 forall (n : Z) (m : nat), Even n -> Even (Zpower_nat n (S m)).
intros n m; elim m.
rewrite Zpower_nat_1; simpl in |- *; auto with zarith.
intros n0 H H0; replace (S (S n0)) with (1 + S n0); auto with arith.
rewrite Zpower_nat_is_exp; rewrite Zpower_nat_1; simpl in |- *;
 auto with zarith.
now apply EvenMult1.
Qed.

Theorem OddExp :
 forall (n : Z) (m : nat), Odd n -> Odd (Zpower_nat n m).
intros n m; elim m; simpl in |- *.
intros; apply Odd1.
intros n0 H H0; replace (S n0) with (1 + n0); auto with arith.
apply OddMult; try easy.
now apply H.
Qed.

Definition Feven (p : float) := Even (Fnum p).

Definition Fodd (p : float) := Odd (Fnum p).

Theorem FevenOrFodd : forall p : float, Feven p \/ Fodd p.
intros p; case (OddEvenDec (Fnum p)); auto.
Qed.

Theorem FevenSucProp :
 forall p : float,
 (Fodd p -> Feven (FSucc b radix precision p)) /\
 (Feven p -> Fodd (FSucc b radix precision p)).
intros p; unfold FSucc, Fodd, Feven in |- *.
generalize (Z_eq_bool_correct (Fnum p) (pPred (vNum b)));
 case (Z_eq_bool (Fnum p) (pPred (vNum b))); intros H'1.
rewrite H'1; simpl in |- *; auto.
unfold pPred in |- *; rewrite pGivesBound; unfold nNormMin in |- *.
case (OddEvenDec radix); auto with zarith.
intros H'; split; intros H'0; auto with zarith.
apply EvenMultInv with (n := radix); auto.
pattern radix at 1 in |- *; rewrite <- Zpower_nat_1;
 rewrite <- Zpower_nat_is_exp.
replace (1 + pred precision) with precision;
 [ idtac | inversion precisionGreaterThanOne; auto ].
rewrite (Zsucc_pred (Zpower_nat radix precision)); auto with zarith.
now apply OddSEven.
now apply OddExp.
intros H'; split; intros H'0; auto with zarith.
replace (pred precision) with (S (pred (pred precision))); auto with zarith.
now apply EvenExp.
contradict H'0; apply OddNEven.
replace (Z.pred (Zpower_nat radix precision)) with
 (Zpower_nat radix precision + - (1))%Z;
 [ idtac | unfold Z.pred in |- *; simpl in |- *; auto ].
replace precision with (S (pred precision));
 [ auto with zarith | inversion precisionGreaterThanOne; auto ].
apply OddPlus2.
now apply EvenExp.
apply OddOpp, Odd1.
generalize (Z_eq_bool_correct (Fnum p) (- nNormMin radix precision));
 case (Z_eq_bool (Fnum p) (- nNormMin radix precision));
 intros H'2.
generalize (Z_eq_bool_correct (Fexp p) (- dExp b));
 case (Z_eq_bool (Fexp p) (- dExp b)); intros H'3.
simpl in |- *; split.
apply OddSEven.
apply EvenSOdd.
simpl in |- *; auto with zarith.
rewrite H'2; unfold pPred, nNormMin in |- *; rewrite pGivesBound.
case (OddEvenDec radix); auto with zarith.
intros H'; split; intros H'0; auto with zarith.
apply EvenOpp; apply OddSEvenInv; rewrite <- Zsucc_pred; auto with zarith.
apply OddExp; easy.
contradict H'0; replace precision with (S (pred precision));
 [ auto with zarith | inversion precisionGreaterThanOne; auto ].
now apply OddNEven, OddOpp, OddExp.
intros H'; split; intros H'0; auto with zarith.
contradict H'0; replace (pred precision) with (S (pred (pred precision)));
 [ auto with zarith | auto with zarith ].
now apply EvenNOdd, EvenOpp, EvenExp.
replace precision with (S (pred precision));
 [ auto with zarith | inversion precisionGreaterThanOne; auto ].
apply OddOpp; apply EvenSOddInv; rewrite <- Zsucc_pred; auto with zarith.
apply EvenExp; easy.
simpl in |- *; split.
apply OddSEven.
apply EvenSOdd.
Qed.

Theorem FoddSuc :
 forall p : float, Fodd p -> Feven (FSucc b radix precision p).
intros p H'; case (FevenSucProp p); auto.
Qed.

Theorem FevenSuc :
 forall p : float, Feven p -> Fodd (FSucc b radix precision p).
intros p H'; case (FevenSucProp p); auto.
Qed.

Theorem FevenFop : forall p : float, Feven p -> Feven (Fopp p).
intros p; unfold Feven, Fopp in |- *; simpl in |- *; auto with zarith.
apply EvenOpp.
Qed.

Definition FNodd (p : float) := Fodd (Fnormalize radix b precision p).

Definition FNeven (p : float) := Feven (Fnormalize radix b precision p).

Theorem FNoddEq :
 forall f1 f2 : float,
 Fbounded b f1 -> Fbounded b f2 -> f1 = f2 :>R -> FNodd f1 -> FNodd f2.
intros f1 f2 H' H'0 H'1 H'2; red in |- *.
rewrite
 FcanonicUnique
                with
                (3 := pGivesBound)
               (p := Fnormalize radix b precision f2)
               (q := Fnormalize radix b precision f1);
 auto with zarith arith.
apply FnormalizeCanonic; auto with zarith.
apply FnormalizeCanonic; auto with zarith.
repeat rewrite FnormalizeCorrect; auto.
Qed.

Theorem FNevenEq :
 forall f1 f2 : float,
 Fbounded b f1 -> Fbounded b f2 -> f1 = f2 :>R -> FNeven f1 -> FNeven f2.
intros f1 f2 H' H'0 H'1 H'2; red in |- *.
rewrite
 FcanonicUnique
                with
                (3 := pGivesBound)
               (p := Fnormalize radix b precision f2)
               (q := Fnormalize radix b precision f1);
 auto with zarith arith.
apply FnormalizeCanonic; auto with zarith.
apply FnormalizeCanonic; auto with zarith.
repeat rewrite FnormalizeCorrect; auto.
Qed.

Theorem FNevenFop : forall p : float, FNeven p -> FNeven (Fopp p).
intros p; unfold FNeven in |- *.
rewrite Fnormalize_Fopp; auto with zarith arith.
intros; apply FevenFop; auto.
Qed.

Theorem FNoddSuc :
 forall p : float,
 Fbounded b p -> FNodd p -> FNeven (FNSucc b radix precision p).
unfold FNodd, FNeven, FNSucc in |- *.
intros p H' H'0.
rewrite FcanonicFnormalizeEq; auto with zarith.
apply FoddSuc; auto with zarith.
apply FSuccCanonic; auto with zarith.
apply FnormalizeCanonic; auto with zarith.
Qed.

Theorem FNevenSuc :
 forall p : float,
 Fbounded b p -> FNeven p -> FNodd (FNSucc b radix precision p).
unfold FNodd, FNeven, FNSucc in |- *.
intros p H' H'0.
rewrite FcanonicFnormalizeEq; auto with zarith arith.
apply FevenSuc; auto.
apply FSuccCanonic; auto with zarith.
apply FnormalizeCanonic; auto with zarith.
Qed.

Theorem FNevenOrFNodd : forall p : float, FNeven p \/ FNodd p.
intros p; unfold FNeven, FNodd in |- *; apply FevenOrFodd.
Qed.

Theorem FnOddNEven : forall n : float, FNodd n -> ~ FNeven n.
intros n H'; unfold FNeven, Feven in |- *; apply OddNEven; auto.
Qed.

End FOdd.

Section FRound.
Variable b : Fbound.
Variable radix : Z.
Variable precision : nat.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.

Hypothesis radixMoreThanOne : (1 < radix)%Z.
Hypothesis precisionGreaterThanOne : 1 < precision.
Hypothesis pGivesBound : Zpos (vNum b) = Zpower_nat radix precision.

Definition TotalP (P : R -> float -> Prop) :=
  forall r : R, exists p : float, P r p.

Definition UniqueP (P : R -> float -> Prop) :=
  forall (r : R) (p q : float), P r p -> P r q -> p = q :>R.

Definition CompatibleP (P : R -> float -> Prop) :=
  forall (r1 r2 : R) (p q : float),
  P r1 p -> r1 = r2 -> p = q :>R -> Fbounded b q -> P r2 q.

Definition MinOrMaxP (P : R -> float -> Prop) :=
  forall (r : R) (p : float), P r p -> isMin b radix r p \/ isMax b radix r p.

Definition RoundedModeP (P : R -> float -> Prop) :=
  TotalP P /\ CompatibleP P /\ MinOrMaxP P /\ MonotoneP radix P.

Theorem RoundedModeP_inv2 : forall P, RoundedModeP P -> CompatibleP P.
intros P H; Casec H; intros H H1; Casec H1; auto.
Qed.

Theorem RoundedModeP_inv4 : forall P, RoundedModeP P -> MonotoneP radix P.
intros P H; Casec H; intros H H1; Casec H1; intros H1 H2; Casec H2; auto.
Qed.

Theorem RoundedProjector : forall P, RoundedModeP P -> ProjectorP b radix P.
intros P H'; red in |- *; simpl in |- *.
intros p q H'0 H'1.
red in H'.
elim H'; intros H'2 H'3; elim H'3; intros H'4 H'5; elim H'5; intros H'6 H'7;
 case (H'6 p q); clear H'5 H'3 H'; auto.
intros H'; apply (ProjectMin b radix p); auto.
intros H'; apply (ProjectMax b radix p); auto.
Qed.

Theorem MinCompatible : CompatibleP (isMin b radix).
red in |- *.
intros r1 r2 p q H' H'0 H'1 H'2; split; auto.
rewrite <- H'0; unfold FtoRradix in H'1; rewrite <- H'1; case H'; auto.
Qed.

Theorem MinRoundedModeP : RoundedModeP (isMin b radix).
split; try red in |- *.
intros r; apply MinEx with (precision := precision); auto with zarith.
split; try exact MinCompatible.
split; try apply MonotoneMin; red in |- *; auto.
Qed.

Theorem MaxCompatible : CompatibleP (isMax b radix).
red in |- *.
intros r1 r2 p q H' H'0 H'1 H'2; split; auto.
rewrite <- H'0; unfold FtoRradix in H'1; rewrite <- H'1; case H'; auto.
Qed.

Theorem MaxRoundedModeP : RoundedModeP (isMax b radix).
split; try red in |- *.
intros r; apply MaxEx with (precision := precision); auto with zarith.
split; try exact MaxCompatible.
split; try apply MonotoneMax; red in |- *; auto.
Qed.

Theorem MinUniqueP : UniqueP (isMin b radix).
red in |- *.
intros r p q H' H'0.
unfold FtoRradix in |- *; apply MinEq with (1 := H'); auto.
Qed.

Theorem MaxUniqueP : UniqueP (isMax b radix).
red in |- *.
intros r p q H' H'0.
unfold FtoRradix in |- *; apply MaxEq with (1 := H'); auto.
Qed.

Theorem MinOrMaxRep :
 forall P,
 MinOrMaxP P ->
 forall p q : float, P p q -> exists m : Z, q = Float m (Fexp p) :>R.
intros P H' p q H'0; case (H' p q); auto; intros H'1.
apply FminRep with (3 := pGivesBound); auto with zarith.
apply FmaxRep with (3 := pGivesBound); auto with zarith.
Qed.

Theorem RoundedModeRep :
 forall P,
 RoundedModeP P ->
 forall p q : float, P p q -> exists m : Z, q = Float m (Fexp p) :>R.
intros P H' p q H'0.
apply MinOrMaxRep with (P := P); auto with zarith.
apply H'.
Qed.

Definition SymmetricP (P : R -> float -> Prop) :=
  forall (r : R) (p : float), P r p -> P (- r)%R (Fopp p).

End FRound.

Inductive Option (A : Set) : Set :=
  | Some : forall x : A, Option A
  | None : Option A.


Fixpoint Pdiv (p q : positive) {struct p} :
 Option positive * Option positive :=
  match p with
  | xH =>
      match q with
      | xH => (Some _ 1%positive, None _)
      | xO r => (None _, Some _ p)
      | xI r => (None _, Some _ p)
      end
  | xI p' =>
      match Pdiv p' q with
      | (None, None) =>
          match (1 - Zpos q)%Z with
          | Z0 => (Some _ 1%positive, None _)
          | Zpos r' => (Some _ 1%positive, Some _ r')
          | Zneg r' => (None _, Some _ 1%positive)
          end
      | (None, Some r1) =>
          match (Zpos (xI r1) - Zpos q)%Z with
          | Z0 => (Some _ 1%positive, None _)
          | Zpos r' => (Some _ 1%positive, Some _ r')
          | Zneg r' => (None _, Some _ (xI r1))
          end
      | (Some q1, None) =>
          match (1 - Zpos q)%Z with
          | Z0 => (Some _ (xI q1), None _)
          | Zpos r' => (Some _ (xI q1), Some _ r')
          | Zneg r' => (Some _ (xO q1), Some _ 1%positive)
          end
      | (Some q1, Some r1) =>
          match (Zpos (xI r1) - Zpos q)%Z with
          | Z0 => (Some _ (xI q1), None _)
          | Zpos r' => (Some _ (xI q1), Some _ r')
          | Zneg r' => (Some _ (xO q1), Some _ (xI r1))
          end
      end
  | xO p' =>
      match Pdiv p' q with
      | (None, None) => (None _, None _)
      | (None, Some r1) =>
          match (Zpos (xO r1) - Zpos q)%Z with
          | Z0 => (Some _ 1%positive, None _)
          | Zpos r' => (Some _ 1%positive, Some _ r')
          | Zneg r' => (None _, Some _ (xO r1))
          end
      | (Some q1, None) => (Some _ (xO q1), None _)
      | (Some q1, Some r1) =>
          match (Zpos (xO r1) - Zpos q)%Z with
          | Z0 => (Some _ (xI q1), None _)
          | Zpos r' => (Some _ (xI q1), Some _ r')
          | Zneg r' => (Some _ (xO q1), Some _ (xO r1))
          end
      end
  end.

Definition oZ h := match h with
                   | None => 0
                   | Some p => nat_of_P p
                   end.

Theorem Pdiv_correct :
 forall p q,
 nat_of_P p =
 oZ (fst (Pdiv p q)) * nat_of_P q + oZ (snd (Pdiv p q)) /\
 oZ (snd (Pdiv p q)) < nat_of_P q.
intros p q; elim p; simpl in |- *; auto with zarith arith.
3: case q; simpl in |- *; try intros q1; split; auto;
    unfold nat_of_P in |- *; simpl in |- *;
    auto with arith.
intros p'; simpl in |- *; case (Pdiv p' q); simpl in |- *;
 intros q1 r1 (H1, H2); split.
unfold nat_of_P in |- *; simpl in |- *.
rewrite ZL6; rewrite H1.
case q1; case r1; simpl in |- *.
intros r2 q2. rewrite Z.pos_sub_spec; unfold Pos.compare.
CaseEq (Pcompare (xI r2) q Datatypes.Eq);
 simpl in |- *; auto.
intros H3; rewrite <- (Pcompare_Eq_eq _ _ H3); simpl in |- *;
 unfold nat_of_P in |- *; simpl in |- *.
apply f_equal with (f := S); repeat rewrite (fun x y => mult_comm x (S y));
 repeat rewrite ZL6; unfold nat_of_P in |- *;
 simpl in |- *; ring.
intros H3; unfold nat_of_P in |- *; simpl in |- *;
 repeat rewrite ZL6; unfold nat_of_P in |- *;
 repeat rewrite (fun x y => plus_comm x (S y)); simpl in |- *;
 apply f_equal with (f := S); ring.
intros H3; case (Pminus_mask_Gt _ _ H3); intros h (H4, (H5, H6));
 unfold Pminus in |- *; rewrite H4.
apply
 trans_equal
  with
    (nat_of_P q + nat_of_P h +
     Pmult_nat q2 2 * Pmult_nat q 1);
 [ rewrite <- nat_of_P_plus_morphism; rewrite H5; simpl in |- *;
    repeat rewrite ZL6; unfold nat_of_P in |- *;
    apply f_equal with (f := S)
 | unfold nat_of_P in |- * ]; ring.
intros r2. rewrite Z.pos_sub_spec; unfold Pos.compare.
CaseEq (Pcompare 1 q Datatypes.Eq); simpl in |- *; auto.
intros H3; rewrite <- (Pcompare_Eq_eq _ _ H3); simpl in |- *;
 repeat rewrite ZL6; unfold nat_of_P in |- *; simpl in |- *;
 apply f_equal with (f := S);ring.
intros H3; unfold nat_of_P in |- *; simpl in |- *;
 repeat rewrite (fun x y => plus_comm x (S y));
 simpl in |- *; apply f_equal with (f := S); repeat rewrite ZL6;
   unfold nat_of_P in |- *; simpl in |- *; ring.
intros H3; case (Pminus_mask_Gt _ _ H3); intros h (H4, (H5, H6));
 unfold Pminus in |- *; rewrite H4;
 apply
  trans_equal
   with
     (nat_of_P q + nat_of_P h +
      Pmult_nat r2 2 * Pmult_nat q 1);
 [ rewrite <- nat_of_P_plus_morphism; rewrite H5; simpl in |- *;
    repeat rewrite ZL6; unfold nat_of_P in |- *;
    apply f_equal with (f := S)
 | unfold nat_of_P in |- * ]; ring.
intros r2. rewrite Z.pos_sub_spec; unfold Pos.compare.
 CaseEq (Pcompare (xI r2) q Datatypes.Eq); simpl in |- *;
 auto.
intros H3; rewrite <- (Pcompare_Eq_eq _ _ H3); simpl in |- *;
 unfold nat_of_P in |- *; simpl in |- *; repeat rewrite ZL6;
 unfold nat_of_P in |- *; apply f_equal with (f := S);
 ring.
intros H3; unfold nat_of_P in |- *; simpl in |- *;
 repeat rewrite ZL6; unfold nat_of_P in |- *;
 apply f_equal with (f := S); ring.
intros H3; case (Pminus_mask_Gt _ _ H3); intros h (H4, (H5, H6));
 unfold Pminus in |- *; rewrite H4;
 apply trans_equal with (nat_of_P q + nat_of_P h);
 [ rewrite <- (nat_of_P_plus_morphism q); rewrite H5;
    unfold nat_of_P in |- *; simpl in |- *;
    repeat rewrite ZL6; unfold nat_of_P in |- *;
    apply f_equal with (f := S)
 | unfold nat_of_P in |- * ]; ring.
case q; simpl in |- *; auto.
generalize H2; case q1; case r1; simpl in |- *; auto.
intros r2 q2. rewrite Z.pos_sub_spec; unfold Pos.compare.
CaseEq (Pcompare (xI r2) q Datatypes.Eq);
 simpl in |- *; auto.
intros; apply lt_O_nat_of_P; auto.
intros H H0; apply nat_of_P_lt_Lt_compare_morphism; auto.
intros H3 H7; case (Pminus_mask_Gt _ _ H3); intros h (H4, (H5, H6));
 unfold Pminus in |- *; rewrite H4;
 apply plus_lt_reg_l with (p := nat_of_P q);
 rewrite <- (nat_of_P_plus_morphism q); rewrite H5;
 unfold nat_of_P in |- *; simpl in |- *; repeat rewrite ZL6;
 unfold nat_of_P in |- *;
 apply le_lt_trans with (Pmult_nat r2 1 + Pmult_nat q 1);
 auto with arith.
intros r2 HH; case q; simpl in |- *; auto.
intros p2; apply nat_of_P_lt_Lt_compare_morphism; easy.
intros p2; apply nat_of_P_lt_Lt_compare_morphism; easy.
intros r2 HH. rewrite Z.pos_sub_spec; unfold Pos.compare.
CaseEq (Pcompare (xI r2) q Datatypes.Eq);
 simpl in |- *.
intros; apply lt_O_nat_of_P; auto.
intros H3; apply nat_of_P_lt_Lt_compare_morphism; auto.
intros H3; case (Pminus_mask_Gt _ _ H3); intros h (H4, (H5, H6));
 unfold Pminus in |- *; rewrite H4;
 apply plus_lt_reg_l with (p := nat_of_P q);
 rewrite <- (nat_of_P_plus_morphism q); rewrite H5;
 unfold nat_of_P in |- *; simpl in |- *; repeat rewrite ZL6;
 unfold nat_of_P in |- *;
 apply le_lt_trans with (Pmult_nat r2 1 + Pmult_nat q 1);
 auto with arith.
intros HH; case q; simpl in |- *; auto.
intros p2; apply nat_of_P_lt_Lt_compare_morphism; easy.
intros p2; apply nat_of_P_lt_Lt_compare_morphism; easy.
intros p'; simpl in |- *; case (Pdiv p' q); simpl in |- *;
 intros q1 r1 (H1, H2); split.
unfold nat_of_P in |- *; simpl in |- *; rewrite ZL6; rewrite H1.
case q1; case r1; simpl in |- *; auto.
intros r2 q2. rewrite Z.pos_sub_spec; unfold Pos.compare.
CaseEq (Pcompare (xO r2) q Datatypes.Eq);
 simpl in |- *; auto.
intros H3; rewrite <- (Pcompare_Eq_eq _ _ H3); simpl in |- *;
 unfold nat_of_P in |- *; simpl in |- *; repeat rewrite ZL6;
 unfold nat_of_P in |- *; ring.
intros H3; unfold nat_of_P in |- *; simpl in |- *;
 repeat rewrite ZL6; unfold nat_of_P in |- *;
 ring.
intros H3; case (Pminus_mask_Gt _ _ H3); intros h (H4, (H5, H6));
 unfold Pminus in |- *; rewrite H4;
 apply
  trans_equal
   with
     (nat_of_P q + nat_of_P h +
      Pmult_nat q2 2 * Pmult_nat q 1);
 [ rewrite <- (nat_of_P_plus_morphism q); rewrite H5;
    unfold nat_of_P in |- *; simpl in |- *;
    repeat rewrite ZL6; unfold nat_of_P in |- *
 | unfold nat_of_P in |- * ]; ring.
intros H3; unfold nat_of_P in |- *; simpl in |- *;
 repeat rewrite ZL6; unfold nat_of_P in |- *;
 ring.
intros r2. rewrite Z.pos_sub_spec; unfold Pos.compare.
CaseEq (Pcompare (xO r2) q Datatypes.Eq); simpl in |- *;
 auto.
intros H3; rewrite <- (Pcompare_Eq_eq _ _ H3); simpl in |- *;
 unfold nat_of_P in |- *; simpl in |- *; repeat rewrite ZL6;
 unfold nat_of_P in |- *; ring.
intros H3; unfold nat_of_P in |- *; simpl in |- *;
 repeat rewrite ZL6; unfold nat_of_P in |- *;
 ring.
intros H3; case (Pminus_mask_Gt _ _ H3); intros h (H4, (H5, H6));
 unfold Pminus in |- *; rewrite H4;
 apply trans_equal with (nat_of_P q + nat_of_P h);
 [ rewrite <- (nat_of_P_plus_morphism q); rewrite H5;
    unfold nat_of_P in |- *; simpl in |- *;
    repeat rewrite ZL6; unfold nat_of_P in |- *
 | unfold nat_of_P in |- * ]; ring.
generalize H2; case q1; case r1; simpl in |- *.
intros r2 q2. rewrite Z.pos_sub_spec; unfold Pos.compare.
CaseEq (Pcompare (xO r2) q Datatypes.Eq);
 simpl in |- *; auto.
intros; apply lt_O_nat_of_P; auto.
intros H H0; apply nat_of_P_lt_Lt_compare_morphism; auto.
intros H3 H7; case (Pminus_mask_Gt _ _ H3); intros h (H4, (H5, H6));
 unfold Pminus in |- *; rewrite H4;
 apply plus_lt_reg_l with (p := nat_of_P q);
 rewrite <- (nat_of_P_plus_morphism q); rewrite H5;
 unfold nat_of_P in |- *; simpl in |- *;
 unfold nat_of_P in |- *; simpl in |- *; repeat rewrite ZL6;
 unfold nat_of_P in |- *;
 apply lt_trans with (Pmult_nat r2 1 + Pmult_nat q 1);
 auto with arith.
intros; apply lt_O_nat_of_P; auto.
intros r2 HH. rewrite Z.pos_sub_spec; unfold Pos.compare.
CaseEq (Pcompare (xO r2) q Datatypes.Eq);
 simpl in |- *.
intros; apply lt_O_nat_of_P; auto.
intros H3; apply nat_of_P_lt_Lt_compare_morphism; auto.
intros H3; case (Pminus_mask_Gt _ _ H3); intros h (H4, (H5, H6));
 unfold Pminus in |- *; rewrite H4;
 apply plus_lt_reg_l with (p := nat_of_P q);
 rewrite <- (nat_of_P_plus_morphism q); rewrite H5;
 unfold nat_of_P in |- *; simpl in |- *; repeat rewrite ZL6;
 unfold nat_of_P in |- *;
 apply lt_trans with (Pmult_nat r2 1 + Pmult_nat q 1);
 auto with arith.
auto.
rewrite ZL6; generalize (Pos2Nat.is_pos q1); auto with zarith arith.
rewrite ZL6; generalize (Pos2Nat.is_pos q1); auto with zarith arith.
Qed.

Transparent Pdiv.

Definition oZ1 (x : Option positive) :=
  match x with
  | None => 0%Z
  | Some z => Zpos z
  end.

Definition Zquotient (n m : Z) :=
  match n, m with
  | Z0, _ => 0%Z
  | _, Z0 => 0%Z
  | Zpos x, Zpos y => match Pdiv x y with
                      | (x, _) => oZ1 x
                      end
  | Zneg x, Zneg y => match Pdiv x y with
                      | (x, _) => oZ1 x
                      end
  | Zpos x, Zneg y => match Pdiv x y with
                      | (x, _) => (- oZ1 x)%Z
                      end
  | Zneg x, Zpos y => match Pdiv x y with
                      | (x, _) => (- oZ1 x)%Z
                      end
  end.

Theorem inj_oZ1 : forall z, oZ1 z = Z_of_nat (oZ z).
intros z; case z; simpl in |- *;
 try (intros; apply sym_equal; apply inject_nat_convert; auto);
 auto.
Qed.

Theorem ZquotientProp :
 forall m n : Z,
 n <> 0%Z ->
 ex
   (fun r : Z =>
    m = (Zquotient m n * n + r)%Z /\
    (Z.abs (Zquotient m n * n) <= Z.abs m)%Z /\ (Z.abs r < Z.abs n)%Z).
intros m n; unfold Zquotient in |- *; case n; simpl in |- *.
intros H; case H; auto.
intros n' Hn'; case m; simpl in |- *; auto.
exists 0%Z; repeat split; simpl in |- *; auto with zarith.
intros m'; generalize (Pdiv_correct m' n'); case (Pdiv m' n'); simpl in |- *;
 auto.
intros q r (H1, H2); exists (oZ1 r); repeat (split; auto with zarith).
rewrite <- (inject_nat_convert (Zpos m') m'); auto.
rewrite H1.
rewrite inj_plus; rewrite inj_mult.
rewrite <- (inject_nat_convert (Zpos n') n'); auto.
repeat rewrite inj_oZ1; auto.
rewrite inj_oZ1; rewrite Z.abs_eq; auto with zarith.
rewrite <- (inject_nat_convert (Zpos n') n'); auto with zarith.
rewrite inj_oZ1; rewrite Z.abs_eq; auto with zarith.
rewrite <- (inject_nat_convert (Zpos n') n'); auto with zarith.
intros m'; generalize (Pdiv_correct m' n'); case (Pdiv m' n'); simpl in |- *;
 auto.
intros q r (H1, H2); exists (- oZ1 r)%Z; repeat (split; auto with zarith).
replace (Zneg m') with (- Zpos m')%Z; [ idtac | simpl in |- *; auto ].
rewrite <- (inject_nat_convert (Zpos m') m'); auto.
rewrite H1.
rewrite inj_plus; rewrite inj_mult.
repeat rewrite inj_oZ1; auto with zarith.
rewrite <- Zopp_mult_distr_l; rewrite Zabs_Zopp.
rewrite inj_oZ1; rewrite Z.abs_eq; auto with zarith.
rewrite Zabs_Zopp.
rewrite inj_oZ1; rewrite Z.abs_eq; auto with zarith.
intros n' Hn'; case m; simpl in |- *; auto.
exists 0%Z; repeat split; simpl in |- *; auto with zarith.
intros m'; generalize (Pdiv_correct m' n'); case (Pdiv m' n'); simpl in |- *;
 auto.
intros q r (H1, H2); exists (oZ1 r); repeat (split; auto with zarith).
replace (Zneg n') with (- Zpos n')%Z; [ idtac | simpl in |- *; auto ].
rewrite <- (inject_nat_convert (Zpos m') m'); auto.
rewrite H1.
rewrite inj_plus; rewrite inj_mult.
rewrite <- (inject_nat_convert (Zpos n') n'); auto.
repeat rewrite inj_oZ1; auto with zarith.
replace (Zneg n') with (- Zpos n')%Z; [ idtac | simpl in |- *; auto ].
rewrite Zmult_opp_opp.
rewrite inj_oZ1; rewrite Z.abs_eq; auto with zarith.
rewrite <- (inject_nat_convert (Zpos n') n'); auto with zarith.
rewrite inj_oZ1; rewrite Z.abs_eq; auto with zarith.
rewrite <- (inject_nat_convert (Zpos n') n'); auto with zarith.
intros m'; generalize (Pdiv_correct m' n'); case (Pdiv m' n'); simpl in |- *;
 auto.
intros q r (H1, H2); exists (- oZ1 r)%Z; repeat (split; auto with zarith).
replace (Zneg m') with (- Zpos m')%Z; [ idtac | simpl in |- *; auto ].
rewrite <- (inject_nat_convert (Zpos m') m'); auto.
replace (Zneg n') with (- Zpos n')%Z; [ idtac | simpl in |- *; auto ].
rewrite H1.
rewrite inj_plus; rewrite inj_mult.
rewrite <- (inject_nat_convert (Zpos n') n'); auto.
repeat rewrite inj_oZ1; auto with zarith.
replace (Zneg n') with (- Zpos n')%Z; [ idtac | simpl in |- *; auto ].
rewrite <- Zopp_mult_distr_r; rewrite Zabs_Zopp.
rewrite inj_oZ1; rewrite Z.abs_eq; auto with zarith.
rewrite Zabs_Zopp.
rewrite inj_oZ1; rewrite Z.abs_eq; auto with zarith.
Qed.

Theorem ZquotientPos :
 forall z1 z2 : Z, (0 <= z1)%Z -> (0 <= z2)%Z -> (0 <= Zquotient z1 z2)%Z.
intros z1 z2 H H0; case (Z.eq_dec z2 0); intros Z1.
rewrite Z1; red in |- *; case z1; simpl in |- *; auto; intros; red in |- *;
 intros; discriminate.
case (ZquotientProp z1 z2); auto; intros r (H1, (H2, H3)).
case (Zle_or_lt 0 (Zquotient z1 z2)); auto; intros Z2.
contradict H3; apply Zle_not_lt.
replace r with (z1 - Zquotient z1 z2 * z2)%Z;
 [ idtac | pattern z1 at 1 in |- *; rewrite H1; ring ].
repeat rewrite Z.abs_eq; auto.
pattern z2 at 1 in |- *; replace z2 with (0 + 1 * z2)%Z; [ idtac | ring ].
unfold Zminus in |- *; apply Z.le_trans with (z1 + 1 * z2)%Z; auto with zarith.
apply Zplus_le_compat_l.
rewrite Zopp_mult_distr_l.
apply Zle_Zmult_comp_r; auto with zarith.
unfold Zminus in |- *; rewrite Zopp_mult_distr_l; auto with zarith.
Qed.

Definition Zdivides (n m : Z) := exists q : Z, n = (m * q)%Z.

Theorem ZdividesZquotient :
 forall n m : Z, m <> 0%Z -> Zdivides n m -> n = (Zquotient n m * m)%Z.
intros n m H' H'0.
case H'0; intros z1 Hz1.
case (ZquotientProp n m); auto; intros z2 (Hz2, (Hz3, Hz4)).
cut (z2 = 0%Z);
 [ intros H1; pattern n at 1 in |- *; rewrite Hz2; rewrite H1; ring | idtac ].
cut (z2 = ((z1 - Zquotient n m) * m)%Z); [ intros H2 | idtac ].
case (Z.eq_dec (z1 - Zquotient n m) 0); intros H3.
rewrite H2; rewrite H3; ring.
contradict Hz4.
replace (Z.abs m) with (1 * Z.abs m)%Z; [ idtac | ring ].
apply Zle_not_lt; rewrite H2.
rewrite Zabs_Zmult; apply Zle_Zmult_comp_r; auto with zarith.
rewrite Zmult_minus_distr_r; rewrite (Zmult_comm z1); rewrite <- Hz1;
 (pattern n at 1 in |- *; rewrite Hz2); ring.
Qed.

Theorem ZdividesZquotientInv :
 forall n m : Z, n = (Zquotient n m * m)%Z -> Zdivides n m.
intros n m H'; red in |- *.
exists (Zquotient n m); auto.
pattern n at 1 in |- *; rewrite H'; auto with zarith.
Qed.

Theorem ZdividesMult :
 forall n m p : Z, Zdivides n m -> Zdivides (p * n) (p * m).
intros n m p H'; red in H'.
elim H'; intros q E.
red in |- *.
exists q.
rewrite E.
auto with zarith.
Qed.

Theorem Zeq_mult_simpl :
 forall a b c : Z, c <> 0%Z -> (a * c)%Z = (b * c)%Z -> a = b.
intros a b c H H0.
case (Zle_or_lt c 0); intros Zl1.
apply Zle_antisym; apply Zmult_le_reg_r with (p := (- c)%Z); try apply Z.lt_gt;
 auto with zarith; repeat rewrite <- Zopp_mult_distr_r;
 rewrite H0; auto with zarith.
apply Zle_antisym; apply Zmult_le_reg_r with (p := c); try apply Z.lt_gt;
 auto with zarith; rewrite H0; auto with zarith.
Qed.

Theorem ZdividesDiv :
 forall n m p : Z, p <> 0%Z -> Zdivides (p * n) (p * m) -> Zdivides n m.
intros n m p H' H'0.
case H'0; intros q E.
exists q.
apply Zeq_mult_simpl with (c := p); auto.
rewrite (Zmult_comm n); rewrite E; ring.
Qed.

Definition ZdividesP : forall n m : Z, {Zdivides n m} + {~ Zdivides n m}.
intros n m; case m.
case n.
left; red in |- *; exists 0%Z; auto with zarith.
intros p; right; red in |- *; intros H; case H; simpl in |- *; intros f H1;
 discriminate.
intros p; right; red in |- *; intros H; case H; simpl in |- *; intros f H1;
 discriminate.
intros p; generalize (Z_eq_bool_correct (Zquotient n (Zpos p) * Zpos p) n);
 case (Z_eq_bool (Zquotient n (Zpos p) * Zpos p) n);
 intros H1.
left; apply ZdividesZquotientInv; auto.
right; contradict H1; apply sym_equal; apply ZdividesZquotient; auto.
red in |- *; intros; discriminate.
intros p; generalize (Z_eq_bool_correct (Zquotient n (Zneg p) * Zneg p) n);
 case (Z_eq_bool (Zquotient n (Zneg p) * Zneg p) n);
 intros H1.
left; apply ZdividesZquotientInv; auto.
right; contradict H1; apply sym_equal; apply ZdividesZquotient; auto.
red in |- *; intros; discriminate.
Defined.

Theorem Zdivides1 : forall m : Z, Zdivides m 1.
intros m; exists m; auto with zarith.
Qed.

Theorem Zabs_eq_case :
 forall z1 z2 : Z, Z.abs z1 = Z.abs z2 -> z1 = z2 \/ z1 = (- z2)%Z.
intros z1 z2; case z1; case z2; simpl in |- *; auto;
 try (intros; discriminate); intros p1 p2 H1; injection H1;
 (intros H2; rewrite H2); auto.
Qed.

Theorem NotDividesDigit :
 forall r v : Z,
 (1 < r)%Z -> v <> 0%Z -> ~ Zdivides v (Zpower_nat r (digit r v)).
intros r v H H'; red in |- *; intros H'0; case H'0; intros q E.
absurd (Z.abs v < Zpower_nat r (digit r v))%Z; auto with zarith.
apply Zle_not_lt.
case (Z.eq_dec q 0); intros Z1.
case H'; rewrite E; rewrite Z1; ring.
pattern v at 2 in |- *; rewrite E.
rewrite Zabs_Zmult.
pattern (Zpower_nat r (digit r v)) at 1 in |- *;
 replace (Zpower_nat r (digit r v)) with (Zpower_nat r (digit r v) * 1)%Z;
 [ idtac | ring ].
rewrite (fun x y => Z.abs_eq (Zpower_nat x y)); auto with zarith.
apply Zle_Zmult_comp_l; auto with zarith.
apply Zpower_NR0; omega.
apply Zpower_NR0; omega.
now apply digitMore.
Qed.

Theorem ZDividesLe :
 forall n m : Z, n <> 0%Z -> Zdivides n m -> (Z.abs m <= Z.abs n)%Z.
intros n m H' H'0; case H'0; intros q E; rewrite E.
rewrite Zabs_Zmult.
pattern (Z.abs m) at 1 in |- *; replace (Z.abs m) with (Z.abs m * 1)%Z;
 [ idtac | ring ].
apply Zle_Zmult_comp_l; auto with zarith.
generalize E H'; case q; simpl in |- *; auto;
 try (intros H1 H2; case H2; rewrite H1; ring; fail);
 intros p; case p; unfold Z.le in |- *; simpl in |- *;
 intros; red in |- *; discriminate.
Qed.


Section mf.
Variable radix : Z.
Hypothesis radixMoreThanOne : (1 < radix)%Z.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ radixMoreThanOne).

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.

Fixpoint maxDiv (v : Z) (p : nat) {struct p} : nat :=
  match p with
  | O => 0
  | S p' =>
      match ZdividesP v (Zpower_nat radix p) with
      | left _ => p
      | right _ => maxDiv v p'
      end
  end.

Theorem maxDivLess : forall (v : Z) (p : nat), maxDiv v p <= p.
intros v p; elim p; simpl in |- *; auto.
intros n H'; case (ZdividesP v (radix * Zpower_nat radix n)); auto.
Qed.

Theorem maxDivLt :
 forall (v : Z) (p : nat),
 ~ Zdivides v (Zpower_nat radix p) -> maxDiv v p < p.
intros v p; case p; simpl in |- *; auto.
intros H'; case H'.
apply Zdivides1.
intros n H'; case (ZdividesP v (radix * Zpower_nat radix n)); auto.
intros H'0; case H'; auto.
intros H'0; generalize (maxDivLess v n); auto with arith.
Qed.

Theorem maxDivCorrect :
 forall (v : Z) (p : nat), Zdivides v (Zpower_nat radix (maxDiv v p)).
intros v p; elim p.
unfold maxDiv in |- *; rewrite Zpower_nat_O; auto.
apply Zdivides1.
simpl in |- *.
intros n H'; case (ZdividesP v (radix * Zpower_nat radix n)); simpl in |- *;
 auto with zarith.
Qed.

Theorem maxDivSimplAux :
 forall (v : Z) (p q : nat),
 p = maxDiv v (S (q + p)) -> p = maxDiv v (S p).
intros v p q; elim q.
simpl in |- *; case (ZdividesP v (radix * Zpower_nat radix p)); auto.
intros n H' H'0.
apply H'; auto; clear H'.
simpl in H'0; generalize H'0; clear H'0.
case (ZdividesP v (radix * (radix * Zpower_nat radix (n + p)))).
2: simpl in |- *; auto.
intros H' H'0; contradict H'0; auto with zarith.
Qed.

Theorem maxDivSimpl :
 forall (v : Z) (p q : nat),
 p < q -> p = maxDiv v q -> p = maxDiv v (S p).
intros v p q H' H'0.
apply maxDivSimplAux with (q := q - S p); auto.
replace (S (q - S p + p)) with q; auto with zarith.
Qed.

Theorem maxDivSimplInvAux :
 forall (v : Z) (p q : nat),
 p = maxDiv v (S p) -> p = maxDiv v (S (q + p)).
intros v p q H'; elim q.
simpl in |- *; auto.
intros n; simpl in |- *.
case (ZdividesP v (radix * Zpower_nat radix (n + p))); auto.
case (ZdividesP v (radix * (radix * Zpower_nat radix (n + p)))); auto.
intros H'0 H'1 H'2; contradict H'2; auto with zarith.
case (ZdividesP v (radix * (radix * Zpower_nat radix (n + p)))); auto.
intros H'0 H'1 H'2; case H'1.
case H'0; intros z1 Hz1; exists (radix * z1)%Z;rewrite Hz1.
unfold Zpower_nat; simpl; ring.
Qed.

Theorem maxDivSimplInv :
 forall (v : Z) (p q : nat),
 p < q -> p = maxDiv v (S p) -> p = maxDiv v q.
intros v p q H' H'0.
replace q with (S (q - S p + p)); auto with zarith.
apply maxDivSimplInvAux; auto.
Qed.

Theorem maxDivUnique :
 forall (v : Z) (p : nat),
 p = maxDiv v (S p) ->
 Zdivides v (Zpower_nat radix p) /\ ~ Zdivides v (Zpower_nat radix (S p)).
intros v p H'; split.
rewrite H'.
apply maxDivCorrect; auto.
red in |- *; intros H'0; generalize H'; clear H'.
simpl in |- *.
case (ZdividesP v (radix * Zpower_nat radix p)); simpl in |- *; auto.
intros H' H'1; contradict H'1; auto with zarith.
Qed.

Theorem maxDivUniqueDigit :
 forall v : Z,
 v <> 0 ->
 Zdivides v (Zpower_nat radix (maxDiv v (digit radix v))) /\
 ~ Zdivides v (Zpower_nat radix (S (maxDiv v (digit radix v)))).
intros v H'.
apply maxDivUnique; auto.
apply maxDivSimpl with (q := digit radix v); auto.
apply maxDivLt; auto.
apply NotDividesDigit; auto.
Qed.

Theorem maxDivUniqueInverse :
 forall (v : Z) (p : nat),
 Zdivides v (Zpower_nat radix p) ->
 ~ Zdivides v (Zpower_nat radix (S p)) -> p = maxDiv v (S p).
intros v p H' H'0; simpl in |- *.
case (ZdividesP v (radix * Zpower_nat radix p)); auto.
intros H'1; case H'0; simpl in |- *; auto.
intros H'1.
generalize H'; case p; simpl in |- *; auto.
intros n H'2; case (ZdividesP v (radix * Zpower_nat radix n)); auto.
intros H'3; case H'3; auto.
Qed.

Theorem maxDivUniqueInverseDigit :
 forall (v : Z) (p : nat),
 v <> 0 ->
 Zdivides v (Zpower_nat radix p) ->
 ~ Zdivides v (Zpower_nat radix (S p)) -> p = maxDiv v (digit radix v).
intros v p H' H'0 H'1.
apply maxDivSimplInv; auto.
2: apply maxDivUniqueInverse; auto.
apply Zpower_nat_anti_monotone_lt with (n := radix); auto.
apply Z.le_lt_trans with (m := Z.abs v); auto.
rewrite <- (fun x => Z.abs_eq (Zpower_nat radix x));
 try apply ZDividesLe; auto.
apply Zpower_NR0; omega.
apply digitMore; auto.
Qed.

Theorem maxDivPlus :
 forall (v : Z) (n : nat),
 v <> 0 ->
 maxDiv (v * Zpower_nat radix n) (digit radix v + n) =
 maxDiv v (digit radix v) + n.
intros v n H.
replace (digit radix v + n) with (digit radix (v * Zpower_nat radix n)); auto.
apply sym_equal.
apply maxDivUniqueInverseDigit; auto.
red in |- *; intros Z1; case (Zmult_integral _ _ Z1); intros Z2.
case H; auto.
absurd (0 < Zpower_nat radix n)%Z; auto with zarith.
now apply Zpower_nat_less.
rewrite Zpower_nat_is_exp.
repeat rewrite (fun x : Z => Zmult_comm x (Zpower_nat radix n)).
apply ZdividesMult; auto.
case (maxDivUniqueDigit v); auto.
replace (S (maxDiv v (digit radix v) + n)) with
 (S (maxDiv v (digit radix v)) + n); auto.
rewrite Zpower_nat_is_exp.
repeat rewrite (fun x : Z => Zmult_comm x (Zpower_nat radix n)).
red in |- *; intros H'.
absurd (Zdivides v (Zpower_nat radix (S (maxDiv v (digit radix v))))).
case (maxDivUniqueDigit v); auto.
apply ZdividesDiv with (p := Zpower_nat radix n); auto with zarith.
apply sym_not_eq, Zlt_not_eq, Zpower_nat_less; easy.
apply digitAdd; auto with zarith.
Qed.

Definition LSB (x : float) :=
  (Z_of_nat (maxDiv (Fnum x) (Fdigit radix x)) + Fexp x)%Z.

Theorem LSB_shift :
 forall (x : float) (n : nat), ~ is_Fzero x -> LSB x = LSB (Fshift radix n x).
intros x n H'; unfold LSB, Fdigit in |- *; simpl in |- *.
rewrite digitAdd; auto with arith.
rewrite maxDivPlus; auto.
rewrite inj_plus; ring.
Qed.

Theorem LSB_comp :
 forall (x y : float) (n : nat), ~ is_Fzero x -> x = y :>R -> LSB x = LSB y.
intros x y H' H'0 H'1.
case (FshiftCorrectSym radix) with (2 := H'1); auto.
intros m1 H'2; elim H'2; intros m2 E; clear H'2.
rewrite (LSB_shift x m1); auto.
rewrite E; auto.
apply sym_equal; apply LSB_shift; auto.
apply (NisFzeroComp radix) with (x := x); auto.
Qed.

Theorem maxDiv_opp :
 forall (v : Z) (p : nat), maxDiv v p = maxDiv (- v) p.
intros v p; elim p; simpl in |- *; auto.
intros n H; case (ZdividesP v (radix * Zpower_nat radix n));
 case (ZdividesP (- v) (radix * Zpower_nat radix n)); auto.
intros Z1 Z2; case Z1.
case Z2; intros z1 Hz1; exists (- z1)%Z; rewrite Hz1; ring.
intros Z1 Z2; case Z2.
case Z1; intros z1 Hz1; exists (- z1)%Z.
rewrite <- (Z.opp_involutive v); rewrite Hz1; ring.
Qed.

Theorem LSB_opp : forall x : float, LSB x = LSB (Fopp x).
intros x; unfold LSB in |- *; simpl in |- *.
rewrite Fdigit_opp; auto.
rewrite maxDiv_opp; auto.
Qed.

Theorem maxDiv_abs :
 forall (v : Z) (p : nat), maxDiv v p = maxDiv (Z.abs v) p.
intros v p; elim p; simpl in |- *; auto.
intros n H; case (ZdividesP v (radix * Zpower_nat radix n));
 case (ZdividesP (Z.abs v) (radix * Zpower_nat radix n));
 auto.
intros Z1 Z2; case Z1.
case Z2; intros z1 Hz1; exists (Z.abs z1); rewrite Hz1.
rewrite Zabs_Zmult; f_equal. apply Z.abs_eq.
apply Z.mul_nonneg_nonneg; try auto with zarith.
apply Zpower_NR0; auto with zarith.
intros Z1 Z2; case Z2.
case Z1; intros z1 Hz1.
case (Zle_or_lt v 0); intros Z4.
exists (- z1)%Z; rewrite <- (Z.opp_involutive v);
 rewrite <- (Zabs_eq_opp v); auto; rewrite Hz1; ring.
exists z1; rewrite <- (Z.abs_eq v); auto with zarith; rewrite Hz1; ring.
Qed.

Theorem LSB_abs : forall x : float, LSB x = LSB (Fabs x).
intros x; unfold LSB in |- *; simpl in |- *.
rewrite Fdigit_abs; auto.
rewrite maxDiv_abs; auto.
Qed.

Definition MSB (x : float) := Z.pred (Z_of_nat (Fdigit radix x) + Fexp x).

Theorem MSB_shift :
 forall (x : float) (n : nat), ~ is_Fzero x -> MSB x = MSB (Fshift radix n x).
intros; unfold MSB, Fshift, Fdigit in |- *; simpl in |- *.
rewrite digitAdd; auto with zarith.
Qed.

Theorem MSB_comp :
 forall (x y : float) (n : nat), ~ is_Fzero x -> x = y :>R -> MSB x = MSB y.
intros x y H' H'0 H'1.
case (FshiftCorrectSym radix) with (2 := H'1); auto.
intros m1 H'2; elim H'2; intros m2 E; clear H'2.
rewrite (MSB_shift x m1); auto.
rewrite E; auto.
apply sym_equal; apply MSB_shift; auto.
apply (NisFzeroComp radix) with (x := x); auto.
Qed.

Theorem MSB_opp : forall x : float, MSB x = MSB (Fopp x).
intros x; unfold MSB in |- *; simpl in |- *.
rewrite Fdigit_opp; auto.
Qed.

Theorem MSB_abs : forall x : float, MSB x = MSB (Fabs x).
intros x; unfold MSB in |- *; simpl in |- *.
rewrite Fdigit_abs; auto.
Qed.

Theorem LSB_le_MSB : forall x : float, ~ is_Fzero x -> (LSB x <= MSB x)%Z.
intros x H'; unfold LSB, MSB in |- *.
apply Zle_Zpred.
cut (maxDiv (Fnum x) (Fdigit radix x) < Fdigit radix x); auto with zarith.
apply maxDivLt; auto.
unfold Fdigit in |- *; apply NotDividesDigit; auto.
Qed.

Theorem Fexp_le_LSB : forall x : float, (Fexp x <= LSB x)%Z.
intros x; unfold LSB in |- *.
auto with zarith.
Qed.

Theorem Ulp_Le_LSigB :
 forall x : float, (Float 1%nat (Fexp x) <= Float 1%nat (LSB x))%R.
intros x; apply (oneExp_le radix); auto.
apply Fexp_le_LSB; auto.
Qed.

Theorem MSB_le_abs :
 forall x : float, ~ is_Fzero x -> (Float 1%nat (MSB x) <= Fabs x)%R.
intros x H'; unfold MSB, FtoRradix, FtoR in |- *; simpl in |- *.
replace (Z.pred (Fdigit radix x + Fexp x)) with
 (Z.pred (Fdigit radix x) + Fexp x)%Z; [ idtac | unfold Z.pred in |- *; ring ].
rewrite powerRZ_add; auto with real zarith arith.
rewrite Rmult_1_l.
repeat rewrite (fun r : R => Rmult_comm r (powerRZ radix (Fexp x))).
apply Rmult_le_compat_l; auto with real zarith.
now apply powerRZ_le, Rlt_IZR.
rewrite <- inj_pred; auto with real zarith.
rewrite <- Zpower_nat_Z_powerRZ; auto.
apply Rle_IZR; auto.
unfold Fdigit in |- *; auto with arith.
apply digitLess; auto.
unfold Fdigit in |- *.
apply not_eq_sym; apply lt_O_neq; apply digitNotZero; auto.
apply IZR_neq; omega.
Qed.

Theorem abs_lt_MSB :
 forall x : float, (Fabs x < Float 1%nat (Z.succ (MSB x)))%R.
intros x.
rewrite (MSB_abs x).
unfold MSB, FtoRradix, FtoR in |- *.
rewrite <- Zsucc_pred; simpl in |- *.
rewrite powerRZ_add; auto with real zarith.
rewrite Rmult_1_l.
repeat rewrite (fun r : R => Rmult_comm r (powerRZ radix (Fexp x))).
apply Rmult_lt_compat_l; auto with real zarith.
apply powerRZ_lt, Rlt_IZR; easy.
rewrite <- Zpower_nat_Z_powerRZ; auto with arith.
apply Rlt_IZR.
unfold Fdigit in |- *; auto with arith.
unfold Fabs in |- *; simpl in |- *.
pattern (Z.abs (Fnum x)) at 1 in |- *; rewrite <- (Z.abs_eq (Z.abs (Fnum x)));
 auto with zarith.
apply digitMore; easy.
apply IZR_neq; omega.
Qed.

Theorem LSB_le_abs :
 forall x : float, ~ is_Fzero x -> (Float 1%nat (LSB x) <= Fabs x)%R.
intros x H'; apply Rle_trans with (FtoRradix (Float 1%nat (MSB x))).
apply (oneExp_le radix); auto.
apply LSB_le_MSB; auto.
apply MSB_le_abs; auto.
Qed.

Theorem MSB_monotoneAux :
 forall x y : float,
 (Fabs x <= Fabs y)%R -> Fexp x = Fexp y -> (MSB x <= MSB y)%Z.
intros x y H' H'0; unfold MSB in |- *.
rewrite <- H'0.
cut (Fdigit radix x <= Fdigit radix y)%Z;
 [ unfold Z.pred in |- *; auto with zarith | idtac ].
unfold Fdigit in |- *; apply inj_le.
apply digit_monotone; auto.
apply le_IZR.
apply Rmult_le_reg_l with (r := powerRZ radix (Fexp x));
 auto with real zarith.
apply powerRZ_lt, Rlt_IZR; omega.
repeat rewrite (Rmult_comm (powerRZ radix (Fexp x))); auto.
pattern (Fexp x) at 2 in |- *; rewrite H'0; auto.
Qed.

Theorem MSB_monotone :
 forall x y : float,
 ~ is_Fzero x -> ~ is_Fzero y -> (Fabs x <= Fabs y)%R -> (MSB x <= MSB y)%Z.
intros x y H' H'0 H'1; rewrite (MSB_abs x); rewrite (MSB_abs y).
case (Zle_or_lt (Fexp (Fabs x)) (Fexp (Fabs y))); simpl in |- *; intros Zle1.
rewrite
 MSB_shift with (x := Fabs y) (n := Z.abs_nat (Fexp (Fabs y) - Fexp (Fabs x))).
apply MSB_monotoneAux; auto.
unfold FtoRradix in |- *; repeat rewrite Fabs_correct; auto with real arith.
rewrite FshiftCorrect; auto with real arith.
repeat rewrite Fabs_correct; auto with real arith.
repeat rewrite Rabs_Rabsolu; repeat rewrite <- Fabs_correct;
 auto with real arith.
unfold Fshift in |- *; simpl in |- *.
rewrite inj_abs; [ ring | auto with zarith ].
apply Fabs_Fzero; auto.
rewrite
 MSB_shift with (x := Fabs x) (n := Z.abs_nat (Fexp (Fabs x) - Fexp (Fabs y))).
apply MSB_monotoneAux; auto.
unfold FtoRradix in |- *; repeat rewrite Fabs_correct; auto with real arith.
rewrite FshiftCorrect; auto with real arith.
repeat rewrite Fabs_correct; auto with real arith.
repeat rewrite Rabs_Rabsolu; repeat rewrite <- Fabs_correct;
 auto with real arith.
unfold Fshift in |- *; simpl in |- *.
rewrite inj_abs; [ ring | auto with zarith ].
apply Fabs_Fzero; auto.
Qed.

Theorem LSB_rep_min :
 forall p : float, exists z : Z, p = Float z (LSB p) :>R.
intros p;
 exists (Zquotient (Fnum p) (Zpower_nat radix (Z.abs_nat (LSB p - Fexp p)))).
unfold FtoRradix, FtoR, LSB in |- *; simpl in |- *.
rewrite powerRZ_add; auto with real zarith.
2: apply IZR_neq; omega.
rewrite <- Rmult_assoc.
replace (maxDiv (Fnum p) (Fdigit radix p) + Fexp p - Fexp p)%Z with
 (Z_of_nat (maxDiv (Fnum p) (Fdigit radix p))); auto.
rewrite absolu_INR.
rewrite <- Zpower_nat_Z_powerRZ; auto with zarith.
rewrite <- mult_IZR.
rewrite <- ZdividesZquotient; auto with zarith.
now apply sym_not_eq, Zlt_not_eq, Zpower_nat_less.
apply maxDivCorrect.
ring.
Qed.
End mf.

Section FRoundP.
Variable b : Fbound.
Variable radix : Z.
Variable precision : nat.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.
Hypothesis radixMoreThanOne : (1 < radix)%Z.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ radixMoreThanOne).

Hypothesis precisionGreaterThanOne : 1 < precision.
Hypothesis pGivesBound : Zpos (vNum b) = Zpower_nat radix precision.

Definition Fulp (p : float) :=
  powerRZ radix (Fexp (Fnormalize radix b precision p)).

Theorem FulpComp :
 forall p q : float,
 Fbounded b p -> Fbounded b q -> p = q :>R -> Fulp p = Fulp q.
intros p q H' H'0 H'1; unfold Fulp in |- *.
rewrite
 FcanonicUnique
                with
                (p := Fnormalize radix b precision p)
               (q := Fnormalize radix b precision q)
               (3 := pGivesBound); auto with zarith.
apply FnormalizeCanonic; auto with zarith.
apply FnormalizeCanonic; auto with zarith.
apply trans_eq with (FtoR radix p).
apply FnormalizeCorrect; auto.
apply trans_eq with (FtoR radix q); auto.
apply sym_eq; apply FnormalizeCorrect; auto.
Qed.

Theorem FulpLe :
 forall p : float, Fbounded b p -> (Fulp p <= Float 1 (Fexp p))%R.
intros p H'; unfold Fulp, FtoRradix, FtoR, Fnormalize in |- *; simpl in |- *;
 rewrite Rmult_1_l.
case (Z_zerop (Fnum p)); simpl in |- *; auto.
intros H'0; apply (Rle_powerRZ radix (- dExp b) (Fexp p));
 auto with real zarith; try apply H'.
apply Rle_IZR; omega.
intros H'0; apply Rle_powerRZ; try apply Rle_IZR; auto with real zarith arith.
Qed.

Theorem Fulp_zero :
 forall x : float, is_Fzero x -> Fulp x = powerRZ radix (- dExp b).
intros x; unfold is_Fzero, Fulp, Fnormalize in |- *; case (Z_zerop (Fnum x));
 simpl in |- *; auto.
intros H' H'0; contradict H'; auto.
Qed.

Theorem FulpLe2 :
 forall p : float,
 Fbounded b p ->
 Fnormal radix b (Fnormalize radix b precision p) ->
 (Fulp p <= Rabs p * powerRZ radix (Z.succ (- precision)))%R.
intros p H1 H2; unfold Fulp in |- *.
replace (FtoRradix p) with (FtoRradix (Fnormalize radix b precision p));
 [ idtac | unfold FtoRradix in |- *; apply FnormalizeCorrect; auto ].
apply Rmult_le_reg_l with (powerRZ radix (Z.pred precision)).
apply powerRZ_lt, Rlt_IZR; auto with real arith.
replace
 (powerRZ radix (Z.pred precision) *
  (Rabs (Fnormalize radix b precision p) *
   powerRZ radix (Z.succ (- precision))))%R with
 (Rabs (Fnormalize radix b precision p)).
unfold FtoRradix in |- *; rewrite <- Fabs_correct; auto with arith real.
unfold Fabs, FtoR in |- *; simpl in |- *.
apply Rmult_le_compat_r; [ apply powerRZ_le, Rlt_IZR | rewrite <- inj_pred ];
 auto with real arith zarith.
rewrite <- Zpower_nat_Z_powerRZ.
replace (Zpower_nat radix (pred precision)) with (nNormMin radix precision);
 auto; apply Rle_IZR.
apply pNormal_absolu_min with b; auto with arith zarith real.
apply
 trans_eq
  with
    (Rabs (Fnormalize radix b precision p) *
     (powerRZ radix (Z.pred precision) * powerRZ radix (Z.succ (- precision))))%R;
 [ idtac | ring ].
rewrite <- powerRZ_add; auto with zarith real.
replace (Z.pred precision + Z.succ (- precision))%Z with 0%Z;
 [ simpl in |- *; ring | unfold Z.succ, Z.pred in |- *; ring ];
 auto with real zarith.
apply IZR_neq; omega.
Qed.

Theorem FulpGe :
 forall p : float,
 Fbounded b p -> (Rabs p <= (powerRZ radix precision - 1) * Fulp p)%R.
intros p H.
replace (FtoRradix p) with (FtoRradix (Fnormalize radix b precision p));
 [ idtac | unfold FtoRradix in |- *; apply FnormalizeCorrect; auto ].
unfold FtoRradix in |- *; rewrite <- Fabs_correct; auto with arith real.
unfold FtoR in |- *; simpl in |- *; unfold Fulp in |- *.
apply Rmult_le_compat_r; [ apply powerRZ_le, Rlt_IZR | idtac ];
 auto with real arith zarith.
apply Rle_trans with (IZR (Z.pred (Zpos (vNum b))));
 [ apply Rle_IZR; auto with zarith | idtac ].
apply Zle_Zpred.
apply FnormalizeBounded; auto with zarith.
unfold Z.pred in |- *; right; rewrite pGivesBound; replace 1%R with (IZR 1);
 auto with real.
rewrite <- Zpower_nat_Z_powerRZ; rewrite Z_R_minus;auto.
Qed.

Theorem LeFulpPos :
 forall x y : float,
 Fbounded b x ->
 Fbounded b y -> (0 <= x)%R -> (x <= y)%R -> (Fulp x <= Fulp y)%R.
intros x y Hx Hy H1 H2; unfold Fulp in |- *.
apply Rle_powerRZ; try apply Rle_IZR; auto with real zarith.
apply Fcanonic_Rle_Zle with radix b precision; auto with zarith arith.
apply FnormalizeCanonic; auto with zarith arith.
apply FnormalizeCanonic; auto with zarith arith.
repeat rewrite FnormalizeCorrect; auto with zarith arith real.
repeat rewrite Rabs_right; auto with zarith arith real.
apply Rge_trans with (FtoRradix x); auto with real.
Qed.

Theorem FulpSucCan :
 forall p : float,
 Fcanonic radix b p -> (FSucc b radix precision p - p <= Fulp p)%R.
intros p H'.
replace (Fulp p) with (powerRZ radix (Fexp p)).
2: unfold Fulp in |- *; replace (Fnormalize radix b precision p) with p; auto;
    apply sym_equal; apply FcanonicUnique with (3 := pGivesBound);
    auto with arith; try apply FnormalizeCanonic; auto with zarith.
2: apply FcanonicBound with (1 := H'); auto with zarith.
2: apply FnormalizeCorrect; easy.
unfold FtoRradix in |- *; rewrite <- Fminus_correct; auto with zarith.
case (Z.eq_dec (Fnum p) (- nNormMin radix precision)); intros H1'.
case (Z.eq_dec (Fexp p) (- dExp b)); intros H2'.
rewrite FSuccDiff2; auto with arith.
unfold FtoR in |- *; simpl in |- *; rewrite Rmult_1_l; auto with real.
rewrite FSuccDiff3; auto with arith zarith.
unfold FtoR in |- *; simpl in |- *; rewrite Rmult_1_l.
apply Rlt_le; apply Rlt_powerRZ; try apply Rlt_IZR; auto with real zarith.
unfold Z.pred in |- *; auto with zarith.
rewrite FSuccDiff1; auto with arith zarith.
unfold FtoR in |- *; simpl in |- *; rewrite Rmult_1_l; auto with real zarith.
Qed.

Theorem FulpSuc :
 forall p : float,
 Fbounded b p -> (FNSucc b radix precision p - p <= Fulp p)%R.
intros p H'.
replace (Fulp p) with (Fulp (Fnormalize radix b precision p)).
replace (FtoRradix p) with (FtoRradix (Fnormalize radix b precision p)).
unfold FNSucc in |- *; apply FulpSucCan; auto with arith.
apply FnormalizeCanonic; auto with zarith.
unfold FtoRradix in |- *; apply FnormalizeCorrect; auto.
apply FulpComp; auto with zarith.
apply FnormalizeBounded; auto with zarith.
unfold FtoRradix in |- *; apply FnormalizeCorrect; auto.
Qed.

Theorem FulpPredCan :
 forall p : float,
 Fcanonic radix b p -> (p - FPred b radix precision p <= Fulp p)%R.
intros p H'.
replace (Fulp p) with (powerRZ radix (Fexp p)).
2: unfold Fulp in |- *; replace (Fnormalize radix b precision p) with p; auto;
    apply sym_equal; apply FcanonicUnique with (3 := pGivesBound);
    auto with arith; try apply FnormalizeCanonic; auto with arith zarith.
2: apply FcanonicBound with (1:=H').
2: apply FnormalizeCorrect; easy.
unfold FtoRradix in |- *; rewrite <- Fminus_correct; auto with arith.
case (Z.eq_dec (Fnum p) (nNormMin radix precision)); intros H1'.
case (Z.eq_dec (Fexp p) (- dExp b)); intros H2'.
rewrite FPredDiff2; auto with zarith.
unfold FtoR in |- *; simpl in |- *; rewrite Rmult_1_l; auto with real.
rewrite FPredDiff3; auto with arith.
unfold FtoR in |- *; simpl in |- *; rewrite Rmult_1_l; auto with real.
apply Rlt_le; apply Rlt_powerRZ; try apply Rlt_IZR; auto with real zarith.
replace 1%R with (INR 1); auto with real arith.
unfold Z.pred in |- *; auto with zarith.
rewrite FPredDiff1; auto with zarith.
unfold FtoR in |- *; simpl in |- *; rewrite Rmult_1_l; auto with real.
Qed.

Theorem FulpPred :
 forall p : float,
 Fbounded b p -> (p - FNPred b radix precision p <= Fulp p)%R.
intros p H'.
replace (Fulp p) with (Fulp (Fnormalize radix b precision p)).
replace (FtoRradix p) with (FtoRradix (Fnormalize radix b precision p)).
unfold FNPred in |- *; apply FulpPredCan; auto with arith.
apply FnormalizeCanonic; auto with zarith.
unfold FtoRradix in |- *; apply FnormalizeCorrect; auto.
apply FulpComp; auto with zarith.
apply FnormalizeBounded; auto with zarith.
unfold FtoRradix in |- *; apply FnormalizeCorrect; auto.
Qed.

Theorem FSuccDiffPos :
 forall x : float,
 (0 <= x)%R ->
 Fminus radix (FSucc b radix precision x) x = Float 1%nat (Fexp x) :>R.
intros x H.
unfold FtoRradix in |- *; apply FSuccDiff1; auto with zarith.
contradict H; unfold FtoRradix, FtoR in |- *; simpl in |- *; rewrite H.
apply Rlt_not_le.
replace 0%R with (0 * powerRZ radix (Fexp x))%R; [ idtac | ring ].
apply Rlt_monotony_exp; auto with real arith.
assert (0 < nNormMin radix precision)%Z.
apply nNormPos; auto with zarith.
replace 0%R with (IZR (- 0)); auto with real zarith arith.
apply Rlt_IZR; omega.
Qed.

Theorem FulpFPredGePos :
 forall f : float,
 Fbounded b f ->
 Fcanonic radix b f ->
 (0 < f)%R -> (Fulp (FPred b radix precision f) <= Fulp f)%R.
intros f Hf1 Hf2 H.
apply LeFulpPos; auto with zarith; unfold FtoRradix in |- *.
apply FBoundedPred; auto with zarith.
apply R0RltRlePred; auto with zarith.
apply Rlt_le; apply FPredLt; auto with zarith.
Qed.

Theorem CanonicFulp :
 forall p : float, Fcanonic radix b p -> Fulp p = Float 1%nat (Fexp p).
intros p H; unfold Fulp in |- *.
rewrite FcanonicFnormalizeEq; auto with zarith.
unfold FtoRradix, FtoR in |- *; simpl in |- *; ring.
Qed.

Theorem FSuccUlpPos :
 forall x : float,
 Fcanonic radix b x ->
 (0 <= x)%R -> Fminus radix (FSucc b radix precision x) x = Fulp x :>R.
intros x H H0; rewrite CanonicFulp; auto.
apply FSuccDiffPos; auto.
Qed.

Theorem FulpFabs : forall f : float, Fulp f = Fulp (Fabs f) :>R.
intros f; unfold Fulp in |- *; case (Rle_or_lt 0 f); intros H'.
replace (Fabs f) with f; auto; unfold Fabs in |- *; apply floatEq;
 simpl in |- *; auto with zarith real.
apply sym_eq; apply Z.abs_eq; apply LeR0Fnum with radix; auto with zarith real.
replace (Fabs f) with (Fopp f);
 [ rewrite Fnormalize_Fopp | apply floatEq; simpl in |- * ];
 auto with zarith.
apply sym_eq; apply Zabs_eq_opp; apply R0LeFnum with radix;
 auto with zarith real.
Qed.

Theorem RoundedModeUlp :
 forall P,
 RoundedModeP b radix P ->
 forall (p : R) (q : float), P p q -> (Rabs (p - q) < Fulp q)%R.
intros P H' p q H'0.
case (Req_dec p q); intros Eq1.
rewrite <- Eq1.
replace (p - p)%R with 0%R; [ idtac | ring ].
rewrite Rabs_R0; auto.
unfold Fulp, FtoRradix, FtoR in |- *; simpl in |- *; auto with real arith.
apply powerRZ_lt, Rlt_IZR; auto with zarith.
case H'.
intros H'1 H'2; elim H'2; intros H'3 H'4; elim H'4; intros H'5 H'6;
 case H'5 with (1 := H'0); clear H'5 H'4 H'2; intros H'5.
rewrite Rabs_right; auto.
cut (Fbounded b q); [ intros B0 | case H'5; auto ].
apply Rlt_le_trans with (2 := FulpSuc q B0).
apply Rplus_lt_reg_l with (r := FtoR radix q).
repeat rewrite Rplus_minus; auto.
case (Rle_or_lt (FNSucc b radix precision q) p); auto.
intros H'2; absurd (FNSucc b radix precision q <= q)%R; auto.
apply Rgt_not_le; red in |- *; unfold FtoRradix in |- *;
 auto with real arith.
apply FNSuccLt; auto with zarith.
case H'5; auto.
intros H'4 H'7; elim H'7; intros H'8 H'9; apply H'9; clear H'7; auto.
apply (FcanonicBound radix b); auto with zarith.
apply FNSuccCanonic; auto with zarith.
apply Rle_ge; apply Rplus_le_reg_l with (r := FtoR radix q).
repeat rewrite Rplus_minus; auto.
rewrite Rplus_0_r; apply isMin_inv1 with (1 := H'5); auto.
rewrite Rabs_left1; auto.
rewrite Ropp_minus_distr; auto.
cut (Fbounded b q); [ intros B0 | case H'5; auto ].
apply Rlt_le_trans with (2 := FulpPred q B0).
apply Ropp_lt_cancel; repeat rewrite Rminus_0_l.
apply Rplus_lt_reg_l with (r := FtoR radix q).
repeat rewrite Rplus_minus; auto.
case (Rle_or_lt p (FNPred b radix precision q)); auto.
intros H'2; absurd (q <= FNPred b radix precision q)%R; auto.
apply Rgt_not_le; red in |- *; unfold FtoRradix in |- *;
 auto with real arith.
apply FNPredLt; auto with zarith.
case H'5; auto.
intros H'4 H'7; elim H'7; intros H'8 H'9; apply H'9; clear H'7; auto.
apply (FcanonicBound radix b); auto with arith.
apply FNPredCanonic; auto with zarith.
intros H1; apply Rplus_lt_compat_l; auto with real; apply Ropp_lt_contravar;
 unfold Rminus in |- *; auto with real.
apply Rplus_le_reg_l with (r := FtoR radix q).
repeat rewrite Rplus_minus; auto.
rewrite Rplus_0_r; apply isMax_inv1 with (1 := H'5).
Qed.

Theorem RoundedModeErrorExpStrict :
 forall P,
 RoundedModeP b radix P ->
 forall (p q : float) (x : R),
 Fbounded b p ->
 Fbounded b q ->
 P x p -> q = (x - p)%R :>R -> q <> 0%R :>R -> (Fexp q < Fexp p)%Z.
intros P H p q x H0 H1 H2 H3 H4.
apply Zlt_powerRZ with (e := IZR radix); try apply Rle_IZR; auto with real zarith.
apply Rle_lt_trans with (FtoRradix (Fabs q)).
replace (powerRZ radix (Fexp q)) with (FtoRradix (Float 1%nat (Fexp q)));
 unfold FtoRradix in |- *;
 [ apply RleFexpFabs; auto with arith
 | unfold FtoR in |- *; simpl in |- *; ring ].
rewrite (Fabs_correct radix); auto with zarith.
fold FtoRradix in |- *; rewrite H3.
apply Rlt_le_trans with (Fulp p);
 [ apply RoundedModeUlp with P; auto | unfold Fulp in |- * ].
apply Rle_powerRZ; try apply Rle_IZR; auto with real zarith.
apply FcanonicLeastExp with radix b precision; auto with real zarith.
apply sym_eq; apply FnormalizeCorrect; auto.
apply FnormalizeCanonic; auto with zarith.
Qed.

Theorem RoundedModeProjectorIdem :
 forall P (p : float), RoundedModeP b radix P -> Fbounded b p -> P p p.
intros P p H' H.
elim H'; intros H'0 H'1; elim H'1; intros H'2 H'3; elim H'3; intros H'4 H'5;
 clear H'3 H'1.
case (H'0 p).
intros x H'6.
apply (H'2 p p x); auto.
apply sym_eq; apply (RoundedProjector _ _ P H'); auto.
Qed.

Theorem RoundedModeBounded :
 forall P (r : R) (q : float),
 RoundedModeP b radix P -> P r q -> Fbounded b q.
intros P r q H' H'0.
case H'.
intros H'1 H'2; elim H'2; intros H'3 H'4; elim H'4; intros H'5 H'6;
 case H'5 with (1 := H'0); clear H'4 H'2; intros H;
 case H; auto.
Qed.

Theorem RoundedModeProjectorIdemEq :
 forall (P : R -> float -> Prop) (p q : float),
 RoundedModeP b radix P -> Fbounded b p -> P (FtoR radix p) q -> p = q :>R.
intros P p q H' H'0 H'1.
cut (MinOrMaxP b radix P);
 [ intros Mn; case (Mn p q); auto; intros Mn1 | apply H'].
apply sym_eq; apply MinEq with (1 := Mn1); auto.
apply (RoundedModeProjectorIdem (isMin b radix)); auto.
apply MinRoundedModeP with (precision := precision); auto.
apply sym_eq; apply MaxEq with (1 := Mn1); auto.
apply (RoundedModeProjectorIdem (isMax b radix)); auto.
apply MaxRoundedModeP with (precision := precision); auto.
Qed.

Theorem RoundedModeMult :
 forall P,
 RoundedModeP b radix P ->
 forall (r : R) (q q' : float),
 P r q -> Fbounded b q' -> (r <= radix * q')%R -> (q <= radix * q')%R.
intros P H' r q q' H'0 H'1.
replace (radix * q')%R with (FtoRradix (Float (Fnum q') (Fexp q' + 1%nat))).
intros H'2; case H'2.
intros H'3; case H'; intros H'4 H'5; elim H'5; intros H'6 H'7; elim H'7;
 intros H'8 H'9; apply H'9 with (1 := H'3); clear H'7 H'5;
 auto.
apply RoundedModeProjectorIdem; auto.
apply FBoundedScale; auto.
intros H'3.
replace (FtoRradix q) with (FtoRradix (Float (Fnum q') (Fexp q' + 1%nat)));
 auto with real.
apply (RoundedProjector _ _ P H'); auto.
apply FBoundedScale; auto.
case H'.
intros H'4 H'5; elim H'5; intros H'6 H'7; clear H'5.
apply (H'6 r (Float (Fnum q') (Fexp q' + 1%nat)) q); auto.
apply RoundedModeBounded with (P := P) (r := r); auto.
rewrite (FvalScale _ radixMoreThanOne b).
rewrite powerRZ_1; auto.
Qed.

Theorem RoundedModeMultLess :
 forall P,
 RoundedModeP b radix P ->
 forall (r : R) (q q' : float),
 P r q -> Fbounded b q' -> (radix * q' <= r)%R -> (radix * q' <= q)%R.
intros P H' r q q' H'0 H'1.
replace (radix * q')%R with (FtoRradix (Float (Fnum q') (Fexp q' + 1%nat))).
intros H'2; case H'2.
intros H'3; case H'; intros H'4 H'5; elim H'5; intros H'6 H'7; elim H'7;
 intros H'8 H'9; apply H'9 with (1 := H'3); clear H'7 H'5;
 auto.
apply RoundedModeProjectorIdem; auto.
apply FBoundedScale; auto.
intros H'3.
replace (FtoRradix q) with (FtoRradix (Float (Fnum q') (Fexp q' + 1%nat)));
 auto with real.
apply (RoundedProjector _ _ P H'); auto.
apply FBoundedScale; auto.
unfold FtoRradix in H'3; rewrite H'3; auto.
case H'.
intros H'4 H'5; elim H'5; intros H'6 H'7; clear H'5.
unfold FtoRradix in |- *; rewrite FvalScale; auto.
rewrite powerRZ_1; auto.
Qed.

Theorem RleMinR0 :
 forall (r : R) (min : float),
 (0 <= r)%R -> isMin b radix r min -> (0 <= min)%R.
intros r min H' H'0.
rewrite <- (FzeroisZero radix b).
case H'; intros H'1.
apply (MonotoneMin b radix) with (p := FtoRradix (Fzero (- dExp b))) (q := r);
 auto.
unfold FtoRradix in |- *; rewrite (FzeroisZero radix b); auto.
apply (RoundedModeProjectorIdem (isMin b radix)); auto.
apply MinRoundedModeP with (precision := precision); auto with zarith.
apply FboundedFzero; auto.
replace (FtoR radix (Fzero (- dExp b))) with (FtoRradix min); auto with real.
apply sym_eq; apply (ProjectMin b radix).
apply FboundedFzero; auto.
rewrite <- H'1 in H'0; rewrite <- (FzeroisZero radix b) in H'0; auto.
Qed.

Theorem RleRoundedR0 :
 forall (P : R -> float -> Prop) (p : float) (r : R),
 RoundedModeP b radix P -> P r p -> (0 <= r)%R -> (0 <= p)%R.
intros P p r H' H'0 H'1.
case H'.
intros H'2 H'3; Elimc H'3; intros H'3 H'4; Elimc H'4; intros H'4 H'5;
 case (H'4 r p); auto; intros H'6.
apply RleMinR0 with (r := r); auto.
apply Rle_trans with r; auto; apply isMax_inv1 with (1 := H'6).
Qed.

Theorem RleMaxR0 :
 forall (r : R) (max : float),
 (r <= 0)%R -> isMax b radix r max -> (max <= 0)%R.
intros r max H' H'0.
rewrite <- (FzeroisZero radix b).
case H'; intros H'1.
apply (MonotoneMax b radix) with (q := FtoRradix (Fzero (- dExp b))) (p := r);
 auto.
unfold FtoRradix in |- *; rewrite FzeroisZero; auto.
apply (RoundedModeProjectorIdem (isMax b radix)); auto.
apply MaxRoundedModeP with (precision := precision); auto.
apply FboundedFzero; auto.
replace (FtoR radix (Fzero (- dExp b))) with (FtoRradix max); auto with real.
apply sym_eq; apply (ProjectMax b radix).
apply FboundedFzero; auto.
rewrite H'1 in H'0; rewrite <- (FzeroisZero radix b) in H'0; auto.
Qed.

Theorem RleRoundedLessR0 :
 forall (P : R -> float -> Prop) (p : float) (r : R),
 RoundedModeP b radix P -> P r p -> (r <= 0)%R -> (p <= 0)%R.
intros P p r H' H'0 H'1.
case H'.
intros H'2 H'3; Elimc H'3; intros H'3 H'4; Elimc H'4; intros H'4 H'5;
 case (H'4 r p); auto; intros H'6.
apply Rle_trans with r; auto; apply isMin_inv1 with (1 := H'6).
apply RleMaxR0 with (r := r); auto.
Qed.

Theorem PminPos :
 forall p min : float,
 (0 <= p)%R ->
 Fbounded b p ->
 isMin b radix (/ 2%nat * p) min ->
 exists c : float, Fbounded b c /\ c = (p - min)%R :>R.
intros p min H' H'0 H'1.
cut (min <= / 2%nat * p)%R;
 [ intros Rl1; Casec Rl1; intros Rl1
 | apply isMin_inv1 with (1 := H'1); auto ].
case (eqExpMax _ radixMoreThanOne b min p); auto.
case H'1; auto.
rewrite Fabs_correct; auto with arith.
rewrite Rabs_right; auto.
apply Rle_trans with (/ 2%nat * p)%R; auto.
apply Rlt_le; auto.
apply Rmult_le_reg_l with (r := INR 2); auto with real.
rewrite <- Rmult_assoc; rewrite Rinv_r; auto with real; rewrite Rmult_ne_r;
 auto with real.
apply Rle_ge; apply RleMinR0 with (r := (/ 2%nat * p)%R); auto.
apply Rmult_le_reg_l with (r := INR 2); auto with real.
rewrite <- Rmult_assoc; rewrite Rinv_r; auto with real; rewrite Rmult_1_l;
 rewrite Rmult_0_r; auto with real.
intros min' H'2; elim H'2; intros H'3 H'4; elim H'4; intros H'5 H'6;
 clear H'4 H'2.
case (FboundNext _ radixMoreThanOne b precision) with (p := min');
 auto with zarith; fold FtoRradix in |- *.
intros Smin H'2; elim H'2; intros H'4 H'7; clear H'2.
exists Smin; split; auto.
rewrite H'7; auto.
unfold FtoRradix in |- *.
rewrite <- H'5; auto.
replace (Float (Z.succ (Fnum min')) (Fexp min')) with
 (Float (Fnum (Fshift radix (Z.abs_nat (Fexp p - Fexp min')) p) - Fnum min')
    (Fexp min')); auto.
unfold FtoRradix, FtoR in |- *; simpl in |- *.
rewrite <- Z_R_minus.
rewrite (fun x y z : R => Rmult_comm (x - y) z); rewrite Rmult_minus_distr_l;
 repeat rewrite (fun x : Z => Rmult_comm (powerRZ radix x)).
rewrite mult_IZR.
rewrite Zpower_nat_powerRZ_absolu; auto with zarith.
rewrite Rmult_assoc.
rewrite <- (powerRZ_add radix); try apply IZR_neq; auto with real zarith.
replace (Fexp p - Fexp min' + Fexp min')%Z with (Fexp p); [ auto | ring ].
apply floatEq; auto; simpl in |- *.
apply Zle_antisym.
apply Zlt_succ_le.
apply Zplus_lt_reg_l with (p := Fnum min'); auto.
cut (forall x y : Z, (x + (y - x))%Z = y);
 [ intros tmp; rewrite tmp; clear tmp | intros; ring ].
replace (Fnum min' + Z.succ (Z.succ (Fnum min')))%Z with
 (2%nat * Z.succ (Fnum min'))%Z.
apply (Rlt_Float_Zlt radix) with (r := Fexp min'); auto;
 fold FtoRradix in |- *.
replace (FtoRradix (Float (2%nat * Z.succ (Fnum min')) (Fexp min'))) with
 (2%nat * Float (Z.succ (Fnum min')) (Fexp min'))%R.
rewrite <- H'7.
replace
 (Float (Fnum p * Zpower_nat radix (Z.abs_nat (Fexp p - Fexp min')))
    (Fexp min')) with (Fshift radix (Z.abs_nat (Fexp p - Fexp min')) p).
unfold FtoRradix in |- *; rewrite FshiftCorrect; auto.
apply Rmult_lt_reg_l with (r := (/ 2%nat)%R); auto with real.
rewrite <- Rmult_assoc; rewrite Rinv_l; auto with real; rewrite Rmult_1_l;
 auto with real.
case (Rle_or_lt Smin (/ 2%nat * FtoR radix p)); auto.
intros H'2; absurd (min < Smin)%R.
apply Rle_not_lt.
case H'1; auto.
intros H'8 H'9; elim H'9; intros H'10 H'11; apply H'11; clear H'9; auto.
rewrite H'7; unfold FtoRradix in |- *; rewrite <- H'5; auto.
unfold FtoR in |- *; simpl in |- *; apply Rlt_monotony_exp;
 auto with real zarith.
apply Rlt_IZR; omega.
unfold Fshift in |- *; simpl in |- *.
replace (Fexp p - Z.abs_nat (Fexp p - Fexp min'))%Z with (Fexp min'); auto.
rewrite inj_abs; auto.
ring.
auto with zarith.
replace (FtoRradix (Float (2%nat * Z.succ (Fnum min')) (Fexp min'))) with
 ((2%nat * Z.succ (Fnum min'))%Z * powerRZ radix (Fexp min'))%R.
rewrite mult_IZR; auto.
unfold FtoRradix, FtoR in |- *; simpl in |- *; ring.
simpl in |- *; auto.
replace (Z_of_nat 2) with (Z.succ (Z.succ 0)).
repeat rewrite <- Zmult_succ_l_reverse; unfold Z.succ in |- *; ring.
simpl in |- *; auto.
apply Zlt_le_succ; auto.
apply Zplus_lt_reg_l with (p := Fnum min'); auto.
cut (forall x y : Z, (x + (y - x))%Z = y);
 [ intros tmp; rewrite tmp; clear tmp | intros; ring ].
replace (Fnum min' + Fnum min')%Z with (2%nat * Fnum min')%Z.
apply (Rlt_Float_Zlt radix) with (r := Fexp min'); auto;
 fold FtoRradix in |- *.
replace (FtoRradix (Float (2%nat * Fnum min') (Fexp min'))) with
 (2%nat * Float (Fnum min') (Fexp min'))%R.
replace
 (Float (Fnum p * Zpower_nat radix (Z.abs_nat (Fexp p - Fexp min')))
    (Fexp min')) with (Fshift radix (Z.abs_nat (Fexp p - Fexp min')) p).
unfold FtoRradix in |- *; rewrite FshiftCorrect; auto.
apply Rmult_lt_reg_l with (r := (/ 2%nat)%R); auto with real.
rewrite <- Rmult_assoc; rewrite Rinv_l; auto with real; rewrite Rmult_1_l;
 auto with real.
replace (FtoR radix (Float (Fnum min') (Fexp min'))) with (FtoR radix min);
 auto.
unfold Fshift in |- *; simpl in |- *.
replace (Fexp p - Z.abs_nat (Fexp p - Fexp min'))%Z with (Fexp min'); auto.
rewrite inj_abs; auto.
ring.
auto with zarith.
replace (FtoRradix (Float (2%nat * Fnum min') (Fexp min'))) with
 ((2%nat * Fnum min')%Z * powerRZ radix (Fexp min'))%R.
rewrite mult_IZR; auto.
unfold FtoRradix, FtoR in |- *; simpl in |- *; ring.
simpl in |- *; auto.
replace (Z_of_nat 2) with (Z.succ (Z.succ 0)).
repeat rewrite <- Zmult_succ_l_reverse; unfold Z.succ in |- *; ring.
simpl in |- *; auto.
exists min; split; auto.
case H'1; auto.
rewrite Rl1.
pattern (FtoRradix p) at 2 in |- *;
 replace (FtoRradix p) with (2%nat * (/ 2%nat * p))%R.
simpl in |- *; ring.
rewrite <- Rmult_assoc; rewrite Rinv_r; auto with real; rewrite Rmult_ne_r;
 auto with real.
Qed.

Theorem div2IsBetweenPos :
 forall p min max : float,
 (0 <= p)%R ->
 Fbounded b p ->
 isMin b radix (/ 2%nat * p) min ->
 isMax b radix (/ 2%nat * p) max -> p = (min + max)%R :>R.
intros p min max P H' H'0 H'1; apply Rle_antisym.
apply Rplus_le_reg_l with (r := (- max)%R).
replace (- max + p)%R with (p - max)%R; [ idtac | ring ].
replace (- max + (min + max))%R with (FtoRradix min); [ idtac | ring ].
rewrite <- (Fminus_correct radix); auto with arith.
case H'0.
intros H'2 H'3; elim H'3; intros H'4 H'5; apply H'5; clear H'3; auto.
apply Sterbenz; auto.
case H'1; auto.
apply Rle_trans with (FtoRradix max); auto.
apply Rmult_le_reg_l with (r := INR 2); auto with real.
rewrite <- Rmult_assoc; rewrite Rinv_r; auto with real; rewrite Rmult_1_l;
 auto with real.
apply Rledouble; auto.
apply Rle_trans with (FtoRradix min); auto.
apply RleMinR0 with (r := (/ 2%nat * p)%R); auto.
apply Rmult_le_reg_l with (r := INR 2); auto with real.
rewrite <- Rmult_assoc; rewrite Rinv_r; auto with real; rewrite Rmult_1_l;
 rewrite Rmult_0_r; auto with real.
apply Rle_trans with (/ 2%nat * p)%R; auto; apply isMax_inv1 with (1 := H'1).
case H'1.
intros H'3 H'6; elim H'6; intros H'7 H'8; apply H'8; clear H'6; auto.
apply Rmult_le_reg_l with (r := INR 2); auto with real.
rewrite <- Rmult_assoc; rewrite Rinv_r; auto with real; rewrite Rmult_1_l;
 auto with real.
apply Rmult_le_reg_l with (r := (/ 2%nat)%R); auto with real.
rewrite <- Rmult_assoc; rewrite Rinv_l; auto with real; rewrite Rmult_1_l;
 auto with real.
apply isMax_inv1 with (1 := H'1).
rewrite Fminus_correct; auto with arith.
apply Rplus_le_reg_l with (r := FtoR radix max).
replace (FtoR radix max + (FtoR radix p - FtoR radix max))%R with
 (FtoR radix p); [ idtac | ring ].
apply Rplus_le_reg_l with (r := (- (/ 2%nat * p))%R).
replace (- (/ 2%nat * p) + FtoR radix p)%R with (/ 2%nat * p)%R.
replace (- (/ 2%nat * p) + (FtoR radix max + / 2%nat * p))%R with
 (FtoR radix max); [ apply isMax_inv1 with (1 := H'1) | ring ].
replace (FtoR radix p) with (2%nat * (/ 2%nat * p))%R.
simpl in |- *; ring.
rewrite <- Rmult_assoc; rewrite Rinv_r; auto with real.
apply Rplus_le_reg_l with (r := (- min)%R).
replace (- min + p)%R with (p - min)%R; [ idtac | ring ].
replace (- min + (min + max))%R with (FtoRradix max); [ idtac | ring ].
case (PminPos p min); auto.
intros x H'2; elim H'2; intros H'3 H'4; elim H'4; clear H'2.
case H'1.
intros H'2 H'5; elim H'5; intros H'6 H'7; apply H'7; clear H'5; auto.
unfold FtoRradix in H'4; rewrite H'4; auto.
fold FtoRradix in |- *; apply Rplus_le_reg_l with (r := FtoRradix min).
replace (min + (p - min))%R with (FtoRradix p); [ idtac | ring ].
apply Rplus_le_reg_l with (r := (- (/ 2%nat * p))%R).
replace (- (/ 2%nat * p) + p)%R with (/ 2%nat * p)%R.
replace (- (/ 2%nat * p) + (min + / 2%nat * p))%R with (FtoRradix min);
 [ apply isMin_inv1 with (1 := H'0) | ring ].
pattern (FtoRradix p) at 3 in |- *;
 replace (FtoRradix p) with (2%nat * (/ 2%nat * p))%R.
simpl in |- *; ring.
rewrite <- Rmult_assoc; rewrite Rinv_r; auto with real.
Qed.

Theorem div2IsBetween :
 forall p min max : float,
 Fbounded b p ->
 isMin b radix (/ 2%nat * p) min ->
 isMax b radix (/ 2%nat * p) max -> p = (min + max)%R :>R.
intros p min max H' H'0 H'1; case (Rle_or_lt 0 p); intros H'2.
apply div2IsBetweenPos; auto.
cut (forall x y : R, (- x)%R = (- y)%R -> x = y);
 [ intros H'3; apply H'3; clear H'3 | idtac ].
replace (- (min + max))%R with (- max + - min)%R; [ idtac | ring ].
repeat rewrite <- (Fopp_correct radix); auto with zarith.
apply div2IsBetweenPos; auto with zarith.
rewrite (Fopp_correct radix); auto.
replace 0%R with (-0)%R; try apply Rlt_le; auto with real.
now apply oppBounded.
replace (/ 2%nat * Fopp p)%R with (- (/ 2%nat * p))%R.
apply MaxOppMin; easy.
rewrite (Fopp_correct radix); auto; fold FtoRradix; ring.
replace (/ 2%nat * Fopp p)%R with (- (/ 2%nat * p))%R.
apply MinOppMax; easy.
rewrite (Fopp_correct radix); auto; fold FtoRradix;ring.
intros x y H'3; rewrite <- (Ropp_involutive x);
 rewrite <- (Ropp_involutive y); rewrite H'3; auto.
Qed.

Theorem RoundedModeMultAbs :
 forall P,
 RoundedModeP b radix P ->
 forall (r : R) (q q' : float),
 P r q ->
 Fbounded b q' -> (Rabs r <= radix * q')%R -> (Rabs q <= radix * q')%R.
intros P H' r q q' H'0 H'1 H'2.
case (Rle_or_lt 0 r); intros Rl0.
rewrite Rabs_right; auto.
apply RoundedModeMult with (P := P) (r := r); auto.
rewrite <- (Rabs_right r); auto with real.
apply Rle_ge; apply RleRoundedR0 with (P := P) (r := r); auto.
rewrite Rabs_left1; auto.
replace (radix * q')%R with (- (radix * - q'))%R;
 [ apply Ropp_le_contravar | ring ].
rewrite <- (Fopp_correct radix).
apply RoundedModeMultLess with (P := P) (r := r); auto.
apply oppBounded; auto.
unfold FtoRradix in |- *; rewrite Fopp_correct.
rewrite <- (Ropp_involutive r).
replace (radix * - FtoR radix q')%R with (- (radix * q'))%R;
 [ apply Ropp_le_contravar | fold FtoRradix;ring ]; auto.
rewrite <- (Rabs_left1 r); auto.
apply Rlt_le; auto.
apply RleRoundedLessR0 with (P := P) (r := r); auto.
apply Rlt_le; auto.
Qed.

Theorem isMinComp :
 forall (r1 r2 : R) (min max : float),
 isMin b radix r1 min ->
 isMax b radix r1 max -> (min < r2)%R -> (r2 < max)%R -> isMin b radix r2 min.
intros r1 r2 min max H' H'0 H'1 H'2; split.
case H'; auto.
split.
apply Rlt_le; auto.
intros f H'3 H'4.
case H'; auto.
intros H'5 H'6; elim H'6; intros H'7 H'8; apply H'8; clear H'6; auto.
case (Rle_or_lt (FtoR radix f) r1); auto; intros C1.
absurd (FtoR radix f < max)%R.
apply Rle_not_lt.
case H'0.
intros H'6 H'9; elim H'9; intros H'10 H'11; apply H'11; clear H'9; auto.
apply Rlt_le; auto.
apply Rle_lt_trans with (2 := H'2); auto.
Qed.

Theorem isMaxComp :
 forall (r1 r2 : R) (min max : float),
 isMin b radix r1 min ->
 isMax b radix r1 max -> (min < r2)%R -> (r2 < max)%R -> isMax b radix r2 max.
intros r1 r2 min max H' H'0 H'1 H'2; split.
case H'0; auto.
split.
apply Rlt_le; auto.
intros f H'3 H'4.
case H'0; auto.
intros H'5 H'6; elim H'6; intros H'7 H'8; apply H'8; clear H'6; auto.
case (Rle_or_lt r1 (FtoR radix f)); auto; intros C1.
absurd (min < FtoR radix f)%R.
apply Rle_not_lt.
case H'.
intros H'6 H'9; elim H'9; intros H'10 H'11; apply H'11; clear H'9; auto.
apply Rlt_le; auto.
apply Rlt_le_trans with (1 := H'1); auto.
Qed.

Theorem RleBoundRoundl :
 forall P,
 RoundedModeP b radix P ->
 forall (p q : float) (r : R),
 Fbounded b p -> (p <= r)%R -> P r q -> (p <= q)%R.
intros P H' p q r H'0 H'1 H'2; case H'1; intros H'3.
cut (MonotoneP radix P);
 [ intros Mn | apply RoundedModeP_inv4 with (1 := H'); auto ].
apply (Mn p r); auto.
apply RoundedModeProjectorIdem with (P := P); auto.
rewrite RoundedModeProjectorIdemEq with (P := P) (p := p) (q := q);
 auto with real.
cut (CompatibleP b radix P);
 [ intros Cp | apply RoundedModeP_inv2 with (1 := H'); auto ].
apply (Cp r p q); auto.
apply RoundedModeBounded with (P := P) (r := r); auto.
Qed.

Theorem RleBoundRoundr :
 forall P,
 RoundedModeP b radix P ->
 forall (p q : float) (r : R),
 Fbounded b p -> (r <= p)%R -> P r q -> (q <= p)%R.
intros P H' p q r H'0 H'1 H'2; case H'1; intros H'3.
cut (MonotoneP radix P);
 [ intros Mn | apply RoundedModeP_inv4 with (1 := H'); auto ].
apply (Mn r p); auto.
apply RoundedModeProjectorIdem with (P := P); auto.
rewrite RoundedModeProjectorIdemEq with (P := P) (p := p) (q := q);
 auto with real.
cut (CompatibleP b radix P);
 [ intros Cp | apply RoundedModeP_inv2 with (1 := H'); auto ].
apply (Cp r p q); auto.
apply RoundedModeBounded with (P := P) (r := r); auto.
Qed.

Theorem RoundAbsMonotoner :
 forall (P : R -> float -> Prop) (p : R) (q r : float),
 RoundedModeP b radix P ->
 Fbounded b r -> P p q -> (Rabs p <= r)%R -> (Rabs q <= r)%R.
intros P p q r H' H'0 H'1 H'2.
case (Rle_or_lt 0 p); intros Rl1.
rewrite Rabs_right; auto with real.
apply RleBoundRoundr with (P := P) (r := p); auto with real.
rewrite <- (Rabs_right p); auto with real.
apply Rle_ge; apply RleRoundedR0 with (P := P) (r := p); auto.
rewrite Rabs_left1; auto.
rewrite <- (Ropp_involutive r); apply Ropp_le_contravar.
rewrite <- (Fopp_correct radix); auto.
apply RleBoundRoundl with (P := P) (r := p); auto.
apply oppBounded; easy.
rewrite (Fopp_correct radix); rewrite <- (Ropp_involutive p);
 rewrite <- (Rabs_left1 p); auto with real;
 apply Rlt_le; auto.
apply RleRoundedLessR0 with (P := P) (r := p); auto; apply Rlt_le; auto.
Qed.

Theorem RoundAbsMonotonel :
 forall (P : R -> float -> Prop) (p : R) (q r : float),
 RoundedModeP b radix P ->
 Fbounded b r -> P p q -> (r <= Rabs p)%R -> (r <= Rabs q)%R.
intros P p q r H' H'0 H'1 H'2.
case (Rle_or_lt 0 p); intros Rl1.
rewrite Rabs_right; auto.
apply RleBoundRoundl with (P := P) (r := p); auto.
rewrite <- (Rabs_right p); auto with real.
apply Rle_ge; apply RleRoundedR0 with (P := P) (r := p); auto.
rewrite Rabs_left1; auto.
rewrite <- (Ropp_involutive r); apply Ropp_le_contravar.
rewrite <- (Fopp_correct radix); auto.
apply RleBoundRoundr with (P := P) (r := p); auto.
apply oppBounded; easy.
rewrite (Fopp_correct radix); rewrite <- (Ropp_involutive p);
 rewrite <- (Rabs_left1 p); auto with real;
 apply Rlt_le; auto.
apply RleRoundedLessR0 with (P := P) (r := p); auto; apply Rlt_le; auto.
Qed.


Theorem FUlp_Le_LSigB :
 forall x : float, Fbounded b x -> (Fulp x <= Float 1%nat (LSB radix x))%R.
intros x H; unfold is_Fzero, Fulp, Fnormalize in |- *;
 case (Z_zerop (Fnum x)); intros ZH.
unfold FtoRradix, FtoR in |- *; simpl in |- *.
rewrite Rmult_1_l.
apply Rle_powerRZ.
apply Rle_IZR; omega.
apply Z.le_trans with (Fexp x); auto.
case H; auto.
apply Fexp_le_LSB; auto.
rewrite
 LSB_shift
           with
           (n :=
             min (precision - Fdigit radix x) (Z.abs_nat (dExp b + Fexp x)));
 auto.
unfold FtoRradix, FtoR in |- *; simpl in |- *.
rewrite Rmult_1_l.
apply Rle_powerRZ; auto with arith.
apply Rle_IZR; omega.
replace 1%R with (INR 1); auto with real arith.
exact
 (Fexp_le_LSB radix
    (Fshift radix
       (min (precision - Fdigit radix x) (Z.abs_nat (dExp b + Fexp x))) x)).
Qed.

End FRoundP.


Section FRoundPP.
Variable b : Fbound.
Variable radix : Z.
Variable precision : nat.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.
Hypothesis radixMoreThanOne : (1 < radix)%Z.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ radixMoreThanOne).

Hypothesis precisionGreaterThanOne : 1 < precision.
Hypothesis pGivesBound : Zpos (vNum b) = Zpower_nat radix precision.

Theorem errorBoundedMultMin :
 forall p q fmin : float,
 Fbounded b p ->
 Fbounded b q ->
 (0 <= p)%R ->
 (0 <= q)%R ->
 (- dExp b <= Fexp p + Fexp q)%Z ->
 isMin b radix (p * q) fmin ->
 exists r : float,
   r = (p * q - fmin)%R :>R /\ Fbounded b r /\ Fexp r = (Fexp p + Fexp q)%Z.
intros p q fmin Fp Fq H' H'0 H'1 H'2.
cut (0 <= Fnum p * Fnum q)%Z;
 [ intros multPos
 | apply Zle_mult_gen; apply (LeR0Fnum radix); auto with arith ].
cut (ex (fun m : Z => FtoRradix fmin = Float m (Fexp (Fmult p q)))).
2: unfold FtoRradix in |- *;
    apply
     RoundedModeRep
      with (b := b) (precision := precision) (P := isMin b radix);
    auto.
2: apply MinRoundedModeP with (precision := precision); auto.
2: rewrite (Fmult_correct radix); auto with zarith.
intros H'3; elim H'3; intros m E; clear H'3.
exists (Fminus radix (Fmult p q) (Float m (Fexp (Fmult p q)))).
split.
rewrite E; unfold FtoRradix in |- *; repeat rewrite Fminus_correct;
 repeat rewrite Fmult_correct; auto with zarith.
split.
cut (fmin <= Fmult p q)%R;
 [ idtac
 | unfold FtoRradix in |- *; rewrite Fmult_correct; auto; case H'2;
    auto with real zarith; (intros H1 H2; case H2; auto with zarith) ].
rewrite E; unfold Fmult, Fminus, Fopp, Fplus in |- *; simpl in |- *; auto.
repeat rewrite Zmin_n_n; repeat rewrite <- Zminus_diag_reverse; auto.
simpl in |- *; repeat rewrite Zpower_nat_O; repeat rewrite Zmult_1_r.
unfold FtoRradix, FtoR in |- *; simpl in |- *.
intros H'3;
 (cut (m <= Fnum p * Fnum q)%Z;
   [ idtac
   | apply le_IZR;
      apply Rmult_le_reg_l with (r := powerRZ radix (Fexp p + Fexp q));
      auto with real zarith;
      repeat rewrite (Rmult_comm (powerRZ radix (Fexp p + Fexp q)));
      auto with zarith; apply powerRZ_lt, Rlt_IZR; omega ]); intros H'4.
repeat split; simpl in |- *; auto.
case (ZquotientProp (Fnum p * Fnum q) (Zpower_nat radix precision));
 auto with zarith.
intros x (H'5, (H'6, H'7)).
cut
 (Zquotient (Fnum p * Fnum q) (Zpower_nat radix precision) *
  powerRZ radix (precision + (Fexp p + Fexp q)) <= fmin)%R;
 [ rewrite E; intros H'8 | idtac ].
cut
 (Zquotient (Fnum p * Fnum q) (Zpower_nat radix precision) *
  powerRZ radix precision <= m)%R; [ intros H'9 | idtac ].
rewrite Z.abs_eq; auto with zarith.
apply Z.le_lt_trans with x; auto.
replace x with
 (Fnum p * Fnum q +
  -
  (Zquotient (Fnum p * Fnum q) (Zpower_nat radix precision) *
   Zpower_nat radix precision))%Z.
apply Zplus_le_compat_l; auto.
apply Zle_Zopp.
apply le_IZR; auto.
rewrite mult_IZR.
rewrite Zpower_nat_Z_powerRZ; auto with zarith.
pattern (Fnum p * Fnum q)%Z at 1 in |- *; rewrite H'5; ring.
rewrite pGivesBound.
rewrite <- (Z.abs_eq (Zpower_nat radix precision)); auto with zarith.
apply Rmult_le_reg_l with (r := powerRZ radix (Fexp p + Fexp q));
 auto with real zarith.
apply powerRZ_lt, Rlt_IZR; omega.
repeat rewrite (Rmult_comm (powerRZ radix (Fexp p + Fexp q))); auto.
rewrite Rmult_assoc; rewrite <- powerRZ_add. auto with real zarith.
apply IZR_neq; omega.
case
 (FboundedMbound _ radixMoreThanOne b precision)
  with
    (z := (precision + (Fexp p + Fexp q))%Z)
    (m := Zquotient (Fnum p * Fnum q) (Zpower_nat radix precision));
 auto with zarith.
apply Zmult_le_reg_r with (p := Zpower_nat radix precision); auto with zarith.
pattern (Zpower_nat radix precision) at 2 in |- *;
 rewrite <- (fun x => Z.abs_eq (Zpower_nat radix x)).
rewrite <- Zabs_Zmult.
apply Z.le_trans with (1 := H'6); auto with zarith.
rewrite Zabs_Zmult.
apply Z.le_trans with (Zpower_nat radix precision * Z.abs (Fnum q))%Z.
apply Zle_Zmult_comp_r; auto with zarith.
apply Zlt_le_weak; rewrite <- pGivesBound; apply Fp.
apply Zle_Zmult_comp_l; auto with zarith.
apply Zlt_le_weak; rewrite <- pGivesBound; apply Fq.
apply Zpower_NR0; omega.
intros x0 (H'8, H'9); rewrite <- H'9.
case H'2.
intros H'10 (H'11, H'12); apply H'12; auto.
rewrite H'9; auto.
rewrite powerRZ_add.
rewrite <- Rmult_assoc.
unfold FtoRradix in |- *; rewrite <- Fmult_correct; auto with zarith.
unfold Fmult, FtoR in |- *; simpl in |- *.
repeat rewrite (fun x => Rmult_comm x (powerRZ radix (Fexp p + Fexp q))).
apply Rmult_le_compat_l; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
rewrite <- Zpower_nat_Z_powerRZ.
pattern (Fnum p * Fnum q)%Z at 2 in |- *;
 rewrite <- (Z.abs_eq (Fnum p * Fnum q)); auto.
rewrite <- mult_IZR; apply Rle_IZR; apply Zle_Zabs_inv2; auto.
apply IZR_neq; omega.
simpl in |- *.
apply Zmin_n_n.
Qed.

Theorem errorBoundedMultMax :
 forall p q fmax : float,
 Fbounded b p ->
 Fbounded b q ->
 (0 <= p)%R ->
 (0 <= q)%R ->
 (- dExp b <= Fexp p + Fexp q)%Z ->
 isMax b radix (p * q) fmax ->
 exists r : float,
   FtoRradix r = (p * q - fmax)%R /\
   Fbounded b r /\ Fexp r = (Fexp p + Fexp q)%Z.
intros p q fmax Fp Fq H' H'0 H'1 H'2.
cut (0 <= Fnum p * Fnum q)%Z;
 [ intros multPos
 | apply Zle_mult_gen; apply (LeR0Fnum radix); auto with arith ].
case (ZquotientProp (Fnum p * Fnum q) (Zpower_nat radix precision));
 auto with zarith.
intros r; intros (H'3, (H'4, H'5)).
cut (0 <= Zquotient (Fnum p * Fnum q) (Zpower_nat radix precision))%Z;
 [ intros Z2 | apply ZquotientPos; auto with zarith ].
cut (0 <= r)%Z;
 [ intros Z3
 | replace r with
    (Fnum p * Fnum q -
     Zquotient (Fnum p * Fnum q) (Zpower_nat radix precision) *
     Zpower_nat radix precision)%Z;
    [ idtac | pattern (Fnum p * Fnum q)%Z at 1 in |- *; rewrite H'3; ring ];
    auto ].
2: apply Zle_Zminus_ZERO; rewrite Z.abs_eq in H'4; auto with zarith;
    rewrite Z.abs_eq in H'4; auto with zarith.
case (Z.eq_dec r 0); intros Z4.
exists (Fzero (Fexp p + Fexp q)); repeat (split; auto with zarith).
replace (FtoRradix (Fzero (Fexp p + Fexp q))) with 0%R;
 [ idtac | unfold Fzero, FtoRradix, FtoR in |- *; simpl in |- *; ring ].
apply Rplus_eq_reg_l with (r := FtoRradix fmax).
replace (fmax + 0)%R with (FtoRradix fmax); [ idtac | ring ].
replace (fmax + (p * q - fmax))%R with (p * q)%R; [ idtac | ring ].
unfold FtoRradix in |- *; rewrite <- (Fmult_correct radix); auto with zarith.
case
 (FboundedMbound _ radixMoreThanOne b precision)
  with
    (z := (precision + (Fexp p + Fexp q))%Z)
    (m := Zquotient (Fnum p * Fnum q) (Zpower_nat radix precision));
 auto with zarith.
apply Zmult_le_reg_r with (p := Zpower_nat radix precision); auto with zarith.
pattern (Zpower_nat radix precision) at 2 in |- *;
 rewrite <- (fun x => Z.abs_eq (Zpower_nat radix x)).
rewrite <- Zabs_Zmult.
apply Z.le_trans with (1 := H'4); auto with zarith.
rewrite Zabs_Zmult.
apply Z.le_trans with (Zpower_nat radix precision * Z.abs (Fnum q))%Z.
apply Zle_Zmult_comp_r; auto with zarith.
apply Zlt_le_weak; rewrite <- pGivesBound; apply Fp.
apply Zle_Zmult_comp_l; auto with zarith.
apply Zlt_le_weak; rewrite <- pGivesBound; apply Fq.
apply Zpower_NR0; omega.
intros x (H'6, H'7).
cut (FtoR radix (Fmult p q) = FtoR radix x).
intros H'8; rewrite H'8.
apply sym_eq; apply (ProjectMax b radix); auto.
rewrite <- H'8; auto.
rewrite Fmult_correct; auto with zarith.
rewrite H'7.
unfold FtoRradix, FtoR in |- *; simpl in |- *.
rewrite powerRZ_add with (n := Z_of_nat precision).
pattern (Fnum p * Fnum q)%Z at 1 in |- *; rewrite H'3.
rewrite plus_IZR; rewrite mult_IZR.
repeat rewrite Zpower_nat_Z_powerRZ; auto with real zarith.
rewrite Z4; simpl;ring.
apply IZR_neq; omega.
cut (ex (fun m : Z => FtoRradix fmax = Float m (Fexp (Fmult p q))));
 [ intros Z5 | idtac ].
2: unfold FtoRradix in |- *;
    apply
     RoundedModeRep
      with (b := b) (precision := precision) (P := isMax b radix);
    auto.
2: apply MaxRoundedModeP with (precision := precision); auto.
2: rewrite (Fmult_correct radix); auto with zarith.
elim Z5; intros m E; clear Z5.
exists (Fopp (Fminus radix (Float m (Fexp (Fmult p q))) (Fmult p q))).
split.
rewrite E; unfold FtoRradix in |- *; repeat rewrite Fopp_correct;
 repeat rewrite Fminus_correct; repeat rewrite Fmult_correct;
 auto with zarith; ring.
cut
 (Fexp (Fopp (Fminus radix (Float m (Fexp (Fmult p q))) (Fmult p q))) =
  (Fexp p + Fexp q)%Z); [ intros Z5 | idtac ].
split; auto.
split; [ idtac | rewrite Z5; auto ].
cut (Fmult p q <= fmax)%R;
 [ idtac
 | unfold FtoRradix in |- *; rewrite Fmult_correct; auto; case H'2;
    auto with real zarith; (intros H1 H2; case H2; auto) ].
cut
 (fmax <=
  Z.succ (Zquotient (Fnum p * Fnum q) (Zpower_nat radix precision)) *
  powerRZ radix (precision + (Fexp p + Fexp q)))%R.
rewrite E; repeat rewrite Zmin_n_n; repeat rewrite <- Zminus_diag_reverse;
 repeat rewrite Zpower_nat_O; repeat rewrite Zmult_1_r;
 auto.
unfold Fmult, Fminus, Fplus, Fopp in |- *; simpl in |- *.
repeat rewrite Zmin_n_n; repeat rewrite <- Zminus_diag_reverse;
 repeat rewrite Zpower_nat_O; repeat rewrite Zmult_1_r;
 auto.
intros H1 H2; rewrite Zabs_Zopp; apply Zlt_Zabs_intro.
apply Z.lt_le_trans with 0%Z; auto with zarith.
cut (Fnum p * Fnum q <= m)%Z; auto with zarith.
apply le_IZR;
 apply (Rle_monotony_contra_exp radix) with (z := (Fexp p + Fexp q)%Z);
 auto with zarith.
cut
 (m <=
  Z.succ (Zquotient (Fnum p * Fnum q) (Zpower_nat radix precision)) *
  Zpower_nat radix precision)%Z; [ intros H'9 | idtac ].
apply Z.le_lt_trans with (Zpower_nat radix precision - r)%Z;
 [ idtac | rewrite pGivesBound; auto with zarith ].
replace r with
 (Fnum p * Fnum q -
  Zquotient (Fnum p * Fnum q) (Zpower_nat radix precision) *
  Zpower_nat radix precision)%Z.
replace
 (Zpower_nat radix precision -
  (Fnum p * Fnum q -
   Zquotient (Fnum p * Fnum q) (Zpower_nat radix precision) *
   Zpower_nat radix precision))%Z with
 (Z.succ (Zquotient (Fnum p * Fnum q) (Zpower_nat radix precision)) *
  Zpower_nat radix precision - Fnum p * Fnum q)%Z;
 auto with zarith.
pattern (Fnum p * Fnum q)%Z at 1 in |- *; rewrite H'3; ring.
apply le_IZR;
 apply (Rle_monotony_contra_exp radix) with (z := (Fexp p + Fexp q)%Z);
 auto with zarith.
replace
 (IZR
    (Z.succ (Zquotient (Fnum p * Fnum q) (Zpower_nat radix precision)) *
     Zpower_nat radix precision) * powerRZ radix (Fexp p + Fexp q))%R with
 (Z.succ (Zquotient (Fnum p * Fnum q) (Zpower_nat radix precision)) *
  powerRZ radix (precision + (Fexp p + Fexp q)))%R;
 [ auto | idtac ].
rewrite powerRZ_add.
repeat rewrite mult_IZR; repeat rewrite Zpower_nat_Z_powerRZ.
ring.
apply IZR_neq; omega.
case
 (FboundedMbound _ radixMoreThanOne b precision)
  with
    (z := (precision + (Fexp p + Fexp q))%Z)
    (m := Z.succ (Zquotient (Fnum p * Fnum q) (Zpower_nat radix precision)));
 auto with zarith.
rewrite Z.abs_eq; auto with zarith.
apply Zlt_le_succ.
case (Zle_lt_or_eq _ _ multPos); intros Eq1.
cut (0 < Z.abs (Fnum p))%Z; [ intros Eq2 | idtac ].
cut (0 < Z.abs (Fnum q))%Z; [ intros Eq3 | idtac ].
apply Zlt_mult_simpl_l with (c := Zpower_nat radix precision);
 auto with zarith.

rewrite (fun x y z => Zmult_comm x (Zquotient y z)).
apply Z.le_lt_trans with (Fnum p * Fnum q)%Z.
rewrite Z.abs_eq in H'4; auto with zarith; rewrite Z.abs_eq in H'4;
 auto with zarith.
cut (Z.abs (Fnum q) < Zpower_nat radix precision)%Z;
 [ intros Eq4| rewrite <- pGivesBound; apply Fq ]; auto with zarith.
cut (Z.abs (Fnum p) < Zpower_nat radix precision)%Z;
 [ intros Eq4' | rewrite <- pGivesBound; case Fp ]; auto with zarith.
apply Z.le_lt_trans with (Z.abs (Fnum p*Fnum q)).
apply Zle_Zabs.
rewrite Zabs_Zmult.
apply Z.lt_trans with (Z.abs (Fnum p) * Zpower_nat radix precision)%Z.
apply Zmult_gt_0_lt_compat_l; auto with zarith.
apply Zmult_gt_0_lt_compat_r; auto with zarith.
case (Zle_lt_or_eq _ _ (Zle_ZERO_Zabs (Fnum q))); auto.
intros Eq3; contradict Eq1; replace (Fnum q) with 0%Z; auto with zarith.
case (Zle_lt_or_eq _ _ (Zle_ZERO_Zabs (Fnum p))); auto.
intros Eq3; contradict Eq1; replace (Fnum p) with 0%Z; auto with zarith.
rewrite <- Eq1; replace (Zquotient 0 (Zpower_nat radix precision)) with 0%Z;
 auto with zarith.
intros f1 (Hf1, Hf2); rewrite <- Hf2.
case H'2; intros L1 (L2, L3); apply L3; auto.
rewrite Hf2; unfold Fmult, FtoRradix, FtoR in |- *.
replace
 (Fnum p * powerRZ radix (Fexp p) * (Fnum q * powerRZ radix (Fexp q)))%R with
 (Fnum p * Fnum q * powerRZ radix (Fexp p + Fexp q))%R.
replace
 (Z.succ (Zquotient (Fnum p * Fnum q) (Zpower_nat radix precision)) *
  powerRZ radix (precision + (Fexp p + Fexp q)))%R with
 ((Zquotient (Fnum p * Fnum q) (Zpower_nat radix precision) *
   Zpower_nat radix precision + Zpower_nat radix precision)%Z *
  powerRZ radix (Fexp p + Fexp q))%R.
apply Rle_monotone_exp; auto with real zarith.
rewrite <- mult_IZR; apply Rle_IZR.
pattern (Fnum p * Fnum q)%Z at 1 in |- *; rewrite H'3;
 cut (r < Zpower_nat radix precision)%Z; auto with zarith.
rewrite Z.abs_eq in H'5; auto with zarith; rewrite Z.abs_eq in H'5;
 auto with zarith.
unfold Z.succ in |- *; repeat rewrite mult_IZR || rewrite plus_IZR;
 simpl in |- *.
rewrite (powerRZ_add radix precision).
rewrite <- (Zpower_nat_Z_powerRZ radix precision); auto with real zarith;
 ring.
apply IZR_neq; omega.
rewrite powerRZ_add. auto with real zarith; try ring.
apply IZR_neq; omega.
unfold Fopp, Fminus, Fmult in |- *; simpl in |- *; auto.
apply Zmin_n_n.
Qed.

Theorem multExpUpperBound :
 forall P,
 RoundedModeP b radix P ->
 forall p q pq : float,
 P (p * q)%R pq ->
 Fbounded b p ->
 Fbounded b q ->
 (- dExp b <= Fexp p + Fexp q)%Z ->
 exists r : float,
   Fbounded b r /\ r = pq :>R /\ (Fexp r <= precision + (Fexp p + Fexp q))%Z.
intros P H' p q pq H'0 H'1 H'2 H'3.
replace (precision + (Fexp p + Fexp q))%Z with
 (Fexp (Float (pPred (vNum b)) (precision + (Fexp p + Fexp q))));
 [ idtac | simpl in |- *; auto ].
unfold FtoRradix in |- *; apply eqExpMax; auto.
apply RoundedModeBounded with (radix := radix) (P := P) (r := (p * q)%R);
 auto; auto.
unfold pPred in |- *; apply maxFbounded; auto.
apply Z.le_trans with (1 := H'3); auto with zarith.
replace (FtoR radix (Float (pPred (vNum b)) (precision + (Fexp p + Fexp q))))
 with (radix * Float (pPred (vNum b)) (pred precision + (Fexp p + Fexp q)))%R.
rewrite Fabs_correct; auto with zarith.
unfold FtoRradix in |- *;
 apply
  RoundedModeMultAbs
   with (b := b) (precision := precision) (P := P) (r := (p * q)%R);
 auto.
unfold pPred in |- *; apply maxFbounded; auto with zarith.
rewrite Rabs_mult; auto.
apply
 Rle_trans
  with
    (FtoR radix
       (Fmult (Float (pPred (vNum b)) (Fexp p))
          (Float (pPred (vNum b)) (Fexp q)))).
rewrite Fmult_correct; auto with arith.
apply Rmult_le_compat; auto with real; try apply Rabs_pos.
rewrite <- (Fabs_correct radix); try apply maxMax1; auto with zarith.
rewrite <- (Fabs_correct radix); try apply maxMax1; auto with zarith.
unfold Fmult, FtoR in |- *; simpl in |- *; auto.
rewrite powerRZ_add with (n := Z_of_nat (pred precision));
 auto with real arith.
repeat rewrite <- Rmult_assoc.
repeat rewrite (fun (z : Z) (x : R) => Rmult_comm x (powerRZ radix z));
 auto.
apply Rmult_le_compat_l; auto with real arith.
apply powerRZ_le, Rlt_IZR; omega.
rewrite <- Rmult_assoc.
rewrite (fun x : R => Rmult_comm x radix).
rewrite <- powerRZ_Zs; auto with real arith.
replace (Z.succ (pred precision)) with (Z_of_nat precision).
rewrite mult_IZR; auto.
apply Rmult_le_compat; auto with real arith.
apply Rle_IZR, Zle_Zpred; auto with zarith.
apply Rle_IZR, Zle_Zpred; auto with zarith.
unfold pPred in |- *; rewrite pGivesBound; rewrite <- Zpower_nat_Z_powerRZ;
 apply Rle_IZR; auto with real zarith.
rewrite inj_pred; auto with arith zarith.
apply IZR_neq; auto with real zarith.
apply IZR_neq; auto with real zarith.
unfold FtoRradix, FtoR in |- *; simpl in |- *.
repeat rewrite (Rmult_comm (pPred (vNum b))).
rewrite <- Rmult_assoc.
rewrite <- powerRZ_Zs.
rewrite inj_pred; auto with arith zarith.
replace (Z.succ (Z.pred precision + (Fexp p + Fexp q))) with
 (precision + (Fexp p + Fexp q))%Z; auto; unfold Z.succ, Z.pred in |- *;
 ring.
apply IZR_neq; auto with real zarith.
Qed.

Theorem errorBoundedMultPos :
 forall P,
 RoundedModeP b radix P ->
 forall p q f : float,
 Fbounded b p ->
 Fbounded b q ->
 (0 <= p)%R ->
 (0 <= q)%R ->
 (- dExp b <= Fexp p + Fexp q)%Z ->
 P (p * q)%R f ->
 exists r : float,
   r = (p * q - f)%R :>R /\ Fbounded b r /\ Fexp r = (Fexp p + Fexp q)%Z.
intros P H p q f H0 H1 H2 H3 H4 H5.
generalize errorBoundedMultMin errorBoundedMultMax; intros H6 H7.
cut (MinOrMaxP b radix P);
 [ intros | case H; intros H'1 (H'2, (H'3, H'4)); auto ].
case (H8 (p * q)%R f); auto.
Qed.

Theorem errorBoundedMultNeg :
 forall P,
 RoundedModeP b radix P ->
 forall p q f : float,
 Fbounded b p ->
 Fbounded b q ->
 (p <= 0)%R ->
 (0 <= q)%R ->
 (- dExp b <= Fexp p + Fexp q)%Z ->
 P (p * q)%R f ->
 exists r : float,
   r = (p * q - f)%R :>R /\ Fbounded b r /\ Fexp r = (Fexp p + Fexp q)%Z.
intros P H p q f H0 H1 H2 H3 H4 H5.
generalize errorBoundedMultMin errorBoundedMultMax; intros H6 H7.
cut (MinOrMaxP b radix P);
 [ intros | case H; intros H'1 (H'2, (H'3, H'4)); auto ].
case (H8 (p * q)%R f); auto; intros H9.
generalize (H7 (Fopp p) q (Fopp f)); intros H12.
lapply H12; auto.
2: now apply oppBounded.
intros H10; clear H12.
lapply H10; auto; intros H12; clear H10.
lapply H12;
 [ intros H10
 | unfold FtoRradix in |- *; rewrite Fopp_correct; auto with real ];
 clear H12.
lapply H10; auto; intros H12; clear H10.
lapply H12; [ intros H10 | simpl in |- *; auto ]; clear H12.
lapply H10; [ intros H12 | idtac ]; clear H10.
2: replace (Fopp p * q)%R with (- (p * q))%R;
    [ apply MinOppMax; auto | idtac ].
2: unfold FtoRradix in |- *; rewrite Fopp_correct; auto with real.
elim H12; intros r H10; clear H12; elim H10; intros H11 H12; clear H10.
elim H12; clear H12; intros H10 H12.
exists (Fopp r); split; [ generalize H11 | split; auto with zarith ].
2: now apply oppBounded.
unfold FtoRradix in |- *; repeat rewrite Fopp_correct; intros H13;
 rewrite H13; ring.
generalize (H6 (Fopp p) q (Fopp f)); intros H12.
lapply H12; auto.
2: now apply oppBounded.
intros H10; clear H12.
lapply H10; auto; intros H12; clear H10.
lapply H12;
 [ intros H10
 | unfold FtoRradix in |- *; rewrite Fopp_correct; auto with real ];
 clear H12.
lapply H10; auto; intros H12; clear H10.
lapply H12; [ intros H10 | simpl in |- *; auto ]; clear H12.
lapply H10; [ intros H12 | idtac ]; clear H10.
2: replace (Fopp p * q)%R with (- (p * q))%R;
    [ apply MaxOppMin; auto | idtac ].
2: unfold FtoRradix in |- *; rewrite Fopp_correct; auto with real.
elim H12; intros r H10; clear H12; elim H10; intros H11 H12; clear H10.
elim H12; clear H12; intros H10 H12.
exists (Fopp r); split; [ generalize H11 | split; auto with zarith ].
unfold FtoRradix in |- *; repeat rewrite Fopp_correct; intros H13;
 rewrite H13; ring.
now apply oppBounded.
Qed.

Theorem errorBoundedMult :
 forall P,
 RoundedModeP b radix P ->
 forall p q f : float,
 Fbounded b p ->
 Fbounded b q ->
 (- dExp b <= Fexp p + Fexp q)%Z ->
 P (p * q)%R f ->
 exists r : float,
   r = (p * q - f)%R :>R /\ Fbounded b r /\ Fexp r = (Fexp p + Fexp q)%Z.
intros P H p q f H0 H1 H2 H3.
case (Rle_or_lt 0 p); intros H4; case (Rle_or_lt 0 q); intros H5.
apply errorBoundedMultPos with P; auto.
replace (Fexp p) with (Fexp (Fopp p)); auto with zarith.
replace (Fexp q) with (Fexp (Fopp q)); auto with zarith.
cut ((p * q)%R = (Fopp p * Fopp q)%R); [ intros H6; rewrite H6 | idtac ].
apply errorBoundedMultNeg with P; auto with real; try apply oppBounded; auto.
unfold FtoRradix in |- *; rewrite Fopp_correct; auto with real.
unfold FtoRradix in |- *; rewrite Fopp_correct; auto with real.
rewrite <- H6; auto.
unfold FtoRradix in |- *; repeat rewrite Fopp_correct; ring.
apply errorBoundedMultNeg with P; auto with real.
replace (Fexp p) with (Fexp (Fopp p)); auto.
replace (Fexp q) with (Fexp (Fopp q)); auto.
cut ((p * q)%R = (Fopp p * Fopp q)%R); [ intros H6; rewrite H6 | idtac ].
apply errorBoundedMultPos with P; try apply oppBounded; auto with real.
unfold FtoRradix in |- *; rewrite Fopp_correct; auto with real.
unfold FtoRradix in |- *; rewrite Fopp_correct; auto with real.
rewrite <- H6; auto.
unfold FtoRradix in |- *; repeat rewrite Fopp_correct; ring.
Qed.

Theorem errorBoundedMultExp_aux :
 forall n1 n2 : Z,
 (Z.abs n1 < Zpos (vNum b))%Z ->
 (Z.abs n2 < Zpos (vNum b))%Z ->
 (exists ny : Z,
    (exists ey : Z,
       (n1 * n2)%R = (ny * powerRZ radix ey)%R :>R /\
       (Z.abs ny < Zpos (vNum b))%Z)) ->
 exists nx : Z,
   (exists ex : Z,
      (n1 * n2)%R = (nx * powerRZ radix ex)%R :>R /\
      (Z.abs nx < Zpos (vNum b))%Z /\
      (0 <= ex)%Z /\ (ex <= precision)%Z).
intros n1 n2 H H0 H1.
case H1; intros ny (ey, (H2, H3)).
case (Zle_or_lt 0 ey); intros Zl1.
case (Zle_or_lt ey precision); intros Zl2.
exists ny; exists ey; repeat (split; auto).
exists (ny * Zpower_nat radix (Z.abs_nat (ey - precision)))%Z;
 exists (Z_of_nat precision); repeat (split; auto with zarith).
replace (IZR (ny * Zpower_nat radix (Z.abs_nat (ey - precision)))) with
 (ny * powerRZ radix (ey - precision))%R.
rewrite Rmult_assoc; rewrite <- powerRZ_add; auto with zarith real.
replace (ey - precision + precision)%Z with ey; [ auto | ring ].
apply IZR_neq; omega.
rewrite mult_IZR.
rewrite Zpower_nat_powerRZ_absolu; auto with real zarith.
rewrite Zabs_Zmult.
apply lt_IZR; apply Rmult_lt_reg_l with (r := powerRZ radix precision);
 auto with real zarith.
apply powerRZ_lt, Rlt_IZR; omega.
repeat rewrite (fun x y => Rmult_comm (powerRZ x y)).
rewrite mult_IZR.
rewrite Rmult_assoc.
rewrite (Z.abs_eq (Zpower_nat radix (Z.abs_nat (ey - precision))));
 auto with zarith.
rewrite Zpower_nat_powerRZ_absolu; auto with real zarith.
rewrite <- powerRZ_add; auto with real zarith.
replace (ey - precision + precision)%Z with ey; [ idtac | ring ].
replace (powerRZ radix precision) with (IZR (Zpos (vNum b)));
 [ idtac
 | rewrite pGivesBound; rewrite <- Zpower_nat_powerRZ_absolu;
    try rewrite absolu_INR; auto with zarith ].
rewrite <- (fun x y => Rabs_pos_eq (powerRZ x y)); auto with real zarith.
rewrite <- Rabs_Zabs; rewrite <- Rabs_mult; rewrite <- H2.
rewrite Rabs_mult; repeat rewrite Rabs_Zabs; auto with real zarith.
case (Zle_lt_or_eq 0 (Z.abs n2)); auto with zarith; intros Z1.
apply Rlt_trans with (Zpos (vNum b) * Z.abs n2)%R;
 auto with real zarith.
apply Rmult_lt_compat_r; apply Rlt_IZR; try easy.
apply Rmult_lt_compat_l; apply Rlt_IZR; try easy.
rewrite <- Z1; auto with real zarith.
replace (Z.abs n1 * 0%Z)%R with (0 * Zpos (vNum b))%R;
 [ auto with real zarith | simpl; ring ].
apply powerRZ_le, Rlt_IZR; omega.
apply IZR_neq; omega.
apply Zpower_NR0; omega.
exists (n1 * n2)%Z; exists 0%Z; repeat (split; auto with zarith).
rewrite mult_IZR; rewrite powerRZ_O; ring.
apply lt_IZR.
rewrite <- Rabs_Zabs; rewrite mult_IZR; rewrite H2.
rewrite Rabs_mult.
apply Rle_lt_trans with (Rabs ny).
pattern (Rabs ny) at 2 in |- *; replace (Rabs ny) with (Rabs ny * 1)%R;
 [ apply Rmult_le_compat_l | ring ]; auto with arith real.
apply Rabs_pos.
rewrite (Rabs_pos_eq (powerRZ radix ey));
 [ idtac | apply powerRZ_le, Rlt_IZR; auto with arith real ].
replace 1%R with (powerRZ radix 0); [ apply Rle_powerRZ; try apply Rlt_IZR | simpl in |- * ];
 auto with real arith zarith.
apply Rle_IZR; omega.
rewrite Rabs_Zabs; try apply Rlt_IZR; auto with real zarith.
Qed.

Theorem errorBoundedMultExpPos :
 forall P,
 RoundedModeP b radix P ->
 forall p q pq : float,
 Fbounded b p ->
 Fbounded b q ->
 (0 <= p)%R ->
 (0 <= q)%R ->
 P (p * q)%R pq ->
 (- dExp b <= Fexp p + Fexp q)%Z ->
 ex
   (fun r : float =>
    ex
      (fun s : float =>
       Fbounded b r /\
       Fbounded b s /\
       r = pq :>R /\
       s = (p * q - r)%R :>R /\
       Fexp s = (Fexp p + Fexp q)%Z :>Z /\
       (Fexp s <= Fexp r)%Z /\ (Fexp r <= precision + (Fexp p + Fexp q))%Z)).
intros P H p q pq H0 H1 H2 H3 H4 H5.
case (multExpUpperBound P H p q pq); auto; intros r (H'1, (H'2, H'3)).
case (Req_dec (p * q - pq) 0); intros H6.
case (Req_dec pq 0); intros H7.
cut (Fbounded b (Fzero (Fexp p + Fexp q))); [ intros Fb1 | idtac ].
exists (Fzero (Fexp p + Fexp q)); exists (Fzero (Fexp p + Fexp q));
 repeat (split; simpl in |- *; auto with zarith).
rewrite H7; unfold FtoRradix in |- *; apply FzeroisReallyZero.
unfold FtoRradix in |- *; rewrite FzeroisReallyZero; rewrite <- H7.
pattern (FtoRradix pq) at 1 in |- *; rewrite H7; auto with real.
repeat (split; auto); simpl in |- *; auto with arith zarith.
case (errorBoundedMultExp_aux (Fnum p) (Fnum q)); auto with real zarith.
apply H0.
apply H1.
exists (Fnum pq); exists (Fexp pq - (Fexp p + Fexp q))%Z; split.
apply Rmult_eq_reg_l with (powerRZ radix (Fexp p + Fexp q));
 auto with real zarith.
repeat rewrite (fun x y => Rmult_comm (powerRZ x y)).
apply trans_eq with (p * q)%R; auto.
rewrite <- (Fmult_correct radix); auto with real zarith;
 unfold FtoRradix, FtoR, Fmult in |- *; simpl in |- *;
 rewrite mult_IZR; auto.
apply trans_eq with (FtoRradix pq); auto with real.
rewrite Rmult_assoc; rewrite <- powerRZ_add; auto with real zarith.
replace (Fexp pq - (Fexp p + Fexp q) + (Fexp p + Fexp q))%Z with (Fexp pq);
 auto; ring.
apply IZR_neq; omega.
apply Rgt_not_eq, powerRZ_lt, Rlt_IZR; omega.
cut (Fbounded b pq); [ intros Z1; case Z1 | idtac ]; auto with real zarith.
apply (RoundedModeBounded b radix P (p * q)); auto.
intros nx (ex, (H'4, (H'5, (H'6, H'7)))).
cut (FtoRradix pq = FtoRradix (Float nx (ex + (Fexp p + Fexp q))) :>R);
 [ intros Eq1 | idtac ].
exists (Float nx (ex + (Fexp p + Fexp q))); exists (Fzero (Fexp p + Fexp q));
 repeat (split; simpl in |- *; auto with real zarith).
rewrite <- Eq1; rewrite H6; apply (FzeroisReallyZero radix).
replace (FtoRradix pq) with (p * q)%R.
unfold FtoRradix in |- *; unfold FtoR in |- *; simpl in |- *.
rewrite powerRZ_add.
2: apply IZR_neq; omega.
repeat rewrite <- Rmult_assoc; rewrite <- H'4; rewrite powerRZ_add.
ring.
apply IZR_neq; omega.
replace (FtoRradix p * FtoRradix q)%R with
 (pq + (FtoRradix p * FtoRradix q - FtoRradix pq))%R;
 [ rewrite H6 | idtac ]; ring.
case (errorBoundedMultPos P H p q pq); auto.
intros s (H'4, (H'5, H'6)).
exists r; exists s; repeat (split; auto with zarith).
rewrite H'2; auto.
apply Zlt_le_weak;
 apply RoundedModeErrorExpStrict with b radix precision P (p * q)%R;
 auto.
cut (CompatibleP b radix P);
 [ intros H'7 | case H; try intros H'7 (H'8, (H'9, H'10)); auto ].
apply H'7 with (p * q)%R pq; auto with real.
fold FtoRradix in |- *; rewrite H'2; auto with real.
fold FtoRradix in |- *; rewrite H'4; auto with real.
Qed.

Theorem errorBoundedMultExp :
 forall P, (RoundedModeP b radix P) ->
 forall p q pq : float,
  (Fbounded b p) -> (Fbounded b q) ->
  (P (p * q)%R pq) ->
   (-(dExp b) <= Fexp p + Fexp q)%Z ->
   exists r : float,
   exists s : float,
      (Fbounded b r) /\ (Fbounded b s) /\
       r = pq :>R /\ s = (p * q - r)%R :>R /\
       (Fexp s = Fexp p + Fexp q)%Z /\
       (Fexp s <= Fexp r)%Z /\
       (Fexp r <= precision + (Fexp p + Fexp q))%Z.
intros P H p q pq H1 H2 H3 H4.
cut (MinOrMaxP b radix P);
 [ intros | case H; intros H'1 (H'2, (H'3, H'4)); auto ].
case H0 with (p*q)%R pq; auto; intros H0'; clear H0.
case (Rle_or_lt 0 p); intros Rl1.
case (Rle_or_lt 0 q); intros Rl2.
apply (errorBoundedMultExpPos P); auto.
case errorBoundedMultExpPos with (isMax b radix) p (Fopp q) (Fopp pq); auto with zarith real.
apply MaxRoundedModeP with precision; auto.
apply oppBounded; easy.
rewrite (Fopp_correct radix); auto with real.
replace (FtoRradix p * FtoRradix (Fopp q))%R with
 (- (FtoRradix p * FtoRradix q))%R; [apply MinOppMax;auto|idtac].
rewrite (Fopp_correct radix);
 fold FtoRradix in |- *; auto with zarith; ring.
intros r (s, (H5, (H6, (H7, (H8, H9))))); exists (Fopp r); exists (Fopp s).
split; try now apply oppBounded.
split; try now apply oppBounded.
split.
generalize H7; unfold FtoRradix; rewrite 2!Fopp_correct.
intros Y; rewrite Y; ring.
split.
generalize H8; unfold FtoRradix; rewrite 3!Fopp_correct.
intros Y; rewrite Y; ring.
split; simpl; try apply H9.
case (Rle_or_lt 0 q); intros Rl2.
case errorBoundedMultExpPos with (isMax b radix) (Fopp p) q (Fopp pq); auto with zarith real.
apply MaxRoundedModeP with precision; auto.
now apply oppBounded.
rewrite (Fopp_correct radix); auto with real.
replace (FtoRradix (Fopp p) * FtoRradix q)%R with
 (- (FtoRradix p * FtoRradix q))%R; [apply MinOppMax;auto|idtac].
rewrite (Fopp_correct radix);
 fold FtoRradix in |- *; auto with zarith; ring.
intros r (s, (H5, (H6, (H7, (H8, H9))))); exists (Fopp r); exists (Fopp s);
 repeat (split; try apply oppBounded; simpl in |- *; auto with real zarith).
repeat rewrite (Fopp_correct radix); auto with zarith; fold FtoRradix in |- *;
 rewrite H7; repeat rewrite (Fopp_correct radix); auto with zarith;
 fold FtoRradix; ring.
repeat rewrite (Fopp_correct radix); auto with zarith; fold FtoRradix in |- *;
 rewrite H8; repeat rewrite (Fopp_correct radix); auto with zarith;
 fold FtoRradix;ring.
case (errorBoundedMultExpPos P H (Fopp p) (Fopp q) pq); auto with zarith real.
apply oppBounded; easy.
apply oppBounded; easy.
rewrite (Fopp_correct radix); auto with real.
rewrite (Fopp_correct radix); auto with real.
replace (FtoRradix (Fopp p) * FtoRradix (Fopp q))%R with
 (FtoRradix p * FtoRradix q)%R; auto; repeat rewrite (Fopp_correct radix);
 fold FtoRradix in |- *; auto with zarith; ring.
intros r (s, (H5, (H6, (H7, (H8, (H9, (H10, H11))))))); exists r; exists s;
 repeat (split; simpl in |- *; auto with real zarith).
fold FtoRradix in |- *; rewrite H8; repeat rewrite (Fopp_correct radix);
 auto with zarith; fold FtoRradix; try ring.
case (Rle_or_lt 0 p); intros Rl1.
case (Rle_or_lt 0 q); intros Rl2.
apply (errorBoundedMultExpPos P); auto.
case errorBoundedMultExpPos with (isMin b radix) p (Fopp q) (Fopp pq);
  auto with zarith real.
apply MinRoundedModeP with precision; auto.
now apply oppBounded.
rewrite (Fopp_correct radix); auto with real.
replace (FtoRradix p * FtoRradix (Fopp q))%R with
 (- (FtoRradix p * FtoRradix q))%R; [apply MaxOppMin;auto|idtac].
rewrite (Fopp_correct radix);
 fold FtoRradix in |- *; auto with zarith; ring.
intros r (s, (H5, (H6, (H7, (H8, H9))))); exists (Fopp r); exists (Fopp s);
 repeat (split; try apply oppBounded; simpl in |- *; auto with real zarith).
repeat rewrite (Fopp_correct radix); auto with zarith; fold FtoRradix in |- *;
 rewrite H7; repeat rewrite (Fopp_correct radix); auto with zarith;
 fold FtoRradix; ring.
repeat rewrite (Fopp_correct radix); auto with zarith; fold FtoRradix in |- *;
 rewrite H8; repeat rewrite (Fopp_correct radix); auto with zarith;
 fold FtoRradix; ring.
case (Rle_or_lt 0 q); intros Rl2.
case errorBoundedMultExpPos with (isMin b radix) (Fopp p) q (Fopp pq); auto with zarith real.
apply MinRoundedModeP with precision; auto.
now apply oppBounded.
rewrite (Fopp_correct radix); auto with real.
replace (FtoRradix (Fopp p) * FtoRradix q)%R with
 (- (FtoRradix p * FtoRradix q))%R; [apply MaxOppMin;auto|idtac].
rewrite (Fopp_correct radix);
 fold FtoRradix in |- *; auto with zarith; ring.
intros r (s, (H5, (H6, (H7, (H8, H9))))); exists (Fopp r); exists (Fopp s);
 repeat (split; try apply oppBounded; simpl in |- *; auto with real zarith).
repeat rewrite (Fopp_correct radix); auto with zarith; fold FtoRradix in |- *;
 rewrite H7; repeat rewrite (Fopp_correct radix); auto with zarith;
 fold FtoRradix; ring.
repeat rewrite (Fopp_correct radix); auto with zarith; fold FtoRradix in |- *;
 rewrite H8; repeat rewrite (Fopp_correct radix); auto with zarith;
 fold FtoRradix;ring.
case (errorBoundedMultExpPos P H (Fopp p) (Fopp q) pq); auto with zarith real.
now apply oppBounded.
now apply oppBounded.
rewrite (Fopp_correct radix); auto with real.
rewrite (Fopp_correct radix); auto with real.
replace (FtoRradix (Fopp p) * FtoRradix (Fopp q))%R with
 (FtoRradix p * FtoRradix q)%R; auto; repeat rewrite (Fopp_correct radix);
 fold FtoRradix in |- *; auto with zarith; ring.
intros r (s, (H5, (H6, (H7, (H8, (H9, (H10, H11))))))); exists r; exists s;
 repeat (split; try apply oppBounded; simpl in |- *; auto with real zarith).
fold FtoRradix in |- *; rewrite H8; repeat rewrite (Fopp_correct radix);
 auto with zarith; fold FtoRradix; ring.
Qed.

End FRoundPP.
Section Fclosest.
Variable b : Fbound.
Variable radix : Z.
Variable precision : nat.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.

Hypothesis radixMoreThanOne : (1 < radix)%Z.
Hypothesis precisionGreaterThanOne : 1 < precision.
Hypothesis pGivesBound : Zpos (vNum b) = Zpower_nat radix precision.

Definition Closest (r : R) (p : float) :=
  Fbounded b p /\
  (forall f : float, Fbounded b f -> (Rabs (p - r) <= Rabs (f - r))%R).

Theorem ClosestTotal : TotalP Closest.
red in |- *; intros r.
case MinEx with (r := r) (3 := pGivesBound); auto with zarith.
intros min H'.
case MaxEx with (r := r) (3 := pGivesBound); auto with zarith.
intros max H'0.
cut (min <= r)%R; [ intros Rl1 | apply isMin_inv1 with (1 := H') ].
cut (r <= max)%R; [ intros Rl2 | apply isMax_inv1 with (1 := H'0) ].
case (Rle_or_lt (Rabs (min - r)) (Rabs (max - r))); intros H'1.
exists min; split.
case H'; auto.
intros f H'2.
case (Rle_or_lt f r); intros H'3.
repeat rewrite Rabs_left1.
apply Ropp_le_contravar; auto.
apply Rplus_le_reg_l with (r := r).
repeat rewrite Rplus_minus; auto.
elim H'; auto.
intros H'4 H'5; elim H'5; intros H'6 H'7; apply H'7; clear H'5; auto.
apply Rplus_le_reg_l with (r := r).
repeat rewrite Rplus_minus; auto.
rewrite Rplus_0_r; auto.
apply Rplus_le_reg_l with (r := r).
repeat rewrite Rplus_minus; auto.
rewrite Rplus_0_r; auto.
apply Rle_trans with (1 := H'1).
repeat rewrite Rabs_right.
apply Rplus_le_reg_l with (r := r).
repeat rewrite Rplus_minus; auto.
elim H'0; auto.
intros H'4 H'5; elim H'5; intros H'6 H'7; apply H'7; clear H'5; auto.
apply Rlt_le; auto.
apply Rle_ge; apply Rplus_le_reg_l with (r := r).
repeat rewrite Rplus_minus; auto.
rewrite Rplus_0_r.
apply Rlt_le; auto.
repeat rewrite Rplus_minus; auto.
apply Rle_ge; apply Rplus_le_reg_l with (r := r).
repeat rewrite Rplus_minus; auto.
rewrite Rplus_0_r; auto.
exists max; split.
case H'0; auto.
intros f H'2.
case (Rle_or_lt f r); intros H'3.
apply Rle_trans with (1 := Rlt_le _ _ H'1).
repeat rewrite Rabs_left1.
apply Ropp_le_contravar; auto.
apply Rplus_le_reg_l with (r := r).
repeat rewrite Rplus_minus; auto.
elim H'; auto.
intros H'4 H'5; elim H'5; intros H'6 H'7; apply H'7; clear H'5; auto.
apply Rplus_le_reg_l with (r := r).
repeat rewrite Rplus_minus; auto.
rewrite Rplus_0_r; auto.
apply Rplus_le_reg_l with (r := r).
repeat rewrite Rplus_minus; auto.
rewrite Rplus_0_r; auto.
repeat rewrite Rabs_right; auto with real.
apply Rplus_le_reg_l with (r := r).
repeat rewrite Rplus_minus; auto.
elim H'0; auto.
intros H'4 H'5; elim H'5; intros H'6 H'7; apply H'7; clear H'5; auto.
apply Rlt_le; auto.
apply Rle_ge; apply Rplus_le_reg_l with (r := r).
repeat rewrite Rplus_minus; auto.
rewrite Rplus_0_r; auto.
apply Rlt_le; auto.
apply Rle_ge; apply Rplus_le_reg_l with (r := r).
repeat rewrite Rplus_minus; auto.
rewrite Rplus_0_r; auto.
Qed.

Theorem ClosestCompatible : CompatibleP b radix Closest.
red in |- *; simpl in |- *.
intros r1 r2 p q H'; case H'.
intros H'0 H'1 H'2 H'3 H'4.
split; auto.
intros f H'5.
unfold FtoRradix in |- *; rewrite <- H'3; rewrite <- H'2; auto.
Qed.

Theorem ClosestMin :
 forall (r : R) (min max : float),
 isMin b radix r min ->
 isMax b radix r max -> (2%nat * r <= min + max)%R -> Closest r min.
intros r min max H' H'0 H'1; split.
case H'; auto.
intros f H'2.
case (Rle_or_lt f r); intros H'3.
repeat rewrite Rabs_left1.
apply Ropp_le_contravar.
apply Rplus_le_reg_l with (r := r).
repeat rewrite Rplus_minus; auto.
case H'; auto.
intros H'4 H'5; elim H'5; intros H'6 H'7; apply H'7; clear H'5; auto.
apply Rplus_le_reg_l with (r := r).
repeat rewrite Rplus_minus; auto.
rewrite Rplus_0_r; auto.
apply Rplus_le_reg_l with (r := r).
repeat rewrite Rplus_minus; auto.
rewrite Rplus_0_r; auto.
apply isMin_inv1 with (1 := H'); auto.
rewrite (Rabs_left1 (min - r)).
rewrite (Rabs_right (f - r)).
apply Rle_trans with (max - r)%R.
cut (forall x y : R, (- y + x)%R = (- (y - x))%R);
 [ intros Eq0; repeat rewrite Eq0; clear Eq0 | intros; ring ].
cut (forall x y : R, (- (x - y))%R = (y - x)%R);
 [ intros Eq0; repeat rewrite Eq0; clear Eq0
 | intros; unfold Rminus in |- *; ring ].
apply Rplus_le_reg_l with (r := FtoR radix min).
repeat rewrite Rplus_minus; auto.
apply Rplus_le_reg_l with (r := r).
replace (r + (FtoR radix min + (max - r)))%R with (min + max)%R.
replace (r + r)%R with (2%nat * r)%R; auto.
simpl in |- *; ring.
simpl in |- *; fold FtoRradix; ring.
apply Rplus_le_reg_l with (r := r).
repeat rewrite Rplus_minus; auto.
case H'0; auto.
intros H'4 H'5; elim H'5; intros H'6 H'7; apply H'7; clear H'5; auto.
apply Rlt_le; auto.
apply Rle_ge; apply Rplus_le_reg_l with (r := r).
repeat rewrite Rplus_minus; auto.
rewrite Rplus_0_r; auto.
apply Rlt_le; auto.
apply Rplus_le_reg_l with (r := r).
repeat rewrite Rplus_minus; auto.
rewrite Rplus_0_r; auto.
apply isMin_inv1 with (1 := H'); auto.
Qed.

Theorem ClosestMax :
 forall (r : R) (min max : float),
 isMin b radix r min ->
 isMax b radix r max -> (min + max <= 2%nat * r)%R -> Closest r max.
intros r min max H' H'0 H'1; split.
case H'0; auto.
intros f H'2.
case (Rle_or_lt f r); intros H'3.
rewrite (Rabs_right (max - r)).
rewrite (Rabs_left1 (f - r)).
apply Rle_trans with (r - min)%R.
apply Rplus_le_reg_l with (r := FtoRradix min).
repeat rewrite Rplus_minus; auto.
apply Rplus_le_reg_l with (r := r).
replace (r + (min + (max - r)))%R with (min + max)%R.
replace (r + r)%R with (2%nat * r)%R; auto.
simpl in |- *; ring.
simpl in |- *; ring.
replace (r - min)%R with (- (min - r))%R.
apply Ropp_le_contravar.
apply Rplus_le_reg_l with (r := r).
repeat rewrite Rplus_minus; auto.
case H'; auto.
intros H'4 H'5; elim H'5; intros H'6 H'7; apply H'7; clear H'5; auto.
simpl in |- *; ring.
apply Rplus_le_reg_l with (r := r).
repeat rewrite Rplus_minus; auto.
rewrite Rplus_0_r; auto.
apply Rle_ge; apply Rplus_le_reg_l with (r := r).
repeat rewrite Rplus_minus; auto.
rewrite Rplus_0_r; auto.
apply isMax_inv1 with (1 := H'0); auto.
repeat rewrite Rabs_right.
apply Rplus_le_reg_l with (r := r).
repeat rewrite Rplus_minus; auto.
case H'0; auto.
intros H'4 H'5; elim H'5; intros H'6 H'7; apply H'7; clear H'5; auto.
apply Rlt_le; auto.
apply Rle_ge; apply Rplus_le_reg_l with (r := r).
repeat rewrite Rplus_minus; auto.
rewrite Rplus_0_r; auto.
apply Rlt_le; auto.
apply Rle_ge; apply Rplus_le_reg_l with (r := r).
repeat rewrite Rplus_minus; auto.
rewrite Rplus_0_r; auto.
apply isMax_inv1 with (1 := H'0); auto.
Qed.

Theorem ClosestMinOrMax : MinOrMaxP b radix Closest.
red in |- *.
intros r p H'.
case (Rle_or_lt p r); intros H'1.
left; split; auto.
case H'; auto.
split; auto.
intros f H'0 H'2.
apply Rplus_le_reg_l with (r := (- r)%R).
cut (forall x y : R, (- y + x)%R = (- (y - x))%R);
 [ intros Eq0; repeat rewrite Eq0; clear Eq0 | intros; ring ].
apply Ropp_le_contravar.
rewrite <- (Rabs_right (r - FtoR radix p)).
rewrite <- (Rabs_right (r - FtoR radix f)).
cut (forall x y : R, Rabs (x - y) = Rabs (y - x));
 [ intros Eq0; repeat rewrite (Eq0 r); clear Eq0
 | intros x y; rewrite <- (Rabs_Ropp (x - y)); rewrite Ropp_minus_distr ];
 auto.
elim H'; auto.
apply Rle_ge; apply Rplus_le_reg_l with (r := FtoR radix f).
repeat rewrite Rplus_minus; auto.
rewrite Rplus_0_r; auto.
apply Rle_ge; apply Rplus_le_reg_l with (r := FtoR radix p).
repeat rewrite Rplus_minus; auto.
rewrite Rplus_0_r; auto.
right; split; auto.
case H'; auto.
split; auto.
apply Rlt_le; auto.
intros f H'0 H'2.
apply Rplus_le_reg_l with (r := (- r)%R).
cut (forall x y : R, (- y + x)%R = (- (y - x))%R);
 [ intros Eq0; repeat rewrite Eq0; clear Eq0 | intros; ring ].
rewrite <- (Rabs_left1 (r - FtoR radix p)).
rewrite <- (Rabs_left1 (r - FtoR radix f)).
cut (forall x y : R, Rabs (x - y) = Rabs (y - x));
 [ intros Eq0; repeat rewrite (Eq0 r); clear Eq0
 | intros x y; rewrite <- (Rabs_Ropp (x - y)); rewrite Ropp_minus_distr ];
 auto.
elim H'; auto.
apply Rplus_le_reg_l with (r := FtoR radix f).
repeat rewrite Rplus_minus; auto.
rewrite Rplus_0_r; auto.
apply Rplus_le_reg_l with (r := FtoR radix p).
repeat rewrite Rplus_minus; auto.
rewrite Rplus_0_r; auto.
apply Rlt_le; auto.
Qed.

Theorem ClosestMinEq :
 forall (r : R) (min max p : float),
 isMin b radix r min ->
 isMax b radix r max ->
 (2%nat * r < min + max)%R -> Closest r p -> p = min :>R.
intros r min max p H' H'0 H'1 H'2.
case (ClosestMinOrMax r p); auto; intros H'3.
unfold FtoRradix in |- *; apply MinEq with (1 := H'3); auto.
absurd (Rabs (max - r) <= Rabs (min - r))%R.
apply Rgt_not_le.
rewrite (Rabs_left1 (min - r)).
rewrite Rabs_right.
replace (- (min - r))%R with (r - min)%R; [ idtac | ring ].
red in |- *; apply Rplus_lt_reg_l with (r := FtoRradix min).
repeat rewrite Rplus_minus; auto.
apply Rplus_lt_reg_l with (r := r).
replace (r + r)%R with (2%nat * r)%R; [ idtac | simpl in |- *; ring ].
replace (r + (min + (max - r)))%R with (min + max)%R; [ idtac | ring ]; auto.
apply Rle_ge; apply Rplus_le_reg_l with (r := r).
repeat rewrite Rplus_minus; auto.
rewrite Rplus_0_r; auto.
apply isMax_inv1 with (1 := H'0); auto.
apply Rplus_le_reg_l with (r := r).
repeat rewrite Rplus_minus; auto.
rewrite Rplus_0_r; auto.
apply isMin_inv1 with (1 := H'); auto.
cut (Closest r max).
intros H'4; case H'4.
intros H'5 H'6; apply H'6; auto.
case H'; auto.
apply ClosestCompatible with (1 := H'2); auto.
apply MaxEq with (1 := H'3); auto.
case H'0; auto.
Qed.

Theorem ClosestMaxEq :
 forall (r : R) (min max p : float),
 isMin b radix r min ->
 isMax b radix r max ->
 (min + max < 2%nat * r)%R -> Closest r p -> p = max :>R.
intros r min max p H' H'0 H'1 H'2.
case (ClosestMinOrMax r p); auto; intros H'3.
2: unfold FtoRradix in |- *; apply MaxEq with (1 := H'3); auto.
absurd (Rabs (min - r) <= Rabs (max - r))%R.
apply Rgt_not_le.
rewrite (Rabs_left1 (min - r)).
rewrite Rabs_right.
replace (- (min - r))%R with (r - min)%R; [ idtac | ring ].
red in |- *; apply Rplus_lt_reg_l with (r := FtoRradix min).
repeat rewrite Rplus_minus; auto.
apply Rplus_lt_reg_l with (r := r).
replace (r + r)%R with (2%nat * r)%R; [ idtac | simpl in |- *; ring ].
replace (r + (min + (max - r)))%R with (min + max)%R; [ idtac | ring ]; auto.
apply Rle_ge; apply Rplus_le_reg_l with (r := r).
repeat rewrite Rplus_minus; auto.
rewrite Rplus_0_r; auto.
apply isMax_inv1 with (1 := H'0); auto.
apply Rplus_le_reg_l with (r := r).
repeat rewrite Rplus_minus; auto.
rewrite Rplus_0_r; auto.
apply isMin_inv1 with (1 := H'); auto.
cut (Closest r min).
intros H'4; case H'4.
intros H'5 H'6; apply H'6; auto.
case H'0; auto.
apply ClosestCompatible with (1 := H'2); auto.
apply MinEq with (1 := H'3); auto.
case H'; auto.
Qed.

Theorem ClosestMonotone : MonotoneP radix Closest.
red in |- *; simpl in |- *.
intros p q p' q' H' H'0 H'1.
change (p' <= q')%R in |- *.
case (Rle_or_lt p p'); intros Rl0.
case (Rle_or_lt p q'); intros Rl1.
apply Rplus_le_reg_l with (r := (- p)%R).
cut (forall x y : R, (- y + x)%R = (- (y - x))%R);
 [ intros Eq0; repeat rewrite Eq0; clear Eq0 | intros; ring ].
rewrite <- (Rabs_left1 (p - p')).
rewrite <- (Rabs_left1 (p - q')).
cut (forall x y : R, Rabs (x - y) = Rabs (y - x));
 [ intros Eq0; repeat rewrite (Eq0 p); clear Eq0
 | intros x y; rewrite <- (Rabs_Ropp (x - y)); rewrite Ropp_minus_distr ];
 auto.
elim H'0; auto.
intros H'2 H'3; apply H'3; auto.
case H'1; auto.
apply Rplus_le_reg_l with (r := FtoR radix q').
repeat rewrite Rplus_minus; auto.
rewrite Rplus_0_r; auto.
apply Rplus_le_reg_l with (r := FtoR radix p').
repeat rewrite Rplus_minus; auto.
rewrite Rplus_0_r; auto.
case (Rle_or_lt p' q); intros Rl2.
apply Rplus_le_reg_l with (r := (- q)%R).
cut (forall x y : R, (- y + x)%R = (- (y - x))%R);
 [ intros Eq0; repeat rewrite Eq0; clear Eq0 | intros; ring ].
rewrite <- (Rabs_right (q - p')).
2: apply Rle_ge; apply Rplus_le_reg_l with (r := FtoR radix p').
2: repeat rewrite Rplus_minus; auto.
2: rewrite Rplus_0_r; auto.
rewrite <- (Rabs_right (q - q')).
2: apply Rle_ge; apply Rplus_le_reg_l with (r := FtoR radix q').
2: repeat rewrite Rplus_minus; auto.
2: rewrite Rplus_0_r; auto.
2: apply Rle_trans with (1 := Rlt_le _ _ Rl1); apply Rlt_le; auto.
cut (forall x y : R, Rabs (x - y) = Rabs (y - x));
 [ intros Eq0; repeat rewrite (Eq0 q); clear Eq0
 | intros x y; rewrite <- (Rabs_Ropp (x - y)); rewrite Ropp_minus_distr ];
 auto.
apply Ropp_le_contravar.
elim H'1; auto.
intros H'2 H'3; apply H'3; auto.
case H'0; auto.
case (Rle_or_lt (p - q') (p' - q)); intros Rl3.
absurd (Rabs (p' - p) <= Rabs (q' - p))%R.
apply Rgt_not_le.
rewrite (Rabs_left1 (q' - p)).
2: apply Rplus_le_reg_l with (r := p).
2: repeat rewrite Rplus_minus; auto.
2: rewrite Rplus_0_r; apply Rlt_le; auto.
rewrite (Rabs_right (p' - p)).
2: apply Rle_ge; apply Rplus_le_reg_l with (r := p).
2: rewrite Rplus_0_r; auto.
cut (forall x y : R, (- (y - x))%R = (x - y)%R);
 [ intros Eq0; repeat rewrite Eq0; clear Eq0 | intros; ring ].
red in |- *; apply Rle_lt_trans with (1 := Rl3).
replace (p' - p)%R with (p' - q + (q - p))%R.
pattern (p' - q)%R at 1 in |- *; replace (p' - q)%R with (p' - q + 0)%R.
apply Rplus_lt_compat_l; auto.
apply Rplus_lt_reg_l with (r := p).
repeat rewrite Rplus_minus; auto.
rewrite Rplus_0_r; auto.
rewrite Rplus_0_r; auto.
ring.
replace (p + (p' - p))%R with (FtoRradix p'); auto; ring.
case H'0; intros H'2 H'3; apply H'3; auto.
case H'1; auto.
absurd (Rabs (q' - q) <= Rabs (p' - q))%R.
apply Rgt_not_le.
rewrite (Rabs_left1 (q' - q)).
2: apply Rplus_le_reg_l with (r := q).
2: repeat rewrite Rplus_minus; auto.
2: rewrite Rplus_0_r; apply Rlt_le; auto.
2: apply Rlt_trans with (1 := Rl1); auto.
rewrite (Rabs_right (p' - q)).
2: apply Rle_ge; apply Rplus_le_reg_l with (r := q).
2: rewrite Rplus_0_r; apply Rlt_le; auto.
cut (forall x y : R, (- (y - x))%R = (x - y)%R);
 [ intros Eq0; repeat rewrite Eq0; clear Eq0 | intros; ring ].
red in |- *; apply Rlt_trans with (1 := Rl3).
replace (q - q')%R with (p - q' + (q - p))%R.
pattern (p - q')%R at 1 in |- *; replace (p - q')%R with (p - q' + 0)%R.
apply Rplus_lt_compat_l; auto.
apply Rplus_lt_reg_l with (r := p).
repeat rewrite Rplus_minus; auto.
rewrite Rplus_0_r; auto.
rewrite Rplus_0_r; auto.
ring.
replace (q + (p' - q))%R with (FtoRradix p'); auto; ring.
case H'1; intros H'2 H'3; apply H'3; auto.
case H'0; auto.
case (Rle_or_lt p q'); intros Rl1.
apply Rle_trans with p; auto.
apply Rlt_le; auto.
apply Rplus_le_reg_l with (r := (- q)%R).
cut (forall x y : R, (- y + x)%R = (- (y - x))%R);
 [ intros Eq0; repeat rewrite Eq0; clear Eq0 | intros; ring ].
rewrite <- (Rabs_right (q - p')).
rewrite <- (Rabs_right (q - q')).
apply Ropp_le_contravar.
cut (forall x y : R, Rabs (x - y) = Rabs (y - x));
 [ intros Eq0; repeat rewrite (Eq0 q); clear Eq0
 | intros x y; rewrite <- (Rabs_Ropp (x - y)); rewrite Ropp_minus_distr ];
 auto.
elim H'1; auto.
intros H'2 H'3; apply H'3; auto.
case H'0; auto.
apply Rle_ge; apply Rplus_le_reg_l with (r := FtoR radix q').
repeat rewrite Rplus_minus; auto.
rewrite Rplus_0_r; auto.
apply Rlt_le; apply Rlt_trans with (1 := Rl1); auto.
apply Rle_ge; apply Rplus_le_reg_l with (r := FtoR radix p').
repeat rewrite Rplus_minus; auto.
rewrite Rplus_0_r; auto.
apply Rlt_le; apply Rlt_trans with (1 := Rl0); auto.
Qed.

Theorem ClosestRoundedModeP : RoundedModeP b radix Closest.
split; try exact ClosestTotal.
split; try exact ClosestCompatible.
split; try exact ClosestMinOrMax.
try exact ClosestMonotone.
Qed.

Definition EvenClosest (r : R) (p : float) :=
  Closest r p /\
  (FNeven b radix precision p \/ (forall q : float, Closest r q -> q = p :>R)).

Theorem EvenClosestTotal : TotalP EvenClosest.
red in |- *; intros r.
case MinEx with (r := r) (3 := pGivesBound); auto with zarith.
intros min H'.
case MaxEx with (r := r) (3 := pGivesBound); auto with zarith.
intros max H'0.
cut (min <= r)%R; [ intros Rl1 | apply isMin_inv1 with (1 := H'); auto ].
cut (r <= max)%R; [ intros Rl2 | apply isMax_inv1 with (1 := H'0) ].
case (Rle_or_lt (r - min) (max - r)); intros H'1.
case H'1; intros H'2; auto.
exists min; split.
apply ClosestMin with (max := max); auto.
replace (2%nat * r)%R with (r + r)%R; [ idtac | simpl in |- *; ring ].
apply Rminus_le; auto.
replace (r + r - (min + max))%R with (r - min - (max - r))%R;
 [ idtac | simpl in |- *; ring ].
apply Rle_minus; auto.
right; intros q H'3.
apply ClosestMinEq with (r := r) (max := max); auto.
replace (2%nat * r)%R with (r + r)%R; [ idtac | simpl in |- *; ring ].
apply Rminus_lt; auto.
replace (r + r - (min + max))%R with (r - min - (max - r))%R;
 [ idtac | simpl in |- *; ring ].
apply Rlt_minus; auto.
case (FNevenOrFNodd b radix precision min); intros Ev0.
exists min; split; auto.
apply ClosestMin with (max := max); auto.
replace (2%nat * r)%R with (r + r)%R; [ idtac | simpl in |- *; ring ].
apply Rminus_le; auto.
replace (r + r - (min + max))%R with (r - min - (max - r))%R;
 [ idtac | simpl in |- *; ring ].
apply Rle_minus; auto.
exists max; split; auto.
apply ClosestMax with (min := min); auto.
replace (2%nat * r)%R with (r + r)%R; [ idtac | simpl in |- *; ring ].
apply Rminus_le; auto.
replace (min + max - (r + r))%R with (max - r - (r - min))%R;
 [ idtac | simpl in |- *; ring ].
apply Rle_minus; auto.
rewrite H'2; auto with real.
case (Req_dec min max); intros H'5.
right; intros q H'3.
case (ClosestMinOrMax _ _ H'3); intros isM0.
rewrite <- H'5.
apply MinEq with (1 := isM0); auto.
apply MaxEq with (1 := isM0); auto.
left.
apply FNevenEq with (f1 := FNSucc b radix precision min); auto.
apply FcanonicBound with (radix := radix).
apply FNSuccCanonic; auto with zarith.
case H'; auto.
case H'0; auto.
apply MaxEq with (b := b) (r := r); auto.
apply MinMax; auto with zarith.
contradict H'5; auto.
fold FtoRradix in H'5; rewrite H'5 in H'2.
replace (FtoRradix max) with (min + (max - min))%R;
 [ rewrite <- H'2 | idtac ]; ring.
apply FNoddSuc; auto.
case H'; auto.
exists max; split; auto.
apply ClosestMax with (min := min); auto.
replace (2%nat * r)%R with (r + r)%R; [ idtac | simpl in |- *; ring ].
apply Rminus_le; auto.
replace (min + max - (r + r))%R with (max - r - (r - min))%R;
 [ idtac | simpl in |- *; ring ].
apply Rle_minus; auto with real.
right; intros q H'2.
apply ClosestMaxEq with (r := r) (min := min); auto.
replace (2%nat * r)%R with (r + r)%R; [ idtac | simpl in |- *; ring ].
apply Rminus_lt; auto.
replace (min + max - (r + r))%R with (max - r - (r - min))%R;
 [ idtac | simpl in |- *; ring ].
apply Rlt_minus; auto.
Qed.

Theorem EvenClosestCompatible : CompatibleP b radix EvenClosest.
red in |- *; simpl in |- *.
intros r1 r2 p q H' H'0 H'1 H'2; red in |- *.
inversion H'.
split.
apply (ClosestCompatible r1 r2 p q); auto.
case H0; intros H1.
left.
apply FNevenEq with (f1 := p); auto.
case H; auto.
right; intros q0 H'3.
unfold FtoRradix in |- *; rewrite <- H'1; auto.
apply H1; auto.
apply (ClosestCompatible r2 r1 q0 q0); auto.
case H'3; auto.
Qed.

Theorem EvenClosestMinOrMax : MinOrMaxP b radix EvenClosest.
red in |- *; intros r p H'; case (ClosestMinOrMax r p); auto.
case H'; auto.
Qed.

Theorem EvenClosestMonotone : MonotoneP radix EvenClosest.
red in |- *; simpl in |- *; intros p q p' q' H' H'0 H'1.
apply (ClosestMonotone p q); auto; case H'0; case H'1; auto.
Qed.

Theorem EvenClosestRoundedModeP : RoundedModeP b radix EvenClosest.
red in |- *; split.
exact EvenClosestTotal.
split.
exact EvenClosestCompatible.
split.
exact EvenClosestMinOrMax.
exact EvenClosestMonotone.
Qed.

Theorem EvenClosestUniqueP : UniqueP radix EvenClosest.
red in |- *; simpl in |- *.
intros r p q H' H'0.
inversion H'; inversion H'0; case H0; case H2; auto.
intros H'1 H'2; case (EvenClosestMinOrMax r p);
 case (EvenClosestMinOrMax r q); auto.
intros H'3 H'4; apply (MinUniqueP b radix r); auto.
intros H'3 H'4; case (Req_dec p q); auto; intros H'5.
contradict H'1; auto.
apply FnOddNEven; auto.
apply FNoddEq with (f1 := FNSucc b radix precision p); auto.
apply FcanonicBound with (radix := radix); auto.
apply FNSuccCanonic; auto with zarith.
case H'4; auto.
case H'3; auto.
apply (MaxUniqueP b radix r); auto.
apply MinMax; auto with zarith.
contradict H'5; auto.
apply
 (RoundedProjector b radix _
    (MaxRoundedModeP _ _ _ radixMoreThanOne precisionGreaterThanOne
       pGivesBound)); auto.
case H'4; auto.
rewrite <- H'5; auto.
apply FNevenSuc; auto.
case H'4; auto.
intros H'3 H'4; case (Req_dec p q); auto; intros H'5.
contradict H'2; auto.
apply FnOddNEven; auto.
apply FNoddEq with (f1 := FNSucc b radix precision q); auto.
apply FcanonicBound with (radix := radix); auto.
apply FNSuccCanonic; auto with zarith.
case H'3; auto.
case H'4; auto.
apply (MaxUniqueP b radix r); auto.
apply MinMax; auto with zarith.
contradict H'5; auto.
apply sym_eq;
 apply
  (RoundedProjector b radix _
     (MaxRoundedModeP _ _ _ radixMoreThanOne precisionGreaterThanOne
        pGivesBound)); auto.
case H'3; auto.
rewrite <- H'5; auto.
apply FNevenSuc; auto.
case H'3; auto.
intros H'3 H'4; apply (MaxUniqueP b radix r); auto.
intros H'1 H'2; apply sym_eq; auto.
Qed.

Theorem ClosestSymmetric : SymmetricP Closest.
red in |- *; intros r p H'; case H'; clear H'.
intros H' H'0; split.
apply oppBounded; auto.
intros f H'1.
replace (Rabs (Fopp p - - r)) with (Rabs (p - r)).
replace (Rabs (f - - r)) with (Rabs (Fopp f - r)).
apply H'0; auto.
apply oppBounded; auto.
unfold FtoRradix in |- *; rewrite Fopp_correct.
pattern r at 1 in |- *; replace r with (- - r)%R; [ idtac | ring ].
replace (- FtoR radix f - - - r)%R with (- (FtoR radix f - - r))%R;
 [ idtac | ring ].
apply Rabs_Ropp; auto.
unfold FtoRradix in |- *; rewrite Fopp_correct.
replace (- FtoR radix p - - r)%R with (- (FtoR radix p - r))%R;
 [ idtac | ring ].
apply sym_eq; apply Rabs_Ropp.
Qed.

Theorem EvenClosestSymmetric : SymmetricP EvenClosest.
red in |- *; intros r p H'; case H'; clear H'.
intros H' H'0; case H'0; clear H'0; intros H'0.
split; auto.
apply (ClosestSymmetric r p); auto.
left.
apply FNevenFop; auto.
split; auto.
apply (ClosestSymmetric r p); auto.
right.
intros q H'1.
cut (Fopp q = p :>R).
intros H'2; unfold FtoRradix in |- *; rewrite Fopp_correct.
unfold FtoRradix in H'2; rewrite <- H'2.
rewrite Fopp_correct; ring.
apply H'0; auto.
replace r with (- - r)%R; [ idtac | ring ].
apply (ClosestSymmetric (- r)%R q); auto.
Qed.
End Fclosest.

Section Fclosestp2.
Variable b : Fbound.
Variable radix : Z.
Variable precision : nat.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.
Hypothesis radixMoreThanOne : (1 < radix)%Z.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ radixMoreThanOne).

Hypothesis precisionGreaterThanOne : 1 < precision.
Hypothesis pGivesBound : Zpos (vNum b) = Zpower_nat radix precision.

Theorem ClosestOpp :
 forall (p : float) (r : R),
 Closest b radix r p -> Closest b radix (- r) (Fopp p).
intros p r H'; split.
apply oppBounded; auto.
case H'; auto.
intros f H'0.
rewrite Fopp_correct.
replace (- FtoR radix p - - r)%R with (- (FtoR radix p - r))%R;
 [ idtac | ring ].
replace (FtoR radix f - - r)%R with (- (- FtoR radix f - r))%R;
 [ idtac | ring ].
rewrite <- Fopp_correct.
repeat rewrite Rabs_Ropp.
apply H', oppBounded; easy.
Qed.

Theorem ClosestFabs :
 forall (p : float) (r : R),
 Closest b radix r p -> Closest b radix (Rabs r) (Fabs p).
intros p r H'; case (Rle_or_lt 0 r); intros Rl0.
rewrite Rabs_right; auto with real.
replace (Fabs p) with p; auto.
unfold Fabs in |- *; apply floatEq; simpl in |- *; auto.
cut (0 <= Fnum p)%Z.
case (Fnum p); simpl in |- *; auto; intros p' H0; contradict H0;
 apply Zlt_not_le; red in |- *; simpl in |- *; auto with zarith.
apply LeR0Fnum with (radix := radix); auto.
apply
 RleRoundedR0
  with (b := b) (precision := precision) (P := Closest b radix) (r := r);
 auto.
apply ClosestRoundedModeP with (precision := precision); auto with real.
rewrite Rabs_left1; auto.
replace (Fabs p) with (Fopp p).
apply ClosestOpp; auto.
unfold Fabs in |- *; apply floatEq; simpl in |- *; auto.
cut (Fnum p <= 0)%Z.
case (Fnum p); simpl in |- *; auto; intros p' H0; contradict H0;
 apply Zlt_not_le; red in |- *; simpl in |- *; auto with zarith.
apply R0LeFnum with (radix := radix); auto.
apply
 RleRoundedLessR0
  with (b := b) (precision := precision) (P := Closest b radix) (r := r);
 auto.
apply ClosestRoundedModeP with (precision := precision); auto.
apply Rlt_le; auto.
apply Rlt_le; auto.
Qed.

Theorem ClosestUlp :
 forall (p : R) (q : float),
 Closest b radix p q -> (2%nat * Rabs (p - q) <= Fulp b radix precision q)%R.
intros p q H'.
case (Req_dec p q); intros Eqpq.
rewrite Eqpq.
replace (Rabs (q - q)) with 0%R;
 [ rewrite Rmult_0_r
 | replace (q - q)%R with 0%R; try ring; rewrite Rabs_right; auto with real ].
unfold Fulp in |- *; apply Rlt_le; auto with real arith.
apply powerRZ_lt, Rlt_IZR; omega.
replace (2%nat * Rabs (p - q))%R with (Rabs (p - q) + Rabs (p - q))%R;
 [ idtac | simpl in |- *; ring ].
case ClosestMinOrMax with (1 := H'); intros H'1.
apply Rle_trans with (Rabs (p - q) + Rabs (FNSucc b radix precision q - p))%R.
apply Rplus_le_compat_l.
rewrite <- (Rabs_Ropp (p - q)).
rewrite Ropp_minus_distr.
elim H'; auto.
intros H'0 H'2; apply H'2; auto.
apply FcanonicBound with (radix := radix); auto with zarith arith.
apply FNSuccCanonic; auto with zarith.
rewrite Rabs_right.
rewrite Rabs_right.
replace (p - q + (FNSucc b radix precision q - p))%R with
 (FNSucc b radix precision q - q)%R; [ idtac | ring ].
unfold FtoRradix in |- *; apply FulpSuc; auto.
case H'1; auto.
apply Rge_minus; apply Rle_ge; auto with real zarith.
case MinMax with (3 := pGivesBound) (r := p) (p := q); auto with zarith.
intros H'0 H'2; elim H'2; intros H'3 H'4; apply H'3; clear H'2; auto.
apply Rge_minus; apply Rle_ge; auto with real.
apply isMin_inv1 with (1 := H'1).
apply Rle_trans with (Rabs (p - q) + Rabs (p - FNPred b radix precision q))%R.
apply Rplus_le_compat_l.
rewrite <- (Rabs_Ropp (p - q));
 rewrite <- (Rabs_Ropp (p - FNPred b radix precision q)).
repeat rewrite Ropp_minus_distr.
elim H'; auto.
intros H'0 H'2; apply H'2; auto.
apply FcanonicBound with (radix := radix); auto with zarith.
apply FNPredCanonic; auto with zarith.
rewrite <- (Rabs_Ropp (p - q)); rewrite Ropp_minus_distr.
rewrite Rabs_right.
rewrite Rabs_right.
replace (q - p + (p - FNPred b radix precision q))%R with
 (q - FNPred b radix precision q)%R; [ idtac | ring ].
unfold FtoRradix in |- *; apply FulpPred; auto.
case H'1; auto.
apply Rge_minus; apply Rle_ge; auto with real.
case MaxMin with (3 := pGivesBound) (r := p) (p := q); auto with zarith.
intros H'0 H'2; elim H'2; intros H'3 H'4; apply H'3; clear H'2; auto.
apply Rge_minus; apply Rle_ge; auto with real.
apply isMax_inv1 with (1 := H'1).
Qed.

Theorem ClosestExp :
 forall (p : R) (q : float),
 Closest b radix p q -> (2%nat * Rabs (p - q) <= powerRZ radix (Fexp q))%R.
intros p q H'.
apply Rle_trans with (Fulp b radix precision q).
apply (ClosestUlp p q); auto.
replace (powerRZ radix (Fexp q)) with (FtoRradix (Float 1%nat (Fexp q))).
apply (FulpLe b radix); auto.
apply
 RoundedModeBounded with (radix := radix) (P := Closest b radix) (r := p);
 auto.
apply ClosestRoundedModeP with (precision := precision); auto.
unfold FtoRradix, FtoR in |- *; simpl in |- *.
ring.
Qed.

Theorem ClosestErrorExpStrict :
 forall (p q : float) (x : R),
 Fbounded b p ->
 Fbounded b q ->
 Closest b radix x p ->
 q = (x - p)%R :>R -> q <> 0%R :>R -> (Fexp q < Fexp p)%Z.
intros.
case (Zle_or_lt (Fexp p) (Fexp q)); auto; intros Z1.
absurd (powerRZ radix (Fexp p) <= powerRZ radix (Fexp q))%R.
2: apply Rle_powerRZ; auto with real arith.
apply Rgt_not_le.
red in |- *; apply Rlt_le_trans with (2%nat * powerRZ radix (Fexp q))%R.
apply Rltdouble; auto with real arith.
apply powerRZ_lt, Rlt_IZR; omega.
apply Rle_trans with (2%nat * Fabs q)%R.
apply Rmult_le_compat_l; auto with real arith.
replace (powerRZ radix (Fexp q)) with (FtoRradix (Float 1%nat (Fexp q)));
 auto.
apply (RleFexpFabs radix); auto with real zarith.
unfold FtoRradix, FtoR in |- *; simpl in |- *; ring.
rewrite (Fabs_correct radix); auto with arith.
replace (FtoR radix q) with (x - p)%R; auto.
apply ClosestExp; auto.
apply Rle_IZR; omega.
Qed.

Theorem ClosestIdem :
 forall p q : float, Fbounded b p -> Closest b radix p q -> p = q :>R.
intros p q H' H'0.
case (Rabs_pos (q - p)); intros H1.
contradict H1; apply Rle_not_lt.
replace 0%R with (Rabs (p - p)); [ case H'0; auto | idtac ].
replace (p - p)%R with 0%R; [ apply Rabs_R0; auto | ring ].
apply Rplus_eq_reg_l with (r := (- p)%R).
apply trans_eq with 0%R; [ ring | idtac ].
apply trans_eq with (q - p)%R; [ idtac | ring ].
generalize H1; unfold Rabs in |- *; case (Rcase_abs (q - p)); auto.
intros r H0; replace 0%R with (-0)%R; [ rewrite H0 | idtac ]; ring.
Qed.

Theorem FmultRadixInv :
 forall (x z : float) (y : R),
 Fbounded b x ->
 Closest b radix y z -> (/ 2%nat * x < y)%R -> (/ 2%nat * x <= z)%R.
intros x z y H' H'0 H'1.
case MinEx with (r := (/ 2%nat * x)%R) (3 := pGivesBound); auto with zarith.
intros min isMin.
case MaxEx with (r := (/ 2%nat * x)%R) (3 := pGivesBound); auto with zarith.
intros max isMax.
case (Rle_or_lt y max); intros Rl1.
case Rl1; clear Rl1; intros Rl1.
replace (FtoRradix z) with (FtoRradix max).
apply isMax_inv1 with (1 := isMax).
apply sym_eq.
unfold FtoRradix in |- *;
 apply ClosestMaxEq with (b := b) (r := y) (min := min);
 auto.
apply isMinComp with (r1 := (/ 2%nat * x)%R) (max := max); auto.
apply Rle_lt_trans with (2 := H'1); auto.
apply isMin_inv1 with (1 := isMin).
apply isMaxComp with (r1 := (/ 2%nat * x)%R) (min := min); auto.
apply Rle_lt_trans with (2 := H'1); auto.
apply isMin_inv1 with (1 := isMin).
replace (FtoR radix min + FtoR radix max)%R with (FtoRradix x).
apply Rmult_lt_reg_l with (r := (/ 2%nat)%R); auto with real.
rewrite <- Rmult_assoc; rewrite Rinv_l; try rewrite Rmult_1_l; auto with real.
unfold FtoRradix in |- *; apply (div2IsBetween b radix precision); auto.
cut (Closest b radix max z); [ intros C0 | idtac ].
replace (FtoRradix z) with (FtoRradix max); auto.
rewrite <- Rl1; auto.
apply Rlt_le; auto.
apply ClosestIdem; auto.
case isMax; auto.
apply (ClosestCompatible b radix y max z z); auto.
case H'0; auto.
apply Rle_trans with (FtoRradix max); auto.
apply isMax_inv1 with (1 := isMax).
apply (ClosestMonotone b radix (FtoRradix max) y); auto.
apply (RoundedModeProjectorIdem b radix (Closest b radix)); auto.
apply ClosestRoundedModeP with (precision := precision); auto.
case isMax; auto.
Qed.

Theorem ClosestErrorBound :
 forall (p q : float) (x : R),
 Fbounded b p ->
 Closest b radix x p ->
 q = (x - p)%R :>R -> (Rabs q <= Float 1%nat (Fexp p) * / 2%nat)%R.
intros p q x H H0 H1.
apply Rle_trans with (Fulp b radix precision p * / 2%nat)%R.
rewrite H1.
replace (Rabs (x - p)) with (2%nat * Rabs (x - p) * / 2%nat)%R;
 [ idtac | field; auto with real ].
apply Rmult_le_compat_r; auto with real.
simpl; auto with real.
apply ClosestUlp; auto.
apply Rmult_le_compat_r.
apply Rlt_le.
apply Rinv_0_lt_compat; auto with real.
unfold FtoRradix in |- *; apply FulpLe; auto.
Qed.

Theorem ClosestErrorBoundNormal_aux :
 forall (x : R) (p : float),
 Closest b radix x p ->
 Fnormal radix b (Fnormalize radix b precision p) ->
 (Rabs (x - p) <= Rabs p * (/ 2%nat * (radix * / Zpos (vNum b))))%R.
intros x p H H'.
apply Rle_trans with (/ 2%nat * Fulp b radix precision p)%R.
replace (Rabs (x - FtoRradix p)) with
 (/ 2%nat * (2%nat * Rabs (x - FtoRradix p)))%R.
apply Rmult_le_compat_l; [simpl; auto with real|idtac].
apply ClosestUlp; auto.
rewrite <- Rmult_assoc; rewrite Rinv_l; simpl in |- *; auto with real.
apply
 Rle_trans with (/ 2%nat * (Rabs p * (radix * / Zpos (vNum b))))%R;
 [ apply Rmult_le_compat_l | right; ring; ring ].
apply Rlt_le; apply Rinv_0_lt_compat; auto with real arith.
unfold Fulp in |- *.
replace (Fexp (Fnormalize radix b precision p)) with
 (Fexp (Fnormalize radix b precision p) + precision + - precision)%Z;
 [ idtac | ring ].
rewrite powerRZ_add; auto with real zarith.
2: apply IZR_neq; omega.
apply Rle_trans with (Rabs p * radix * powerRZ radix (- precision))%R;
 [ apply Rmult_le_compat_r | right ]; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
2: rewrite pGivesBound; simpl in |- *.
2: rewrite powerRZ_Zopp; auto with real zarith.
2: rewrite Zpower_nat_Z_powerRZ; auto with real zarith; ring.
2: apply IZR_neq; omega.
replace (FtoRradix p) with (FtoRradix (Fnormalize radix b precision p));
 [ idtac | apply (FnormalizeCorrect radix) ]; auto.
rewrite <- (Fabs_correct radix); unfold FtoR in |- *; simpl in |- *;
 auto with arith.
rewrite powerRZ_add; auto with real zarith.
2: apply IZR_neq; omega.
replace
 (Z.abs (Fnum (Fnormalize radix b precision p)) *
  powerRZ radix (Fexp (Fnormalize radix b precision p)) * radix)%R with
 (powerRZ radix (Fexp (Fnormalize radix b precision p)) *
  (Z.abs (Fnum (Fnormalize radix b precision p)) * radix))%R;
 [ idtac | ring ].
apply Rmult_le_compat_l; auto with arith real.
apply powerRZ_le, Rlt_IZR; omega.
rewrite <- Zpower_nat_Z_powerRZ; auto with real zarith.
rewrite <- mult_IZR; apply Rle_IZR.
rewrite <- pGivesBound; pattern radix at 2 in |- *;
 rewrite <- (Z.abs_eq radix); auto with zarith.
rewrite <- Zabs_Zmult.
rewrite Zmult_comm; elim H'; auto.
Qed.

Theorem ClosestErrorBoundNormal :
 forall (x : R) (p : float),
 Closest b radix x p ->
 Fnormal radix b (Fnormalize radix b precision p) ->
 (Rabs (x - p) <= Rabs p * (/ 2%nat * powerRZ radix (Z.succ (- precision))))%R.
intros x p H H1.
apply
 Rle_trans
  with (Rabs (FtoRradix p) * (/ 2%nat * (radix * / Zpos (vNum b))))%R;
 [ apply ClosestErrorBoundNormal_aux; auto | right ].
replace (powerRZ radix (Z.succ (- precision))) with
 (radix * / Zpos (vNum b))%R; auto with real.
rewrite pGivesBound; rewrite Zpower_nat_Z_powerRZ.
rewrite Rinv_powerRZ; auto with real zarith.
rewrite powerRZ_Zs; auto with real zarith.
apply IZR_neq; omega.
apply IZR_neq; omega.
Qed.

Theorem FpredUlpPos :
 forall x : float,
 Fcanonic radix b x ->
 (0 < x)%R ->
 (FPred b radix precision x +
  Fulp b radix precision (FPred b radix precision x))%R = x.
intros x Hx H.
apply sym_eq;
 apply Rplus_eq_reg_l with (- FtoRradix (FPred b radix precision x))%R.
apply trans_eq with (Fulp b radix precision (FPred b radix precision x));
 [ idtac | ring ].
apply trans_eq with (FtoRradix x - FtoRradix (FPred b radix precision x))%R;
 [ ring | idtac ].
unfold FtoRradix in |- *; rewrite <- Fminus_correct; auto with zarith;
 fold FtoRradix in |- *.
pattern x at 1 in |- *;
 replace x with (FSucc b radix precision (FPred b radix precision x));
 [ idtac | apply FSucPred; auto with zarith arith ].
unfold FtoRradix in |- *; apply FSuccUlpPos; auto with zarith arith.
apply FPredCanonic; auto with zarith arith.
apply R0RltRlePred; auto with zarith arith real.
Qed.

Theorem FulpFPredLe :
 forall f : float,
 Fbounded b f ->
 Fcanonic radix b f ->
 (Fulp b radix precision f <=
  radix * Fulp b radix precision (FPred b radix precision f))%R.
intros f Hf1 Hf2; unfold Fulp in |- *.
replace (Fnormalize radix b precision f) with f;
 [ idtac
 | apply
    FcanonicUnique with (radix := radix) (b := b) (precision := precision);
    auto with arith zarith ].
2: apply FnormalizeCanonic; auto with zarith.
2: apply sym_eq; apply FnormalizeCorrect; auto with arith zarith.
replace (Fnormalize radix b precision (FPred b radix precision f)) with
 (FPred b radix precision f);
 [ idtac
 | apply
    FcanonicUnique with (radix := radix) (b := b) (precision := precision);
    auto with arith zarith ].
2: apply FPredCanonic; auto with zarith.
2: apply FnormalizeCanonic; auto with zarith.
2: apply FBoundedPred; auto with zarith.
2: apply sym_eq; apply FnormalizeCorrect; auto with arith zarith.
pattern (IZR radix) at 2 in |- *; replace (IZR radix) with (powerRZ radix 1);
 [ idtac | simpl in |- *; auto with arith zarith real ].
rewrite <- powerRZ_add; auto with zarith real.
2: apply IZR_neq; omega.
apply Rle_powerRZ; try apply Rle_IZR; auto with zarith real.
replace (1 + Fexp (FPred b radix precision f))%Z with
 (Z.succ (Fexp (FPred b radix precision f))); auto with zarith.
unfold FPred in |- *.
generalize (Z_eq_bool_correct (Fnum f) (- pPred (vNum b)));
 case (Z_eq_bool (Fnum f) (- pPred (vNum b))); intros H1;
 [ simpl in |- *; auto with zarith | idtac ].
generalize (Z_eq_bool_correct (Fnum f) (nNormMin radix precision));
 case (Z_eq_bool (Fnum f) (nNormMin radix precision));
 intros H2; [ idtac | simpl in |- *; auto with zarith ].
generalize (Z_eq_bool_correct (Fexp f) (- dExp b));
 case (Z_eq_bool (Fexp f) (- dExp b)); intros H3; simpl in |- *;
 auto with zarith.
Qed.

End Fclosestp2.


Section FRoundPM.
Variable b : Fbound.
Variable radix : Z.
Variable precision : nat.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.

Hypothesis radixMoreThanOne : (1 < radix)%Z.
Hypothesis precisionGreaterThanOne : 1 < precision.
Hypothesis pGivesBound : Zpos (vNum b) = Zpower_nat radix precision.

Theorem errorBoundedMultClosest_aux :
 forall p q pq : float,
 Fbounded b p ->
 Fbounded b q ->
 Closest b radix (p * q) pq ->
 (- dExp b <= Fexp p + Fexp q)%Z ->
 (p * q - pq)%R <> 0%R :>R ->
 ex
   (fun r : float =>
    ex
      (fun s : float =>
       Fcanonic radix b r /\
       Fbounded b r /\
       Fbounded b s /\
       r = pq :>R /\
       s = (p * q - r)%R :>R /\
       Fexp s = (Fexp p + Fexp q)%Z :>Z /\
       (Fexp s <= Fexp r)%Z /\ (Fexp r <= precision + (Fexp p + Fexp q))%Z)).
intros p q pq Hp Hq H1 H2 H3.
cut (RoundedModeP b radix (Closest b radix));
 [ intros H4 | apply ClosestRoundedModeP with precision; auto ].
lapply (errorBoundedMultExp b radix precision);
 [ intros H'2; lapply H'2;
    [ intros H'3; lapply H'3;
       [ intros H'4; lapply (H'4 (Closest b radix));
             [ intros H'7; elim (H'7 p q pq);
                [ intros r E; elim E; intros s E0; elim E0; intros H'15 H'16;
                   elim H'16; intros H'17 H'18; elim H'18;
                   intros H'19 H'20; elim H'20; intros H'21 H'22;
                   elim H'22; intros H'23 H'24; elim H'24;
                   intros H'25 H'26;
                   clear H'24 H'22 H'20 H'18 H'16 E0 E H'3 H'2
                | clear H'3 H'2
                | clear H'3 H'2
                | clear H'3 H'2
                | clear H'3 H'2 ]
             | clear H'3 H'2 ]
          | clear H'3 H'2 ]
    | clear H'2 ]
  | idtac]; auto.
exists (Fnormalize radix b precision r); exists s.
cut (Fbounded b (Fnormalize radix b precision r));
 [ intros H5 | apply FnormalizeBounded; auto with zarith ].
split; [ apply FnormalizeCanonic; auto with zarith | idtac ].
repeat (split; auto).
unfold FtoRradix in |- *; rewrite <- H'19; unfold FtoRradix in |- *;
 apply FnormalizeCorrect; auto.
unfold FtoRradix in |- *; rewrite FnormalizeCorrect; auto with arith.
apply Zlt_le_weak.
apply
 RoundedModeErrorExpStrict with b radix precision (Closest b radix) (p * q)%R;
 auto with arith.
generalize ClosestCompatible; unfold CompatibleP in |- *; intros H6.
generalize
 (H6 b radix (FtoRradix p * FtoRradix q)%R (FtoRradix p * FtoRradix q)%R pq);
 intros H9; apply H9; auto.
rewrite FnormalizeCorrect; auto with real arith.
rewrite FnormalizeCorrect; auto with real arith.
rewrite H'21; rewrite H'19; auto.
apply Z.le_trans with (Fexp r); auto.
apply FcanonicLeastExp with radix b precision; auto with zarith.
rewrite FnormalizeCorrect; auto with real arith.
apply FnormalizeCanonic; auto with zarith.
Qed.

Theorem errorBoundedMultClosest :
 forall p q pq : float,
 Fbounded b p ->
 Fbounded b q ->
 Closest b radix (p * q) pq ->
 (- dExp b <= Fexp p + Fexp q)%Z ->
 (- dExp b <= Fexp (Fnormalize radix b precision pq) - precision)%Z ->
 ex
   (fun r : float =>
    ex
      (fun s : float =>
       Fcanonic radix b r /\
       Fbounded b r /\
       Fbounded b s /\
       r = pq :>R /\
       s = (p * q - r)%R :>R /\ Fexp s = (Fexp r - precision)%Z :>Z)).
intros.
cut (RoundedModeP b radix (Closest b radix));
 [ intros G1 | apply ClosestRoundedModeP with precision; auto ].
case (Req_dec (p * q - pq) 0); intros U.
exists (Fnormalize radix b precision pq);
 exists (Fzero (Fexp (Fnormalize radix b precision pq) - precision)).
cut (Fbounded b pq);
 [ intros G2
 | apply RoundedModeBounded with radix (Closest b radix) (p * q)%R; auto ].
cut (Fcanonic radix b (Fnormalize radix b precision pq));
 [ intros G3 | apply FnormalizeCanonic; auto with zarith ].
cut (Fbounded b (Fnormalize radix b precision pq));
 [ intros G4 | apply FnormalizeBounded; auto with zarith ].
cut (Fnormalize radix b precision pq = pq :>R);
 [ intros G5
 | unfold FtoRradix in |- *; apply FnormalizeCorrect; auto with arith ].
repeat (split; auto).
rewrite G5; unfold FtoRradix in |- *; rewrite FzeroisReallyZero;
 auto with real.
lapply (errorBoundedMultClosest_aux p q pq); auto; intros H5.
lapply H5; auto; intros H6; clear H5.
lapply H6; auto; intros H5; clear H6.
lapply H5; auto; intros H6; clear H5.
lapply H6; auto; intros H5; clear H6.
elim H5; intros r H6; clear H5.
elim H6; intros s H5; clear H6.
elim H5; intros H7 H6; clear H5.
elim H6; intros H8 H9; clear H6.
elim H9; intros H6 H10; clear H9.
elim H10; intros H9 H11; clear H10.
elim H11; intros H10 H12; clear H11.
elim H12; intros H11 H13; clear H12.
elim H13; intros H12 H14; clear H13.
cut
 (ex
    (fun m : Z =>
     s = Float m (Fexp r - precision) :>R /\ (Z.abs m <= pPred (vNum b))%Z)).
intros H13; elim H13; intros m H15; elim H15; intros H16 H17; clear H15 H13.
exists r; exists (Float m (Fexp r - precision)).
split; auto.
split; auto.
split.
2: repeat (split; auto).
2: rewrite <- H16; auto.
split; simpl in |- *.
generalize H17; unfold pPred in |- *; apply Zle_Zpred_inv.
replace r with (Fnormalize radix b precision pq); auto with zarith.
apply FcanonicUnique with radix b precision; auto with zarith.
apply FnormalizeCanonic; auto with zarith; elim H1; auto.
rewrite FnormalizeCorrect; auto with real zarith.
cut (radix <> 0%Z :>Z); [ intros V | auto with arith real zarith ].
cut (0 < radix)%Z; [ intros V2 | auto with arith real zarith ].
rewrite H10; unfold FtoRradix in |- *; rewrite <- Fmult_correct; auto.
rewrite <- Fminus_correct; fold FtoRradix in |- *; auto.
unfold Fmult in |- *; unfold Fminus in |- *; unfold Fopp in |- *;
 unfold Fplus in |- *; simpl in |- *.
unfold FtoRradix in |- *; unfold FtoR in |- *; simpl in |- *.
rewrite Zmin_le1; auto with zarith.
replace
 (Fnum p * Fnum q *
  Zpower_nat radix (Z.abs_nat (Fexp p + Fexp q - (Fexp p + Fexp q))))%Z with
 (Fnum p * Fnum q)%Z.
2: replace (Fexp p + Fexp q - (Fexp p + Fexp q))%Z with 0%Z;
    auto with zarith arith; simpl in |- *.
2: auto with zarith.
exists
 ((Fnum p * Fnum q +
   - Fnum r * Zpower_nat radix (Z.abs_nat (Fexp r - (Fexp p + Fexp q)))) *
  Zpower_nat radix (Z.abs_nat (Fexp p + Fexp q + (precision - Fexp r))))%Z;
 split.
rewrite plus_IZR.
repeat rewrite mult_IZR.
rewrite plus_IZR.
repeat rewrite mult_IZR.
rewrite (Zpower_nat_powerRZ_absolu radix (Fexp r - (Fexp p + Fexp q))).
2: auto with zarith arith.
rewrite
 (Zpower_nat_powerRZ_absolu radix (Fexp p + Fexp q + (precision - Fexp r)))
 .
2: auto with zarith arith.
cut (radix <> 0%R :>R). intros W.
unfold Zminus in |- *.
repeat rewrite powerRZ_add; try rewrite <- INR_IZR_INZ; auto.
apply
 trans_eq
  with
    ((Fnum p * Fnum q +
      (- Fnum r)%Z *
      (powerRZ radix (Fexp r) * powerRZ radix (- (Fexp p + Fexp q)))) *
     (powerRZ radix (Fexp p) * powerRZ radix (Fexp q)))%R.
ring; ring.
apply
 trans_eq
  with
    ((Fnum p * Fnum q +
      (- Fnum r)%Z *
      (powerRZ radix (Fexp r) * powerRZ radix (- (Fexp p + Fexp q)))) *
     (powerRZ radix (Fexp p) * powerRZ radix (Fexp q) *
      (powerRZ radix precision * powerRZ radix (- precision))) *
     (powerRZ radix (Fexp r) * powerRZ radix (- Fexp r)))%R.
2: ring; ring.
replace (powerRZ radix precision * powerRZ radix (- precision))%R with 1%R.
replace (powerRZ radix (Fexp r) * powerRZ radix (- Fexp r))%R with 1%R;
 try ring.
rewrite <- powerRZ_add; try rewrite <- INR_IZR_INZ; auto.
rewrite Zplus_opp_r; simpl in |- *; auto.
rewrite <- powerRZ_add; try rewrite <- INR_IZR_INZ; auto.
rewrite Zplus_opp_r; simpl in |- *; auto.
apply IZR_neq; omega.
apply le_IZR; rewrite <- Rabs_Zabs.
rewrite mult_IZR; rewrite plus_IZR.
repeat rewrite mult_IZR.
rewrite
 (Zpower_nat_powerRZ_absolu radix (Fexp p + Fexp q + (precision - Fexp r)))
 .
2: auto with zarith arith.
rewrite (Zpower_nat_powerRZ_absolu radix (Fexp r - (Fexp p + Fexp q))).
2: auto with zarith arith.
rewrite powerRZ_add; try rewrite <- INR_IZR_INZ; auto with real arith.
replace
 ((Fnum p * Fnum q +
   (- Fnum r)%Z * powerRZ radix (Fexp r - (Fexp p + Fexp q))) *
  (powerRZ radix (Fexp p + Fexp q) * powerRZ radix (precision - Fexp r)))%R
 with
 ((Fnum p * Fnum q +
   (- Fnum r)%Z * powerRZ radix (Fexp r - (Fexp p + Fexp q))) *
  powerRZ radix (Fexp p + Fexp q) * powerRZ radix (precision - Fexp r))%R;
 [ idtac | ring; ring ].
rewrite Rabs_mult.
rewrite (Rabs_right (powerRZ radix (precision - Fexp r))).
2: apply Rle_ge; apply powerRZ_le; auto with real zarith.
apply Rmult_le_reg_l with (powerRZ radix (Fexp r - precision)).
apply powerRZ_lt, Rlt_IZR; auto with real arith.
rewrite Rmult_comm; rewrite Rmult_assoc; rewrite <- powerRZ_add.
2: apply IZR_neq; auto with zarith arith real.
2: apply Rlt_IZR; auto with zarith arith real.
2: apply IZR_neq; auto with zarith arith real.
replace (precision - Fexp r + (Fexp r - precision))%Z with 0%Z;
 [ simpl in |- * | ring ].
apply
 Rle_trans
  with
    (Rabs
       ((Fnum p * Fnum q +
         (- Fnum r)%Z * powerRZ radix (Fexp r - (Fexp p + Fexp q))) *
        powerRZ radix (Fexp p + Fexp q))); [ right; ring | idtac ].
replace
 ((Fnum p * Fnum q +
   (- Fnum r)%Z * powerRZ radix (Fexp r - (Fexp p + Fexp q))) *
  powerRZ radix (Fexp p + Fexp q))%R with (p * q - r)%R.
2: unfold FtoRradix in |- *; unfold FtoR in |- *; simpl in |- *;
    unfold Rminus in |- *.
2: unfold Zminus in |- *; repeat rewrite Ropp_Ropp_IZR.
2: repeat rewrite powerRZ_add; auto with real arith.
2: apply
    trans_eq
     with
       (Fnum p * Fnum q * (powerRZ radix (Fexp p) * powerRZ radix (Fexp q)) +
        - Fnum r *
        (powerRZ radix (Fexp r) *
         (powerRZ radix (- (Fexp p + Fexp q)) *
          (powerRZ radix (Fexp p) * powerRZ radix (Fexp q)))))%R;
    [ idtac | ring ].
2: replace
    (powerRZ radix (- (Fexp p + Fexp q)) *
     (powerRZ radix (Fexp p) * powerRZ radix (Fexp q)))%R with 1%R;
    try ring.
2: repeat rewrite <- powerRZ_add; auto with real arith.
2: replace (- (Fexp p + Fexp q) + (Fexp p + Fexp q))%Z with 0%Z;
    simpl in |- *; simpl; ring.
apply Rle_trans with (powerRZ radix (Fexp r) * / 2%nat)%R.
rewrite <- H10;
 replace (powerRZ radix (Fexp r)) with (FtoRradix (Float 1%nat (Fexp r)));
 unfold FtoRradix in |- *;
 [ idtac | unfold FtoR in |- *; simpl in |- *; ring ].
apply ClosestErrorBound with b precision (p * q)%R; auto.
apply (ClosestCompatible b radix (p * q)%R (p * q)%R pq); auto.
unfold Zminus in |- *; rewrite powerRZ_add; auto with real arith.
rewrite Rmult_assoc; apply Rmult_le_compat_l.
apply powerRZ_le; auto with real arith.
apply Rlt_IZR; omega.
unfold pPred, Z.pred in |- *; rewrite pGivesBound.
rewrite plus_IZR; rewrite Zpower_nat_Z_powerRZ.
replace (powerRZ radix (- precision) * (powerRZ radix precision + (-1)%Z))%R
 with (1 + - powerRZ radix (- precision))%R.
apply Rle_trans with (1 + - powerRZ radix (- 1%nat))%R.
simpl in |- *.
replace (radix * 1)%R with (IZR radix); [ idtac | ring ].
replace (/ (1+1))%R with (1 + - / 2)%R.
apply Rplus_le_compat_l; apply Ropp_le_contravar.
apply Rle_Rinv; auto with real arith zarith.
apply Rle_IZR; omega.
field.
apply Rplus_le_compat_l; apply Ropp_le_contravar.
apply Rle_powerRZ; auto with real arith zarith.
apply Rle_IZR; omega.
rewrite Rmult_plus_distr_l.
rewrite <- powerRZ_add; auto with real arith.
replace (- precision + precision)%Z with 0%Z; simpl in |- *; ring.
apply IZR_neq; omega.
apply IZR_neq; omega.
apply IZR_neq; omega.
apply IZR_neq; omega.
apply IZR_neq; omega.
apply IZR_neq; omega.
Qed.
End FRoundPM.

Section finduct.
Variable b : Fbound.
Variable radix : Z.
Variable precision : nat.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.

Hypothesis radixMoreThanOne : (1 < radix)%Z.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ radixMoreThanOne).
Hypothesis precisionNotZero : precision <> 0.
Hypothesis pGivesBound : Zpos (vNum b) = Zpower_nat radix precision.

Definition Fweight (p : float) :=
  (Fnum p + Fexp p * Zpower_nat radix precision)%Z.

Theorem FweightLt :
 forall p q : float,
 Fcanonic radix b p ->
 Fcanonic radix b q -> (0 <= p)%R -> (p < q)%R -> (Fweight p < Fweight q)%Z.
intros p q H' H'0 H'1 H'2.
cut (Fbounded b p); [ intros Fb1 | apply FcanonicBound with (1 := H') ]; auto.
cut (Fbounded b q); [ intros Fb2 | apply FcanonicBound with (1 := H'0) ];
 auto.
case (FcanonicLtPos _ radixMoreThanOne b precision) with (p := p) (q := q);
 auto with arith; intros Zl1.
unfold Fweight in |- *; simpl in |- *.
replace (Fexp q) with (Fexp q - Fexp p + Fexp p)%Z; [ idtac | ring ].
rewrite Zmult_plus_distr_l.
rewrite Zplus_assoc.
repeat rewrite (fun x y z : Z => Zplus_comm x (y * z)).
apply Zplus_lt_compat_l.
apply Z.lt_le_trans with (Zpower_nat radix precision); auto with zarith.
apply Z.le_lt_trans with (Z.pred (Zpower_nat radix precision));
 auto with zarith.
apply Zle_Zabs_inv2; auto with zarith.
apply Zle_Zpred; auto with zarith.
rewrite <- pGivesBound; apply Fb1.
apply Z.le_trans with ((Fexp q - Fexp p) * Zpower_nat radix precision)%Z;
 auto with zarith.
pattern (Zpower_nat radix precision) at 1 in |- *;
 replace (Zpower_nat radix precision) with
  (Z.succ 0 * Zpower_nat radix precision)%Z; auto.
apply Zle_Zmult_comp_r; auto with zarith.
unfold Z.succ in |- *; ring.
cut (0 <= Fnum q)%Z; auto with zarith.
apply (LeR0Fnum radix); auto.
apply Rle_trans with (FtoRradix p); auto; apply Rlt_le; auto.
elim Zl1; intros H'3 H'4; clear Zl1.
unfold Fweight in |- *; simpl in |- *.
rewrite <- H'3.
repeat rewrite (fun x y z : Z => Zplus_comm x (y * z)).
apply Zplus_lt_compat_l; auto.
Qed.

Theorem FweightEq :
 forall p q : float,
 Fcanonic radix b p ->
 Fcanonic radix b q -> p = q :>R -> Fweight p = Fweight q.
intros p q H' H'0 H'1.
rewrite
 (FcanonicUnique _ radixMoreThanOne b precision) with (p := p) (q := q);
 auto with arith.
Qed.

Theorem FweightZle :
 forall p q : float,
 Fcanonic radix b p ->
 Fcanonic radix b q -> (0 <= p)%R -> (p <= q)%R -> (Fweight p <= Fweight q)%Z.
intros p q H' H'0 H'1 H'2; Casec H'2; intros H'2.
apply Zlt_le_weak.
apply FweightLt; auto.
rewrite (FweightEq p q); auto with zarith.
Qed.

Theorem FinductNegAux :
 forall (P : float -> Prop) (p : float),
 (0 <= p)%R ->
 Fcanonic radix b p ->
 P p ->
 (forall q : float,
  Fcanonic radix b q ->
  (0 < q)%R -> (q <= p)%R -> P q -> P (FPred b radix precision q)) ->
 forall x : Z,
 (0 <= x)%Z ->
 forall q : float,
 x = (Fweight p - Fweight q)%Z ->
 Fcanonic radix b q -> (0 <= q)%R -> (q <= p)%R -> P q.
intros P p H' H'0 H'1 H'2 x H'3; pattern x in |- *.
apply Z_lt_induction; auto.
intros x0 H'4 q H'5 H'6 H'7 H'8.
Casec H'8; intros H'8.
cut (FSucc b radix precision q <= p)%R; [ intros Rle1 | idtac ].
cut (P (FSucc b radix precision q)); [ intros P1 | idtac ].
rewrite <- (FPredSuc b radix precision) with (x := q); auto with arith.
apply H'2; auto with zarith arith.
apply FSuccCanonic; auto with zarith.
apply Rle_lt_trans with (FtoRradix q); auto.
apply (FSuccLt b radix); auto with arith.
apply H'4 with (y := (Fweight p - Fweight (FSucc b radix precision q))%Z);
 auto.
split.
cut (Fweight (FSucc b radix precision q) <= Fweight p)%Z; auto with zarith.
apply FweightZle; auto.
apply FSuccCanonic; auto with arith.
apply Rle_trans with (FtoRradix q); auto; apply Rlt_le.
apply (FSuccLt b radix); auto with arith.
rewrite H'5.
cut (Fweight q < Fweight (FSucc b radix precision q))%Z;
 [ auto with zarith | idtac ].
apply FweightLt; auto with zarith.
apply FSuccCanonic; auto with arith.
apply (FSuccLt b radix); auto with arith.
apply FSuccCanonic; auto with arith.
apply Rle_trans with (FtoRradix q); auto; apply Rlt_le.
apply (FSuccLt b radix); auto with arith.
apply (FSuccProp b radix); auto with arith.
rewrite <-
 (FcanonicUnique _ radixMoreThanOne b precision) with (p := p) (q := q);
 auto with arith.
Qed.

Theorem FinductNeg :
 forall (P : float -> Prop) (p : float),
 (0 <= p)%R ->
 Fcanonic radix b p ->
 P p ->
 (forall q : float,
  Fcanonic radix b q ->
  (0 < q)%R -> (q <= p)%R -> P q -> P (FPred b radix precision q)) ->
 forall q : float, Fcanonic radix b q -> (0 <= q)%R -> (q <= p)%R -> P q.
intros P p H' H'0 H'1 H'2 q H'3 H'4 H'5.
apply FinductNegAux with (p := p) (x := (Fweight p - Fweight q)%Z); auto.
cut (Fweight q <= Fweight p)%Z; [ auto with zarith | idtac ].
apply FweightZle; auto with zarith.
Qed.

Theorem radixRangeBoundExp :
 forall p q : float,
 Fcanonic radix b p ->
 Fcanonic radix b q ->
 (0 <= p)%R ->
 (p < q)%R -> (q < radix * p)%R -> Fexp p = Fexp q \/ Z.succ (Fexp p) = Fexp q.
intros p q H' H'0 H'1 H'2 H'3.
case (FcanonicLtPos _ radixMoreThanOne b precision) with (p := p) (q := q);
 auto with arith.
2: intros H'4; elim H'4; intros H'5 H'6; clear H'4; auto.
intros H'4; right.
Casec H'; intros H'.
case
 (FcanonicLtPos _ radixMoreThanOne b precision)
  with (p := q) (q := Float (Fnum p) (Z.succ (Fexp p)));
 auto with arith.
left.
case H'; intros H1 H2; red in H1.
repeat split; simpl in |- *; auto with zarith.
apply Rle_trans with (FtoRradix p); auto; apply Rlt_le; auto.
unfold FtoR in |- *; simpl in |- *.
rewrite powerRZ_Zs; auto with real zarith; auto.
rewrite <- Rmult_assoc;
 rewrite (fun (x : R) (y : Z) => Rmult_comm x y);
 rewrite Rmult_assoc; auto.
apply IZR_neq; omega.
simpl in |- *; intros; apply Zle_antisym; auto with zarith.
simpl in |- *; auto.
intros H'5; elim H'5; intros H'6 H'7; auto.
case
 (FcanonicLtPos _ radixMoreThanOne b precision)
  with (p := q) (q := Float (nNormMin radix precision) (Z.succ (Fexp p)));
 auto with arith.
left; repeat split; simpl in |- *.
rewrite Z.abs_eq; auto with zarith.
apply ZltNormMinVnum; auto with zarith.
unfold nNormMin in |- *; apply Zpower_NR0; auto with zarith.
apply Z.le_trans with (Fexp p); auto with zarith; apply H'.
case H'; auto.
rewrite <- (PosNormMin radix b precision); auto with zarith.
apply Rle_trans with (1 := H'1); auto with real.
apply Rlt_trans with (1 := H'3).
unfold FtoR in |- *; simpl in |- *.
rewrite powerRZ_Zs; auto with real zarith; auto.
rewrite <- Rmult_assoc;
 rewrite (fun (x : R) (y : Z) => Rmult_comm x y);
 rewrite Rmult_assoc; auto.
apply Rmult_lt_compat_l; auto with real arith.
apply Rlt_IZR; omega.
case H'.
intros H'5 H'6; elim H'6; intros H'7 H'8; rewrite H'7; clear H'6.
change (p < firstNormalPos radix b precision)%R in |- *.
apply (FsubnormalLtFirstNormalPos radix); auto with arith.
apply IZR_neq; omega.
simpl in |- *; intros; apply Zle_antisym; auto with zarith.
intros H'5; elim H'5; intros H'6 H'7; rewrite H'6; clear H'5; auto.
Qed.
End finduct.

Section FRoundPN.
Variable b : Fbound.
Variable radix : Z.
Variable precision : nat.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.
Hypothesis radixMoreThanOne : (1 < radix)%Z.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ radixMoreThanOne).

Hypothesis precisionGreaterThanOne : 1 < precision.
Hypothesis pGivesBound : Zpos (vNum b) = Zpower_nat radix precision.

Theorem plusExpMin :
 forall P,
 RoundedModeP b radix P ->
 forall p q pq : float,
 P (p + q)%R pq ->
 exists s : float,
   Fbounded b s /\ s = pq :>R /\ (Z.min (Fexp p) (Fexp q) <= Fexp s)%Z.
intros P H' p q pq H'0.
case
 (RoundedModeRep b radix precision)
  with (p := Fplus radix p q) (q := pq) (P := P); auto with zarith arith.
rewrite Fplus_correct; auto with arith.
simpl in |- *; intros x H'1.
case
 (eqExpLess _ radixMoreThanOne b)
  with (p := pq) (q := Float x (Fexp (Fplus radix p q)));
 auto.
apply (RoundedModeBounded b radix) with (P := P) (r := (p + q)%R); auto.
simpl in |- *; intros x0 H'2; elim H'2; intros H'3 H'4; elim H'4;
 intros H'5 H'6; clear H'4 H'2.
exists x0; split; [ idtac | split ]; auto.
unfold FtoRradix in |- *; rewrite H'5; auto.
apply le_IZR; auto.
Qed.

Theorem plusExpUpperBound :
 forall P,
 RoundedModeP b radix P ->
 forall p q pq : float,
 P (p + q)%R pq ->
 Fbounded b p ->
 Fbounded b q ->
 exists r : float,
   Fbounded b r /\ r = pq :>R /\ (Fexp r <= Z.succ (Zmax (Fexp p) (Fexp q)))%Z.
intros P H' p q pq H'0 H'1 H'2.
replace (Z.succ (Zmax (Fexp p) (Fexp q))) with
 (Fexp (Float (pPred (vNum b)) (Z.succ (Zmax (Fexp p) (Fexp q)))));
 [ idtac | simpl in |- *; auto ].
unfold FtoRradix in |- *; apply eqExpMax; auto.
apply RoundedModeBounded with (radix := radix) (P := P) (r := (p + q)%R);
 auto with arith.
unfold pPred in |- *; apply maxFbounded; auto.
apply Z.le_trans with (Fexp p); try apply H'1.
apply Z.le_trans with (Z.succ (Fexp p)); auto with zarith.
apply Zsucc_le_compat, ZmaxLe1.
replace
 (FtoR radix (Float (pPred (vNum b)) (Z.succ (Zmax (Fexp p) (Fexp q))))) with
 (radix * Float (pPred (vNum b)) (Zmax (Fexp p) (Fexp q)))%R.
rewrite Fabs_correct; auto with zarith.
unfold FtoRradix in |- *;
 apply
  RoundedModeMultAbs
   with (b := b) (precision := precision) (P := P) (r := (p + q)%R);
 auto.
unfold pPred in |- *; apply maxFbounded; auto.
apply Z.le_trans with (Fexp p); try apply H'1; apply ZmaxLe1.
apply Rle_trans with (Rabs p + Rabs q)%R.
apply Rabs_triang; auto.
apply
 Rle_trans
  with
    (2%nat * FtoR radix (Float (pPred (vNum b)) (Zmax (Fexp p) (Fexp q))))%R;
 auto.
cut (forall r : R, (2%nat * r)%R = (r + r)%R);
 [ intros tmp; rewrite tmp; clear tmp | intros; simpl in |- *; ring ].
apply Rplus_le_compat; auto.
rewrite <- (Fabs_correct radix); auto with arith; apply maxMax1; auto;
 apply ZmaxLe1.
rewrite <- (Fabs_correct radix); auto with arith; apply maxMax1; auto;
 apply ZmaxLe2.
apply Rmult_le_compat; auto with real arith.
replace 0%R with (INR 0); auto with real arith.
apply LeFnumZERO; simpl in |- *; auto; replace 0%Z with (Z_of_nat 0);
 auto with zarith.
unfold pPred in |- *; apply Zle_Zpred; auto with zarith.
rewrite INR_IZR_INZ; apply Rle_IZR; simpl in |- *; auto with zarith.
cut (1 < radix)%Z; auto with zarith;intros.
unfold FtoRradix, FtoR in |- *; simpl in |- *.
rewrite powerRZ_Zs; auto with real zarith; try ring.
apply IZR_neq; omega.
Qed.

Theorem plusExpBound :
 forall P,
 RoundedModeP b radix P ->
 forall p q pq : float,
 P (p + q)%R pq ->
 Fbounded b p ->
 Fbounded b q ->
 exists r : float,
   Fbounded b r /\
   r = pq :>R /\
   (Z.min (Fexp p) (Fexp q) <= Fexp r)%Z /\
   (Fexp r <= Z.succ (Zmax (Fexp p) (Fexp q)))%Z.
intros P H' p q pq H'0 H'1 H'2.
case (plusExpMin P H' _ _ _ H'0).
intros r' H'3; elim H'3; intros H'4 H'5; elim H'5; intros H'6 H'7;
 clear H'5 H'3.
case (Zle_or_lt (Fexp r') (Z.succ (Zmax (Fexp p) (Fexp q)))); intros Zl1.
exists r'; repeat (split; auto).
case (plusExpUpperBound P H' _ _ _ H'0); auto.
intros r'' H'3; elim H'3; intros H'5 H'8; elim H'8; intros H'9 H'10;
 clear H'8 H'3.
exists
 (Fshift radix (Z.abs_nat (Fexp r' - Z.succ (Zmax (Fexp p) (Fexp q)))) r');
 split.
apply FboundedShiftLess with (n := Z.abs_nat (Fexp r' - Fexp r'')); auto.
apply ZleLe; auto.
repeat rewrite <- Zabs_absolu.
repeat rewrite Z.abs_eq; auto with zarith.
rewrite FshiftCorrectInv; auto.
apply trans_eq with (FtoRradix pq); auto.
apply Z.le_trans with (1 := H'10); auto with zarith.
split.
unfold FtoRradix in |- *; rewrite FshiftCorrect; auto.
split.
simpl in |- *.
repeat rewrite inj_abs; auto with zarith arith.
apply Z.le_trans with (Zmax (Fexp p) (Fexp q)); auto with zarith.
apply Zmin_Zmax; auto.
simpl in |- *.
repeat rewrite inj_abs; auto with zarith arith.
Qed.

Theorem minusRoundRep :
 forall P,
 RoundedModeP b radix P ->
 forall p q qmp qmmp : float,
 (0 <= p)%R ->
 (p <= q)%R ->
 P (q - p)%R qmp ->
 Fbounded b p ->
 Fbounded b q -> exists r : float, Fbounded b r /\ r = (q - qmp)%R :>R.
intros P H' p q qmp H'0 H'1 H'2 H'3 H'4 H'5.
case (Rle_or_lt (/ 2%nat * q) p); intros Rle1.
exists p; split; auto.
replace (FtoRradix qmp) with (FtoRradix (Fminus radix q p)).
rewrite (Fminus_correct radix); auto with arith; unfold FtoRradix in |- *;
 ring.
apply (RoundedModeProjectorIdemEq b radix precision) with (P := P); auto.
rewrite <- Fopp_Fminus.
apply oppBounded; auto.
apply Sterbenz; auto.
apply Rle_trans with (FtoRradix q); auto with real.
apply Rledouble; auto.
apply Rle_trans with (FtoRradix p); auto with real.
cut (CompatibleP b radix P);
 [ intros Cp | apply RoundedModeP_inv2 with (1 := H'); auto ].
apply (Cp (q - p)%R (Fminus radix q p) qmp); auto.
rewrite (Fminus_correct radix); auto with arith.
apply RoundedModeBounded with (radix := radix) (P := P) (r := (q - p)%R);
 auto; auto.
exists (Fminus radix q qmp); split.
rewrite <- Fopp_Fminus.
apply oppBounded; auto.
apply Sterbenz; auto.
apply RoundedModeBounded with (radix := radix) (P := P) (r := (q - p)%R);
 auto; auto.
case MaxEx with (r := (/ 2%nat * FtoR radix q)%R) (3 := pGivesBound);
 auto with zarith.
intros max H'6.
apply Rle_trans with (FtoRradix max);
 [ apply isMax_inv1 with (1 := H'6); auto | idtac ].
apply (RleBoundRoundl b radix precision) with (P := P) (r := (q - p)%R); auto;
 fold FtoRradix in |- *.
case H'6; auto.
case MinEx with (r := (/ 2%nat * FtoR radix q)%R) (3 := pGivesBound);
 auto with zarith.
intros min H'7.
replace (FtoRradix max) with (q - min)%R.
apply Rplus_le_reg_l with (r := (- q)%R).
cut (forall p q : R, (- p + (p - q))%R = (- q)%R);
 [ intros tmp; repeat rewrite tmp; clear tmp | intros; ring ].
apply Ropp_le_contravar.
case H'7.
intros H'8 H'9; elim H'9; intros H'10 H'11; apply H'11; clear H'9; auto.
apply Rlt_le; auto.
unfold FtoRradix in |- *;
 rewrite (div2IsBetween b radix precision) with (5 := H'7) (6 := H'6);
 auto.
ring.
apply Rle_trans with (FtoRradix q); auto with real.
apply (RleBoundRoundr b radix precision) with (P := P) (r := (q - p)%R); auto;
 fold FtoRradix in |- *.
apply Rplus_le_reg_l with (r := (- q)%R).
cut (forall p q : R, (- p + (p - q))%R = (- q)%R);
 [ intros tmp; repeat rewrite tmp; clear tmp | intros; ring ].
replace (- q + q)%R with (-0)%R; [ auto with real | ring ].
apply Rle_trans with (FtoRradix q); auto with real.
apply Rledouble; auto.
apply Rle_trans with (FtoRradix p); auto with real.
apply (Fminus_correct radix); auto with arith.
Qed.

Theorem ExactMinusIntervalAux :
 forall P,
 RoundedModeP b radix P ->
 forall p q : float,
 (0 < p)%R ->
 (2%nat * p < q)%R ->
 Fcanonic radix b p ->
 Fcanonic radix b q ->
 (exists r : float, Fbounded b r /\ r = (q - p)%R :>R) ->
 forall r : float,
 Fcanonic radix b r ->
 (2%nat * p < r)%R ->
 (r <= q)%R -> exists r' : float, Fbounded b r' /\ r' = (r - p)%R :>R.
intros P H' p q H'0 H'1 H'2 H'3 H'4 r H'5 H'6 H'7.
cut (0 <= p)%R; [ intros Rle0 | apply Rlt_le; auto ].
cut (0 <= r)%R; [ intros Rle1 | apply Rle_trans with (2%nat * p)%R; auto ].
2: apply Rle_trans with (FtoRradix p); auto with zarith.
2: apply Rledouble; auto.
2: apply Rlt_le; auto.
generalize H'6; clear H'6; pattern r in |- *;
 apply (FinductNeg b radix precision) with (p := q);
 auto with zarith.
apply Rle_trans with (FtoRradix r); auto.
intros q0 H'6 H'8 H'9 H'10 H'11.
elim H'10;
 [ intros r' E; elim E; intros H'13 H'14; clear E H'10 | clear H'10 ];
 auto.
2: apply Rlt_trans with (1 := H'11); auto; apply (FPredLt b radix precision);
    auto with zarith.
cut (0 <= Fnormalize radix b precision r')%R; [ intros Rle2 | idtac ].
2: rewrite (FnormalizeCorrect radix); auto with arith.
2: unfold FtoRradix in H'14; rewrite H'14.
2: apply Rplus_le_reg_l with (r := FtoR radix p).
2: replace (FtoR radix p + 0)%R with (FtoR radix p); [ idtac | ring ].
2: replace (FtoR radix p + (FtoR radix q0 - FtoR radix p))%R with
    (FtoR radix q0); [ idtac | ring ].
2: apply Rle_trans with (2%nat * p)%R; auto.
2: apply Rledouble; auto with real zarith.
2: apply Rlt_le; apply Rlt_trans with (1 := H'11); auto.
2: apply (FPredLt b radix precision); auto with zarith.
cut (Fnormalize radix b precision r' < q0)%R; [ intros Rle3 | idtac ].
2: rewrite (FnormalizeCorrect radix); auto with arith.
2: unfold FtoRradix in H'14; rewrite H'14.
2: apply Rplus_lt_reg_l with (r := (- q0)%R).
2: replace (- q0 + (FtoR radix q0 - FtoR radix p))%R with (- p)%R;
    [ idtac | unfold FtoRradix in |- *; ring; ring ].
2: replace (- q0 + q0)%R with (-0)%R; [ auto with real | ring ].
case radixRangeBoundExp with (b:=b) (radix:=radix) (precision:=precision) (p := Fnormalize radix b precision r') (q := q0);
 auto with zarith; fold FtoRradix in |- *.
apply FnormalizeCanonic; auto with zarith.
rewrite (FnormalizeCorrect radix); auto with arith.
apply Rlt_le_trans with (2%nat * r')%R; auto.
rewrite H'14.
rewrite Rmult_minus_distr_l.
pattern (FtoRradix q0) at 1 in |- *;
 (replace (FtoRradix q0) with (2%nat * q0 - q0)%R;
   [ idtac | simpl in |- *; ring ]).
unfold Rminus in |- *; apply Rplus_lt_compat_l; apply Ropp_lt_contravar.
apply Rlt_trans with (1 := H'11).
apply (FPredLt b radix precision); auto with zarith.
apply Rmult_le_compat_r; auto with real arith.
unfold FtoRradix in Rle2; rewrite (FnormalizeCorrect radix) in Rle2;
 auto with arith.
rewrite INR_IZR_INZ; apply Rle_IZR; simpl; omega.
intros H'10.
case
 (FcanonicLtPos _ radixMoreThanOne b precision)
  with (p := Fnormalize radix b precision r') (q := q0);
 auto with zarith.
apply FnormalizeCanonic; auto with zarith.
intros; contradict H'10; auto with zarith.
intros H'12; elim H'12; intros H'15 H'16; clear H'12.
exists
 (Float (Z.pred (Fnum (Fnormalize radix b precision r')))
    (Fexp (Fnormalize radix b precision r'))).
split.
cut (Fbounded b (Fnormalize radix b precision r')); [ intros Fb0 | idtac ].
repeat split; simpl in |- *; auto.
case Rle2; intros Z1.
apply Z.le_lt_trans with (Z.abs (Fnum (Fnormalize radix b precision r')));
 auto with zarith.
repeat rewrite Z.abs_eq; auto with zarith.
apply (LeR0Fnum radix); auto with zarith.
apply Zle_Zpred; apply (LtR0Fnum radix); auto with zarith.
apply Fb0.
replace (Fnum (Fnormalize radix b precision r')) with 0%Z; simpl in |- *;
 auto with zarith.
apply (vNumbMoreThanOne radix) with (precision := precision);
 auto with zarith.
apply sym_equal; change (is_Fzero (Fnormalize radix b precision r')) in |- *;
 apply (is_Fzero_rep2 radix); auto with zarith.
apply FcanonicBound with (radix := radix); auto.
apply FnormalizeCanonic; auto with zarith.
apply FnormalizeBounded; auto with zarith.
replace
 (Float (Z.pred (Fnum (Fnormalize radix b precision r')))
    (Fexp (Fnormalize radix b precision r'))) with
 (Fminus radix (Fnormalize radix b precision r')
    (Fminus radix q0 (FPred b radix precision q0))).
repeat rewrite (Fopp_correct radix); repeat rewrite (Fminus_correct radix);
 auto with arith.
rewrite (FnormalizeCorrect radix); auto with arith.
unfold FtoRradix in H'14; rewrite H'14.
unfold FtoRradix in |- *; ring; ring.
replace (FPred b radix precision q0) with (Float (Z.pred (Fnum q0)) (Fexp q0));
 auto.
unfold Fminus, Fopp, Fplus in |- *; simpl in |- *.
repeat rewrite Zmin_n_n; repeat rewrite <- Zminus_diag_reverse; simpl in |- *;
 auto.
rewrite H'10.
repeat rewrite Zmin_n_n; repeat rewrite <- Zminus_diag_reverse; simpl in |- *;
 auto.
repeat rewrite Zmult_1_r.
apply floatEq; simpl in |- *; auto; unfold Z.pred in |- *; ring.
case (Z.eq_dec (Fnum q0) (nNormMin radix precision)); intros Zeq2.
case (Z.eq_dec (Fexp q0) (- dExp b)); intros Zeq1.
rewrite Zeq1; rewrite Zeq2; rewrite <- (FPredSimpl3 b radix); auto with zarith;
 rewrite <- Zeq1; rewrite <- Zeq2; auto.
contradict H'16.
apply Zle_not_lt.
rewrite Zeq2.
rewrite <- (Z.abs_eq (Fnum (Fnormalize radix b precision r')));
 auto with zarith.
apply pNormal_absolu_min with (b := b); auto with zarith.
cut (Fcanonic radix b (Fnormalize radix b precision r'));
 [ intros Ca1; case Ca1; auto | apply FnormalizeCanonic; auto with zarith ].
intros H'12; case Zeq1; rewrite <- H'10.
apply H'12.
apply (LeR0Fnum radix); auto.
rewrite FPredSimpl4; auto.
contradict H'16; rewrite H'16.
apply Zle_not_lt.
unfold pPred in |- *; rewrite Zopp_Zpred_Zs; apply Zlt_le_succ.
apply Zlt_Zabs_inv1.
cut (Fbounded b (Fnormalize radix b precision r'));
 [ intros T; apply T | idtac ].
apply FnormalizeBounded; auto with zarith.
intros H'10.
case (Z.eq_dec (Fnum q0) (nNormMin radix precision)); intros Zeq2.
exists
 (Float (Z.pred (Fnum (Fnormalize radix b precision r')))
    (Fexp (Fnormalize radix b precision r'))).
cut (Fbounded b (Fnormalize radix b precision r')); [ intros Fb1 | idtac ].
repeat split; simpl in |- *; auto with zarith.
case Rle2; intros Z1.
apply Z.lt_trans with (Z.abs (Fnum (Fnormalize radix b precision r'))).
repeat rewrite Z.abs_eq; auto with zarith.
apply (LeR0Fnum radix); auto.
apply Zle_Zpred; apply (LtR0Fnum radix); auto.
case Fb1; auto.
replace (Fnum (Fnormalize radix b precision r')) with 0%Z.
simpl in |- *; apply (vNumbMoreThanOne radix) with (precision := precision);
 auto with zarith.
apply sym_equal; change (is_Fzero (Fnormalize radix b precision r')) in |- *;
 apply (is_Fzero_rep2 radix); auto with zarith.
apply Fb1.
rewrite FPredSimpl2; auto with zarith.
rewrite <- H'10.
cut (forall z : Z, Z.pred (Z.succ z) = z);
 [ intros tmp; rewrite tmp; clear tmp
 | intros; unfold Z.succ, Z.pred in |- *; ring ].
unfold FtoRradix, FtoR in |- *; simpl in |- *.
cut (forall x : Z, Z.pred x = (x - 1%nat)%Z);
 [ intros tmp; rewrite tmp; clear tmp
 | intros; unfold Z.pred in |- *; simpl in |- *; ring ].
unfold FtoRradix, FtoR in |- *; simpl in |- *.
rewrite <- Z_R_minus; auto.
rewrite (fun x y => Rmult_comm (x - y)); rewrite Rmult_minus_distr_l;
 repeat rewrite (fun x y => Rmult_comm (powerRZ x y)).
replace
 (Fnum (Fnormalize radix b precision r') *
  powerRZ radix (Fexp (Fnormalize radix b precision r')))%R with
 (FtoRradix (Fnormalize radix b precision r')).
rewrite (FnormalizeCorrect radix); auto.
unfold FtoRradix in H'14; rewrite H'14.
unfold FtoR in |- *; simpl in |- *.
pattern (Fexp q0) at 1 in |- *; rewrite <- H'10.
rewrite Zeq2; rewrite powerRZ_Zs.
2: apply IZR_neq; omega.
rewrite <- Rmult_assoc.
replace (nNormMin radix precision * radix)%R with (powerRZ radix precision).
unfold pPred, nNormMin, Z.pred in |- *; rewrite pGivesBound.
rewrite plus_IZR; repeat rewrite Zpower_nat_Z_powerRZ; simpl in |- *; try ring.
rewrite <- Zpower_nat_Z_powerRZ; auto with zarith; rewrite <- mult_IZR;
 rewrite Zmult_comm; rewrite <- (PosNormMin radix b precision);
 auto with real zarith.
rewrite pGivesBound; easy.
auto.
red in |- *; intros H'12;
 absurd (- dExp b <= Fexp (Fnormalize radix b precision r'))%Z;
 auto with zarith.
apply Fb1.
apply FnormalizeBounded; auto with zarith.
exists
 (Float (Fnum (Fnormalize radix b precision r') - radix)
    (Fexp (Fnormalize radix b precision r'))).
cut (Fbounded b (Fnormalize radix b precision r')); [ intros Fb1 | idtac ].
repeat split; simpl in |- *; auto with zarith; try apply Fb1.
case (Zle_or_lt (Fnum (Fnormalize radix b precision r')) radix); intros Z1.
apply Z.le_lt_trans with radix.
rewrite Zabs_eq_opp; auto with zarith.
cut (0 <= Fnum (Fnormalize radix b precision r'))%Z; auto with zarith.
apply (LeR0Fnum radix); auto.
rewrite <- (Zpower_nat_1 radix); rewrite pGivesBound; auto with zarith.
apply Zpower_nat_monotone_lt; omega.
apply Z.le_lt_trans with (Z.abs (Fnum (Fnormalize radix b precision r'))).
repeat rewrite Z.abs_eq; auto with zarith.
case Fb1; auto.
rewrite FPredSimpl4; auto with arith.
rewrite <- H'10.
unfold FtoRradix, FtoR in |- *; simpl in |- *.
cut (forall x : Z, Z.pred x = (x - 1%nat)%Z);
 [ intros tmp; rewrite tmp; clear tmp
 | intros; unfold Z.pred in |- *; simpl in |- *; ring ].
repeat rewrite <- Z_R_minus; auto.
repeat rewrite (fun x y => Rmult_comm (x - y));
 repeat rewrite Rmult_minus_distr_l;
 repeat rewrite (fun x y => Rmult_comm (powerRZ x y)).
replace
 (Fnum (Fnormalize radix b precision r') *
  powerRZ radix (Fexp (Fnormalize radix b precision r')))%R with
 (FtoRradix (Fnormalize radix b precision r')).
rewrite (FnormalizeCorrect radix); auto.
unfold FtoRradix in H'14; rewrite H'14.
unfold FtoR in |- *; simpl in |- *.
rewrite <- H'10.
repeat rewrite powerRZ_Zs.
ring.
apply IZR_neq; omega.
auto with real zarith.
unfold FtoR in |- *; simpl in |- *.
red in |- *; intros H'12; absurd (0 <= Fnum q0)%Z; auto.
apply Zlt_not_le.
rewrite H'12.
replace 0%Z with (- 0%nat)%Z; [ apply Zlt_Zopp | simpl in |- *; auto ].
unfold pPred in |- *; apply Zlt_succ_pred; simpl in |- *; auto with zarith.
apply (vNumbMoreThanOne radix) with (precision := precision);
 auto with zarith.
apply (LeR0Fnum radix); auto.
apply Rlt_le; auto.
apply (FcanonicBound radix b); auto with arith.
apply FnormalizeCanonic; auto with zarith.
Qed.

Theorem ExactMinusIntervalAux1 :
 forall P,
 RoundedModeP b radix P ->
 forall p q : float,
 (0 <= p)%R ->
 (p <= q)%R ->
 Fcanonic radix b p ->
 Fcanonic radix b q ->
 (exists r : float, Fbounded b r /\ r = (q - p)%R :>R) ->
 forall r : float,
 Fcanonic radix b r ->
 (p <= r)%R ->
 (r <= q)%R -> exists r' : float, Fbounded b r' /\ r' = (r - p)%R :>R.
intros P H' p q H'0 H'1 H'2 H'3 H'4 r H'5 H'6 H'7.
Casec H'0; intros H'0.
case (Rle_or_lt q (2%nat * p)); intros Rl1.
exists (Fminus radix r p); split; auto.
rewrite <- Fopp_Fminus.
apply oppBounded.
apply Sterbenz; auto.
apply (FcanonicBound radix b); auto with arith.
apply (FcanonicBound radix b); auto with arith.
apply Rmult_le_reg_l with (r := INR 2); auto with real.
rewrite <- Rmult_assoc; rewrite Rinv_r; auto with real arith;
 rewrite Rmult_1_l; auto.
apply Rle_trans with (1 := H'7); auto.
apply Rle_trans with (1 := H'6); auto.
apply Rledouble; auto.
apply Rle_trans with (2 := H'6); apply Rlt_le; auto.
rewrite (Fminus_correct radix); auto with arith.
case (Rle_or_lt r (2%nat * p)); intros Rl2.
exists (Fminus radix r p); split; auto.
rewrite <- Fopp_Fminus.
apply oppBounded.
apply Sterbenz; auto.
apply (FcanonicBound radix b); auto with arith.
apply (FcanonicBound radix b); auto with arith.
apply Rmult_le_reg_l with (r := INR 2); auto with real.
rewrite <- Rmult_assoc; rewrite Rinv_r; auto with real arith;
 rewrite Rmult_1_l; auto.
apply Rle_trans with (1 := H'6); auto.
apply Rledouble; auto.
apply Rle_trans with (2 := H'6); apply Rlt_le; auto.
rewrite (Fminus_correct radix); auto with arith.
apply ExactMinusIntervalAux with (P := P) (q := q); auto.
exists r; split; auto.
apply (FcanonicBound radix b); auto with arith.
rewrite <- H'0; ring.
Qed.

Theorem ExactMinusInterval :
 forall P,
 RoundedModeP b radix P ->
 forall p q : float,
 (0 <= p)%R ->
 (p <= q)%R ->
 Fbounded b p ->
 Fbounded b q ->
 (exists r : float, Fbounded b r /\ r = (q - p)%R :>R) ->
 forall r : float,
 Fbounded b r ->
 (p <= r)%R ->
 (r <= q)%R -> exists r' : float, Fbounded b r' /\ r' = (r - p)%R :>R.
intros P H' p q H'0 H'1 H'2 H'3 H'4 r H'5 H'6 H'7.
replace (FtoRradix r) with (FtoRradix (Fnormalize radix b precision r));
 [ idtac | apply (FnormalizeCorrect radix) ]; auto.
replace (FtoRradix p) with (FtoRradix (Fnormalize radix b precision p));
 [ idtac | apply (FnormalizeCorrect radix) ]; auto.
apply
 ExactMinusIntervalAux1 with (P := P) (q := Fnormalize radix b precision q);
 auto; try repeat rewrite (FnormalizeCorrect radix);
 auto; apply FnormalizeCanonic; auto with zarith.
Qed.

Theorem MSBroundLSB :
 forall P : R -> float -> Prop,
 RoundedModeP b radix P ->
 forall f1 f2 : float,
 P f1 f2 ->
 ~ is_Fzero (Fminus radix f1 f2) ->
 (MSB radix (Fminus radix f1 f2) < LSB radix f2)%Z.
intros P H' f1 f2 H'0 HZ.
apply (oneExp_Zlt radix); auto.
apply Rlt_le_trans with (Fulp b radix precision f2).
apply Rle_lt_trans with (FtoRradix (Fabs (Fminus radix f1 f2))).
unfold FtoRradix in |- *; apply MSB_le_abs; auto.
unfold FtoRradix in |- *; rewrite Fabs_correct; auto with arith;
 rewrite Fminus_correct; auto with arith.
apply RoundedModeUlp with (4 := H'); auto.
apply FUlp_Le_LSigB; auto.
apply RoundedModeBounded with (1 := H') (2 := H'0); auto.
Qed.

Theorem LSBMinus :
 forall p q : float,
 ~ is_Fzero (Fminus radix p q) ->
 (Z.min (LSB radix p) (LSB radix q) <= LSB radix (Fminus radix p q))%Z.
intros p q H'1.
elim (LSB_rep_min radix) with (p := p); auto; intros z E.
elim (LSB_rep_min radix) with (p := q); auto; intros z0 E0.
replace (LSB radix (Fminus radix p q)) with
 (LSB radix (Fminus radix (Float z (LSB radix p)) (Float z0 (LSB radix q)))).
replace (Z.min (LSB radix p) (LSB radix q)) with
 (Fexp (Fminus radix (Float z (LSB radix p)) (Float z0 (LSB radix q))));
 [ idtac | simpl in |- *; auto ].
apply Fexp_le_LSB; auto.
apply sym_equal; apply LSB_comp; auto.
repeat rewrite Fminus_correct; auto with arith.
unfold FtoRradix in E; unfold FtoRradix in E0; rewrite E; rewrite E0; auto.
Qed.

Theorem LSBPlus :
 forall p q : float,
 ~ is_Fzero (Fplus radix p q) ->
 (Z.min (LSB radix p) (LSB radix q) <= LSB radix (Fplus radix p q))%Z.
intros p q H'.
elim (LSB_rep_min _ radixMoreThanOne p); intros z E.
elim (LSB_rep_min _ radixMoreThanOne q); intros z0 E0.
replace (LSB radix (Fplus radix p q)) with
 (LSB radix (Fplus radix (Float z (LSB radix p)) (Float z0 (LSB radix q)))).
replace (Z.min (LSB radix p) (LSB radix q)) with
 (Fexp (Fplus radix (Float z (LSB radix p)) (Float z0 (LSB radix q))));
 [ idtac | simpl in |- *; auto ].
apply Fexp_le_LSB; auto.
apply sym_equal; apply LSB_comp; auto.
repeat rewrite Fplus_correct; auto with arith.
unfold FtoRradix in E; unfold FtoRradix in E0; rewrite E; rewrite E0; auto.
Qed.

End FRoundPN.

Section ClosestP.
Variable b : Fbound.
Variable radix : Z.
Variable precision : nat.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.

Hypothesis radixMoreThanOne : (1 < radix)%Z.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ radixMoreThanOne).

Hypothesis precisionGreaterThanOne : 1 < precision.
Hypothesis pGivesBound : Zpos (vNum b) = Zpower_nat radix precision.

Theorem errorBoundedPlusLe :
 forall p q pq : float,
 Fbounded b p ->
 Fbounded b q ->
 (Fexp p <= Fexp q)%Z ->
 Closest b radix (p + q) pq ->
 exists error : float,
   error = Rabs (p + q - pq) :>R /\
   Fbounded b error /\ Fexp error = Z.min (Fexp p) (Fexp q).
intros p q pq H' H'0 H'1 H'2.
cut (ex (fun m : Z => pq = Float m (Fexp (Fplus radix p q)) :>R)).
2: unfold FtoRradix in |- *;
    apply
     RoundedModeRep
      with (b := b) (precision := precision) (P := Closest b radix);
    auto.
2: apply ClosestRoundedModeP with (precision := precision); auto.
2: rewrite (Fplus_correct radix); auto with arith.
intros H'3; elim H'3; intros m E; clear H'3.
exists
 (Fabs (Fminus radix q (Fminus radix (Float m (Fexp (Fplus radix p q))) p))).
cut (forall A B : Prop, A -> (A -> B) -> A /\ B);
 [ intros tmp; apply tmp; clear tmp | auto ].
unfold FtoRradix in |- *; rewrite Fabs_correct; auto with arith.
cut (forall p q : R, p = q -> Rabs p = Rabs q);
 [ intros tmp; apply tmp; clear tmp | intros p' q' H; rewrite H; auto ].
unfold FtoRradix in |- *; repeat rewrite Fminus_correct; auto with arith.
unfold FtoRradix in E; rewrite E; auto.
ring.
intros H'4.
cut (Rabs (pq - (p + q)) <= Rabs (q - (p + q)))%R.
2: elim H'2; auto.
replace (q - (p + q))%R with (- FtoRradix p)%R.
2: ring.
rewrite Rabs_Ropp.
unfold FtoRradix in |- *; rewrite <- Fabs_correct; auto with arith.
rewrite <- Rabs_Ropp; rewrite Ropp_minus_distr.
unfold FtoRradix in H'4; rewrite <- H'4.
simpl in |- *.
rewrite Zmin_le1; auto.
generalize H'1 H'; case p; case q; unfold Fabs, Fminus, Fopp, Fplus in |- *;
 simpl in |- *; clear H'1 H'.
intros Fnum1 Fexp1 Fnum2 Fexp2 H'5 H'6.
repeat rewrite Zmin_n_n; auto.
repeat rewrite (Zmin_le2 _ _ H'5); auto with zarith.
replace (Z.abs_nat (Fexp2 - Fexp2)) with 0.
rewrite Zpower_nat_O.
cut (forall z : Z, (z * 1%nat)%Z = z);
 [ intros tmp; repeat rewrite tmp; clear tmp | auto with zarith ].
unfold FtoRradix, FtoR in |- *; simpl in |- *.
intros H'.
repeat split; simpl in |- *.
rewrite (fun x => Z.abs_eq (Z.abs x)); auto with zarith.
apply Z.le_lt_trans with (Z.abs Fnum2); auto.
apply le_IZR.
apply (Rle_monotony_contra_exp radix) with (z := Fexp2); auto.
case H'6; auto.
case H'6; auto.
replace (Fexp2 - Fexp2)%Z with 0%Z; simpl in |- *; auto with zarith.
Qed.

Theorem errorBoundedPlusAbs :
 forall p q pq : float,
 Fbounded b p ->
 Fbounded b q ->
 Closest b radix (p + q) pq ->
 exists error : float,
   error = Rabs (p + q - pq) :>R /\
   Fbounded b error /\ Fexp error = Z.min (Fexp p) (Fexp q).
intros p q pq H' H'0 H'1.
case (Zle_or_lt (Fexp p) (Fexp q)); intros H'2.
apply errorBoundedPlusLe; auto.
replace (p + q)%R with (q + p)%R; [ idtac | ring ].
replace (Z.min (Fexp p) (Fexp q)) with (Z.min (Fexp q) (Fexp p));
 [ idtac | apply Zmin_sym ].
apply errorBoundedPlusLe; auto.
auto with zarith.
apply (ClosestCompatible b radix (p + q)%R (q + p)%R pq); auto.
ring.
case H'1; auto.
Qed.

Theorem errorBoundedPlus :
 forall p q pq : float,
 (Fbounded b p) ->
 (Fbounded b q) ->
 (Closest b radix (p + q) pq) ->
 exists error : float,
   error = (p + q - pq)%R :>R /\
   (Fbounded b error) /\ (Fexp error) = (Z.min (Fexp p) (Fexp q)).
intros p q pq H' H'0 H'1.
case (errorBoundedPlusAbs p q pq); auto.
intros x H'2; elim H'2; intros H'3 H'4; elim H'4; intros H'5 H'6;
 clear H'4 H'2.
generalize H'3; clear H'3.
unfold Rabs in |- *; case (Rcase_abs (p + q - pq)).
intros H'2 H'3; exists (Fopp x); split; auto.
unfold FtoRradix in |- *; rewrite Fopp_correct; auto.
unfold FtoRradix in H'3; rewrite H'3; ring.
split.
apply oppBounded; auto.
rewrite <- H'6; auto.
intros H'2 H'3; exists x; split; auto.
Qed.

Theorem plusExact1 :
 forall p q r : float,
 Fbounded b p ->
 Fbounded b q ->
 Closest b radix (p + q) r ->
 (Fexp r <= Z.min (Fexp p) (Fexp q))%Z -> r = (p + q)%R :>R.
intros p q r H' H'0 H'1 H'2.
cut
 (2%nat * Rabs (FtoR radix (Fplus radix p q) - FtoR radix r) <=
  Float 1%nat (Fexp r))%R;
 [ rewrite Fplus_correct; auto with zarith; intros Rl1 | idtac ].
case errorBoundedPlus with (p := p) (q := q) (pq := r); auto.
intros x H'3; elim H'3; intros H'4 H'5; elim H'5; intros H'6 H'7;
 clear H'5 H'3.
unfold FtoRradix in H'4; rewrite <- H'4 in Rl1.
2: apply Rle_trans with (Fulp b radix precision r); auto.
2: apply (ClosestUlp b radix precision); auto.
2: rewrite Fplus_correct; auto with zarith.
2: unfold FtoRradix in |- *; apply FulpLe; auto.
2: apply
    RoundedModeBounded
     with (radix := radix) (P := Closest b radix) (r := (p + q)%R);
    auto.
2: apply ClosestRoundedModeP with (precision := precision); auto.
cut (x = 0%R :>R); [ unfold FtoRradix in |- *; intros Eq1 | idtac ].
replace (FtoR radix r) with (FtoR radix r + 0)%R; [ idtac | ring ].
rewrite <- Eq1.
rewrite H'4; ring.
apply (is_Fzero_rep1 radix).
case (Z_zerop (Fnum x)); simpl in |- *; auto.
intros H'3; contradict Rl1.
apply Rgt_not_le.
red in |- *; apply Rle_lt_trans with (Rabs (FtoR radix x)).
unfold FtoRradix, FtoR in |- *; simpl in |- *; auto.
rewrite Rabs_mult.
apply Rmult_le_compat; auto with real arith.
apply powerRZ_le, IZR_lt; omega.
rewrite Rabs_Zabs.
apply IZR_le.
assert (0 < Z.abs (Fnum x))%Z; [idtac|auto with zarith].
now apply <- Z.abs_pos.
rewrite Rabs_right; auto with real arith.
apply Rle_powerRZ; auto with real arith.
apply Rle_IZR; auto with zarith.
auto with zarith.
apply Rle_ge, powerRZ_le, Rlt_IZR; omega.
cut (forall r : R, (2%nat * r)%R = (r + r)%R);
 [ intros tmp; rewrite tmp; clear tmp | intros f; simpl in |- *; ring ].
pattern (Rabs (FtoR radix x)) at 1 in |- *;
 replace (Rabs (FtoR radix x)) with (Rabs (FtoR radix x) + 0)%R;
 [ idtac | ring ].
apply Rplus_lt_compat_l; auto.
case (Rabs_pos (FtoR radix x)); auto.
rewrite <- Fabs_correct; auto with arith.
intros H'5; contradict H'3.
cut (Fnum (Fabs x) = 0%Z).
unfold Fabs in |- *; simpl in |- *; case (Fnum x); simpl in |- *; auto;
 intros; discriminate.
change (is_Fzero (Fabs x)) in |- *.
apply (is_Fzero_rep2 radix); auto with arith.
Qed.

Theorem plusExact2Aux :
 forall p q r : float,
 (0 <= p)%R ->
 Fcanonic radix b p ->
 Fbounded b q ->
 Closest b radix (p + q) r ->
 (Fexp r < Z.pred (Fexp p))%Z -> r = (p + q)%R :>R.
intros p q r H' H'0 H'1 H'2 H'3.
apply plusExact1; auto.
apply FcanonicBound with (1 := H'0); auto.
case (Zle_or_lt (Fexp p) (Fexp q)); intros Zl1.
rewrite Zmin_le1; auto with zarith.
apply Z.le_trans with (Z.pred (Fexp p)); auto with zarith.
unfold Z.pred in |- *; auto with zarith.
rewrite Zmin_le2; auto with zarith.
case (Zlt_next _ _ Zl1); intros Zl2.
rewrite Zl2 in H'3.
replace (Fexp q) with (Z.pred (Z.succ (Fexp q))); auto with zarith;
 unfold Z.pred, Z.succ in |- *; ring.
case H'0; clear H'0; intros H'0.
absurd (r < Float (nNormMin radix precision) (Z.pred (Fexp p)))%R.
apply Rle_not_lt; auto.
unfold FtoRradix in |- *;
 apply
  (ClosestMonotone b radix
     (Float (nNormMin radix precision) (Z.pred (Fexp p))) (
     p + q)%R); auto; auto.
cut (Float (nNormMin radix precision) (Fexp p) <= p)%R;
 [ intros Eq1 | idtac ].
case (Rle_or_lt 0 q); intros Rl1.
apply Rlt_le_trans with (FtoRradix p).
apply
 Rlt_le_trans with (FtoRradix (Float (nNormMin radix precision) (Fexp p)));
 auto.
unfold FtoRradix, FtoR in |- *; simpl in |- *; auto.
apply Rmult_lt_compat_l; auto with real arith.
apply Rlt_IZR; apply nNormPos; auto with zarith.
apply Rlt_powerRZ; try apply Rlt_IZR; omega.
pattern (FtoRradix p) at 1 in |- *; replace (FtoRradix p) with (p + 0)%R;
 auto with real.
apply Rplus_lt_reg_l with (r := (- q)%R); auto.
replace (- q + (p + q))%R with (FtoRradix p); [ idtac | ring ].
apply
 Rlt_le_trans with (FtoRradix (Float (nNormMin radix precision) (Fexp p)));
 auto.
apply
 Rlt_le_trans
  with (2%nat * Float (nNormMin radix precision) (Z.pred (Fexp p)))%R;
 auto.
cut (forall r : R, (2%nat * r)%R = (r + r)%R);
 [ intros tmp; rewrite tmp; clear tmp | intros; simpl in |- *; ring ].
rewrite (Rplus_comm (- q)).
apply Rplus_lt_compat_l.
rewrite <- Rabs_left1; auto.
rewrite <- (Fabs_correct radix); auto with arith.
unfold FtoRradix in |- *; apply maxMaxBis with (b := b); auto with zarith.
apply Rlt_le; auto.
apply
 Rle_trans with (radix * Float (nNormMin radix precision) (Z.pred (Fexp p)))%R.
apply Rmult_le_compat_r; auto.
apply (LeFnumZERO radix); simpl in |- *; auto with arith.
apply Zlt_le_weak; apply nNormPos; auto with zarith.
rewrite INR_IZR_INZ; apply Rle_IZR; simpl in |- *; cut (1 < radix)%Z;
 auto with real zarith.
pattern (Fexp p) at 2 in |- *; replace (Fexp p) with (Z.succ (Z.pred (Fexp p)));
 [ idtac | unfold Z.succ, Z.pred in |- *; ring ].
unfold FtoRradix, FtoR in |- *; simpl in |- *.
rewrite powerRZ_Zs; auto with real zarith.
repeat rewrite <- Rmult_assoc.
rewrite (Rmult_comm radix); auto with real.
apply IZR_neq; omega.
unfold FtoRradix, FtoR in |- *; simpl in |- *; auto.
apply Rmult_le_compat_r; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
apply Rle_IZR.
rewrite <- (Z.abs_eq (Fnum p)); auto with zarith.
apply pNormal_absolu_min with (b := b); auto with zarith.
apply (LeR0Fnum radix); auto with arith.
apply (RoundedModeProjectorIdem b radix (Closest b radix)); auto.
apply ClosestRoundedModeP with (precision := precision); auto.
repeat split; simpl in |- *.
rewrite Z.abs_eq; auto with zarith.
apply ZltNormMinVnum; auto with zarith.
apply Zlt_le_weak; apply nNormPos; auto with zarith.
apply Z.le_trans with (Fexp q); auto with zarith; apply H'1.
case (Rle_or_lt 0 r); intros Rl1.
rewrite <- (Rabs_right r); auto with real.
rewrite <- (Fabs_correct radix); auto with arith.
unfold FtoRradix in |- *; apply maxMaxBis with (b := b); auto with zarith.
apply
 RoundedModeBounded
  with (radix := radix) (P := Closest b radix) (r := (p + q)%R);
 auto.
apply ClosestRoundedModeP with (precision := precision); auto with real.
apply Rlt_le_trans with 0%R; auto.
apply (LeFnumZERO radix); simpl in |- *; auto with arith.
apply Zlt_le_weak; apply nNormPos; auto with zarith.
absurd (- dExp b <= Fexp q)%Z; try apply H'1.
apply Zlt_not_le.
case H'0; intros Z1 (Z2, Z3); rewrite <- Z2; auto with zarith.
Qed.

Theorem plusExact2 :
 forall p q r : float,
 Fcanonic radix b p ->
 Fbounded b q ->
 Closest b radix (p + q) r ->
 (Fexp r < Z.pred (Fexp p))%Z -> r = (p + q)%R :>R.
intros p q r H' H'0 H'1 H'2.
case (Rle_or_lt 0 p); intros Rl1.
apply plusExact2Aux; auto.
replace (p + q)%R with (- (Fopp p + Fopp q))%R.
rewrite <- (plusExact2Aux (Fopp p) (Fopp q) (Fopp r)); auto.
unfold FtoRradix in |- *; rewrite Fopp_correct; ring.
unfold FtoRradix in |- *; rewrite Fopp_correct.
apply Rlt_le; replace 0%R with (-0)%R; auto with real.
apply FcanonicFopp; auto with arith.
apply oppBounded; auto.
replace (Fopp p + Fopp q)%R with (- (p + q))%R.
apply ClosestOpp; auto.
unfold FtoRradix in |- *; repeat rewrite Fopp_correct; ring.
unfold FtoRradix in |- *; repeat rewrite Fopp_correct; ring.
Qed.

Theorem plusExactR0 :
 forall p q r : float,
 Fbounded b p ->
 Fbounded b q ->
 Closest b radix (p + q) r -> r = 0%R :>R -> r = (p + q)%R :>R.
intros p q r H' H'0 H'1 H'2.
cut (r = FtoRradix (Fzero (- dExp b)) :>R);
 [ intros Eq1; rewrite Eq1
 | rewrite H'2; apply sym_eq; unfold FtoRradix in |- *; apply FzeroisZero ].
apply plusExact1; auto.
apply (ClosestCompatible b radix (p + q)%R (p + q)%R r); auto.
apply FboundedFzero; auto.
simpl in |- *; auto.
unfold Z.min in |- *; case (Fexp p ?= Fexp q)%Z; auto with zarith; try apply H'; try apply H'0.
Qed.

Theorem pPredMoreThanOne : (0 < pPred (vNum b))%Z.
unfold pPred in |- *; apply Zlt_succ_pred; simpl in |- *.
apply (vNumbMoreThanOne radix) with (precision := precision);
 auto with zarith.
Qed.

Theorem pPredMoreThanRadix : (radix < pPred (vNum b))%Z.
apply Z.le_lt_trans with (nNormMin radix precision).
pattern radix at 1 in |- *; rewrite <- (Zpower_nat_1 radix);
 unfold nNormMin in |- *; auto with zarith.
apply Zpower_nat_monotone_le; omega.
apply nNormMimLtvNum; auto with zarith.
Qed.

Theorem plusExactExp :
 forall p q pq : float,
 Fbounded b p ->
 Fbounded b q ->
 Closest b radix (p + q) pq ->
 ex
   (fun r : float =>
    ex
      (fun s : float =>
       Fbounded b r /\
       Fbounded b s /\
       s = pq :>R /\
       r = (p + q - s)%R :>R /\
       Fexp r = Z.min (Fexp p) (Fexp q) :>Z /\
       (Fexp r <= Fexp s)%Z /\ (Fexp s <= Z.succ (Zmax (Fexp p) (Fexp q)))%Z)).
intros p q pq H H0 H1.
case (plusExpBound b radix precision) with (P := Closest b radix) (5 := H1);
 auto with zarith.
apply (ClosestRoundedModeP b radix precision); auto with zarith.
intros r (H2, (H3, (H4, H5))); fold FtoRradix in H3.
case (Req_dec (p + q - pq) 0); intros Hr.
cut (Fbounded b (Fzero (Z.min (Fexp p) (Fexp q)))); [ intros Fbs | idtac ].
exists (Fzero (Z.min (Fexp p) (Fexp q))); exists r; repeat (split; auto).
rewrite (FzeroisReallyZero radix); rewrite <- Hr; rewrite <- H3; auto.
case (Zmin_or (Fexp p) (Fexp q)); intros Hz; rewrite Hz;
 apply FboundedZeroSameExp; auto.
case (errorBoundedPlus p q pq); auto.
intros error (H6, (H7, H8)).
exists error; exists r; repeat (split; auto).
rewrite H3; auto.
rewrite H8; auto.
Qed.

End ClosestP.


Section MinOrMax_def.
Variable radix : Z.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.
Variable b : Fbound.
Variable precision : nat.
Hypothesis radixMoreThanOne : (1 < radix)%Z.
Hypothesis precisionGreaterThanOne : 1 < precision.
Hypothesis pGivesBound : Zpos (vNum b) = Zpower_nat radix precision.

Definition MinOrMax (z : R) (f : float) :=
  isMin b radix z f \/ isMax b radix z f.

Theorem MinOrMax_Fopp :
 forall (x : R) (f : float), MinOrMax (- x) (Fopp f) -> MinOrMax x f.
unfold MinOrMax in |- *; intros x f H.
rewrite <- (Ropp_involutive x); rewrite <- (Fopp_Fopp f).
case H; intros H1.
right; apply MinOppMax; auto.
left; apply MaxOppMin; auto.
Qed.

Theorem MinOrMax1 :
 forall (z : R) (p : float),
 Fbounded b p ->
 Fcanonic radix b p ->
 (0 < p)%R ->
 (Rabs (z - p) < Fulp b radix precision (FPred b radix precision p))%R ->
 MinOrMax z p.
intros z p Hp Hcan Hblop.
case (Rcase_abs (z - p)); intros H1;
 [ rewrite Rabs_left | rewrite Rabs_right ]; auto;
 intros H.
right; unfold isMax in |- *.
split; auto; split.
fold FtoRradix in |- *; apply Rplus_le_reg_l with (- p)%R.
ring_simplify (- p + p)%R; rewrite Rplus_comm; auto with real.
intros f Hf H2.
replace (FtoR radix f) with (FtoR radix (Fnormalize radix b precision f));
 [ idtac | apply FnormalizeCorrect; auto with zarith arith ].
replace p with (FSucc b radix precision (FPred b radix precision p));
 [ idtac | apply FSucPred; auto with zarith arith ].
apply FSuccProp; auto with zarith arith.
apply FPredCanonic; auto with arith zarith.
apply FnormalizeCanonic; auto with arith zarith.
apply Rlt_le_trans with z.
2: rewrite FnormalizeCorrect; auto with arith zarith.
apply
 Rle_lt_trans with (p - Fulp b radix precision (FPred b radix precision p))%R.
2: apply Ropp_lt_cancel.
2: apply Rplus_lt_reg_r with (FtoRradix p).
2: ring_simplify.
2: apply Rle_lt_trans with (2 := H); right; ring.
replace (FtoRradix p) with
 (FPred b radix precision p +
  Fulp b radix precision (FPred b radix precision p))%R;
 [ right; fold FtoRradix; ring
 | unfold FtoRradix in |- *; apply FpredUlpPos; auto with zarith arith ].
left; unfold isMin in |- *.
split; auto; split.
fold FtoRradix in |- *; apply Rplus_le_reg_l with (- p)%R.
ring_simplify (- p + p)%R; rewrite Rplus_comm; auto with real.
intros f Hf H2.
replace (FtoR radix f) with (FtoR radix (Fnormalize radix b precision f));
 [ idtac | apply FnormalizeCorrect; auto with zarith arith ].
replace p with (FPred b radix precision (FSucc b radix precision p));
 [ idtac | apply FPredSuc; auto with zarith arith ].
apply FPredProp; auto with zarith arith.
apply FnormalizeCanonic; auto with arith zarith.
apply FSuccCanonic; auto with arith zarith.
apply Rle_lt_trans with z.
rewrite FnormalizeCorrect; auto with arith zarith.
apply
 Rlt_le_trans with (p + Fulp b radix precision (FPred b radix precision p))%R.
apply Rplus_lt_reg_r with (- FtoRradix p)%R.
ring_simplify.
apply Rle_lt_trans with (2 := H); right; ring.
apply Rle_trans with (FtoRradix p + Fulp b radix precision p)%R;
 [ apply Rplus_le_compat_l | idtac ].
apply LeFulpPos; auto with arith zarith.
apply FBoundedPred; auto with arith zarith.
unfold FtoRradix in |- *; apply R0RltRlePred; auto with arith zarith.
apply Rlt_le; unfold FtoRradix in |- *; apply FPredLt; auto with arith zarith.
pattern p at -3 in |- *;
 replace p with (FPred b radix precision (FSucc b radix precision p));
 [ idtac | apply FPredSuc; auto with zarith arith ].
unfold FtoRradix in |- *;
 rewrite FpredUlpPos with (x := FSucc b radix precision p);
 auto with zarith arith real.
apply FSuccCanonic; auto with zarith arith.
apply Rlt_trans with (FtoRradix p); auto with real.
unfold FtoRradix in |- *; apply FSuccLt; auto with zarith real.
Qed.

Theorem MinOrMax2 :
 forall (z : R) (p : float),
 Fbounded b p ->
 Fcanonic radix b p ->
 (0 < p)%R ->
 (Rabs (z - p) < Fulp b radix precision p)%R -> (p <= z)%R -> MinOrMax z p.
intros z p Hp1 Hp2 H H1 H2.
unfold MinOrMax in |- *; left; unfold isMin in |- *.
split; auto; split; auto.
intros f Hf H3.
replace (FtoR radix f) with (FtoR radix (Fnormalize radix b precision f));
 [ idtac | apply FnormalizeCorrect; auto with zarith arith ].
replace p with (FPred b radix precision (FSucc b radix precision p));
 [ idtac | apply FPredSuc; auto with zarith arith ].
apply FPredProp; auto with zarith arith.
apply FnormalizeCanonic; auto with arith zarith.
apply FSuccCanonic; auto with arith zarith.
apply Rle_lt_trans with z.
rewrite FnormalizeCorrect; auto with arith zarith.
apply Rlt_le_trans with (FtoRradix p + Fulp b radix precision p)%R.
apply Rplus_lt_reg_r with (- FtoRradix p)%R.
ring_simplify.
apply Rle_lt_trans with (2 := H1); right.
rewrite Rabs_right; try ring.
apply Rle_ge; apply Rplus_le_reg_l with (FtoRradix p); auto with real.
ring_simplify; auto with real.
pattern p at -3 in |- *;
 replace p with (FPred b radix precision (FSucc b radix precision p));
 [ idtac | apply FPredSuc; auto with zarith arith ].
unfold FtoRradix in |- *;
 rewrite FpredUlpPos with (x := FSucc b radix precision p);
 auto with zarith arith real.
apply FSuccCanonic; auto with zarith arith.
apply Rlt_trans with (FtoRradix p); auto with real.
unfold FtoRradix in |- *; apply FSuccLt; auto with zarith real.
Qed.

Theorem MinOrMax3_aux :
 forall (z : R) (p : float),
 Fbounded b p ->
 Fcanonic radix b p ->
 0%R = p ->
 (z <= 0)%R ->
 (- z < Fulp b radix precision (FPred b radix precision p))%R -> MinOrMax z p.
intros z p Hp Hcan Hzero H.
right; unfold isMax in |- *.
fold FtoRradix in |- *; rewrite <- Hzero.
split; auto; split; auto with real.
intros f Hf H2.
unfold FtoRradix in |- *;
 replace (FtoR radix f) with (FtoR radix (Fnormalize radix b precision f));
 [ idtac | apply FnormalizeCorrect; auto with zarith arith ].
cut (Fcanonic radix b (Fzero (- dExp b))); [ intros H3 | idtac ].
2: right; repeat (split; simpl in |- *; auto with zarith).
cut (p = Fzero (- dExp b)); [ intros H4 | idtac ].
2: apply
    FcanonicUnique with (radix := radix) (precision := precision) (b := b);
    auto with zarith.
2: fold FtoRradix in |- *; rewrite <- Hzero; auto with real.
2: unfold FtoRradix in |- *; apply sym_eq; apply FzeroisZero.
cut (0 < pPred (vNum b))%Z;
 [ intros V | apply pPredMoreThanOne with radix precision; auto with zarith ].
cut (1 < Zpos (vNum b))%Z;
 [ intros V'
 | apply vNumbMoreThanOne with radix precision; auto with zarith ].
replace 0%R with
 (FtoR radix
    (FSucc b radix precision (FPred b radix precision (Fzero (- dExp b))))).
apply FSuccProp; auto with zarith arith.
apply FPredCanonic; auto with arith zarith.
apply FnormalizeCanonic; auto with arith zarith.
apply Rlt_le_trans with z.
2: rewrite FnormalizeCorrect; auto with arith zarith.
apply Ropp_lt_cancel; apply Rlt_le_trans with (1 := H0).
rewrite H4; right.
rewrite FPredSimpl4; simpl in |- *; auto with zarith.
2: unfold nNormMin in |- *; auto with zarith.
unfold Fulp in |- *.
replace (Fnormalize radix b precision (Float (-1) (- dExp b))) with
 (Float (-1) (- dExp b));
 [ unfold FtoRradix, FtoR in |- *; simpl in |- *; ring | idtac ].
apply FcanonicUnique with (radix := radix) (precision := precision) (b := b);
 auto with zarith.
right; repeat (split; simpl in |- *; auto with zarith).
rewrite pGivesBound.
rewrite <- Zabs_Zopp; rewrite Z.abs_eq; auto with zarith.
replace (- (radix * -1))%Z with (Zpower_nat radix 1); auto with arith zarith.
apply Zpower_nat_monotone_lt; omega.
unfold Zpower_nat in |- *; simpl in |- *; ring.
apply FnormalizeCanonic; auto with zarith.
repeat (split; simpl in |- *; auto with zarith).
apply sym_eq; apply FnormalizeCorrect; auto with zarith.
now apply Zlt_not_eq, Zpower_nat_less.
rewrite FPredSimpl4; simpl in |- *; auto with zarith.
rewrite FSuccSimpl4; simpl in |- *; auto with zarith.
simpl in |- *; unfold FtoR in |- *; simpl in |- *; ring.
unfold nNormMin in |- *.
replace (-1)%Z with (- Zpower_nat radix (pred 1))%Z; auto with zarith arith.
apply sym_not_eq, Zlt_not_eq, Zlt_Zopp.
apply Zpower_nat_monotone_lt; omega.
unfold nNormMin in |- *; auto with zarith arith.
now apply Zlt_not_eq, Zpower_nat_less.
Qed.

Theorem MinOrMax3 :
 forall (z : R) (p : float),
 Fbounded b p ->
 Fcanonic radix b p ->
 0%R = p ->
 (Rabs (z - p) < Fulp b radix precision (FPred b radix precision p))%R ->
 MinOrMax z p.
intros z p Hp Hcan Hzero.
replace (z - FtoRradix p)%R with z; [ idtac | rewrite <- Hzero; ring ].
case (Rcase_abs z); intros H1; [ rewrite Rabs_left | rewrite Rabs_right ];
 auto; intros H.
apply MinOrMax3_aux; auto with real.
apply MinOrMax_Fopp.
apply MinOrMax3_aux; auto with real.
apply oppBounded; easy.
apply FcanonicFopp; easy.
unfold FtoRradix in |- *; rewrite Fopp_correct; auto with real.
rewrite Ropp_involutive.
replace (Fopp p) with p; auto.
apply floatEq; auto; simpl in |- *.
cut (is_Fzero p);
 [ unfold is_Fzero in |- *; intros H2; repeat rewrite H2
 | apply is_Fzero_rep2 with radix ]; auto with zarith real.
Qed.
End MinOrMax_def.

Section AxpyMisc.

Let radix := 2%Z.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.
Variable b : Fbound.
Variable precision : nat.
Hypothesis precisionGreaterThanOne : 1 < precision.
Hypothesis pGivesBound : Zpos (vNum b) = Zpower_nat radix precision.

Theorem TwoMoreThanOne : (1 < radix)%Z.
unfold radix in |- *; red in |- *; simpl in |- *; auto.
Qed.

Theorem TwoMoreThanOneR : (1 < radix)%R.
replace 1%R with (IZR 1); auto with real.
Qed.

Theorem FulpLeGeneral :
 forall p : float,
 Fbounded b p ->
 (Fulp b radix precision p <=
  Rabs (FtoRradix p) * powerRZ radix (Z.succ (- precision)) +
  powerRZ radix (- dExp b))%R.
intros p Hp.
cut (Fcanonic radix b (Fnormalize radix b precision p));
 [ intros H | apply FnormalizeCanonic; auto with arith zarith; try apply TwoMoreThanOne ].
case H; intros H1.
apply
 Rle_trans with (Rabs (FtoR radix p) * powerRZ radix (Z.succ (- precision)))%R.
apply FulpLe2; auto with zarith.
try apply TwoMoreThanOne.
apply Rle_trans with (Rabs p * powerRZ radix (Z.succ (- precision)) + 0)%R;
 [ right; fold FtoRradix; ring | apply Rplus_le_compat_l; auto with real zarith ].
apply powerRZ_le, Rlt_IZR, TwoMoreThanOne.
apply Rle_trans with (powerRZ radix (- dExp b)).
right; unfold Fulp in |- *; simpl in |- *.
replace (Fexp (Fnormalize radix b precision p)) with (- dExp b)%Z;
 [ auto with real zarith | elim H1; intuition ].
apply Rle_trans with (0 + powerRZ radix (- dExp b))%R;
 [ right; ring | apply Rplus_le_compat_r ].
apply Rmult_le_pos; try apply Rabs_pos.
apply powerRZ_le, Rlt_IZR, TwoMoreThanOne.
Qed.

Theorem RoundLeGeneral :
 forall (p : float) (z : R),
 Fbounded b p ->
 Closest b radix z p ->
 (Rabs p <=
  Rabs z * / (1 - powerRZ radix (- precision)) +
  powerRZ radix (Z.pred (- dExp b)) * / (1 - powerRZ radix (- precision)))%R.
intros p z Hp H.
cut (0 < 1 - powerRZ radix (- precision))%R; [ intros H1 | idtac ].
2: apply Rplus_lt_reg_r with (powerRZ radix (- precision)).
2: ring_simplify.
2: replace 1%R with (powerRZ radix 0);
    [ auto with real zarith | simpl in |- *; auto ].
2: apply Rlt_powerRZ; try omega; try apply TwoMoreThanOneR.
apply Rmult_le_reg_l with (1 - powerRZ radix (- precision))%R; auto.
ring_simplify ((1 - powerRZ radix (- precision)) * Rabs p)%R.
apply
 Rle_trans
  with
    (Rabs z *
     ((1 - powerRZ radix (- precision)) * / (1 - powerRZ radix (- precision))) +
     powerRZ radix (Z.pred (- dExp b)) *
     ((1 - powerRZ radix (- precision)) * / (1 - powerRZ radix (- precision))))%R;
 [ idtac | right; ring; ring ].
repeat rewrite Rinv_r; auto with real.
ring_simplify.
apply Rplus_le_reg_l with (- powerRZ radix (Z.pred (- dExp b)))%R.
ring_simplify.
apply
 Rle_trans
  with
    (Rabs p +
     -
     (Rabs p * powerRZ radix (- precision) + powerRZ radix (Z.pred (- dExp b))))%R;
 [ right; ring | idtac ].
apply Rle_trans with (Rabs p + - (/ radix * Fulp b radix precision p))%R.
apply Rplus_le_compat_l; apply Ropp_le_contravar.
apply Rmult_le_reg_l with (IZR radix); auto with real zarith.
rewrite <- Rmult_assoc; rewrite Rinv_r; auto with real zarith.
ring_simplify (1 * Fulp b radix precision p)%R.
apply
 Rle_trans
  with
    (Rabs p * (powerRZ radix 1 * powerRZ radix (- precision)) +
     powerRZ radix 1 * powerRZ radix (Z.pred (- dExp b)))%R;
 [ idtac | right; simpl in |- *; ring ].
repeat rewrite <- powerRZ_add; auto with zarith real.
replace (1 + - precision)%Z with (Z.succ (- precision)); auto with zarith.
replace (1 + Z.pred (- dExp b))%Z with (- dExp b)%Z; auto with zarith.
apply FulpLeGeneral; auto.
unfold Z.pred in |- *; auto with zarith.
apply Rplus_le_reg_l with (- Rabs z + / radix * Fulp b radix precision p)%R.
ring_simplify.
apply Rle_trans with (Rabs (p - z)).
apply Rle_trans with (Rabs p - Rabs z)%R;
 [ right; ring | apply Rabs_triang_inv ].
rewrite <- Rabs_Ropp.
replace (- (p - z))%R with (z - p)%R; [ idtac | ring ].
apply Rmult_le_reg_l with (IZR radix); auto with real zarith.
apply Rle_trans with (Fulp b radix precision p * (radix * / radix))%R;
 [ rewrite Rinv_r; auto with real zarith | right; ring ].
ring_simplify (Fulp b radix precision p * 1)%R.
unfold FtoRradix in |- *; apply ClosestUlp; auto.
try apply TwoMoreThanOne.
Qed.

Theorem ExactSum_Near :
 forall p q f : float,
 Fbounded b p ->
 Fbounded b q ->
 Fbounded b f ->
 Closest b radix (p + q) f ->
 Fexp p = (- dExp b)%Z ->
 (Rabs (p + q - f) < powerRZ radix (- dExp b))%R -> (p + q)%R = f.
intros p q f Hp Hq Hf H H12 H1.
case
 errorBoundedPlus
  with
    (b := b)
    (radix := radix)
    (precision := precision)
    (p := p)
    (q := q)
    (pq := f); auto with zarith.
try apply TwoMoreThanOne.
intros x0 H'; elim H'; intros H2 H'1; elim H'1; intros H3 H4; clear H' H'1.
apply Rplus_eq_reg_l with (- FtoRradix f)%R; ring_simplify (-f+f)%R.
apply trans_eq with (FtoR radix p + FtoR radix q - FtoR radix f)%R;
 [ fold FtoRradix; ring | rewrite <- H2 ].
generalize H1; unfold FtoRradix in |- *; rewrite <- H2; intros H5.
cut (forall r : R, Rabs r = 0%R -> r = 0%R);
 [ intros V; apply V | auto with real ].
rewrite <- Fabs_correct; auto with zarith.
apply is_Fzero_rep1; unfold is_Fzero in |- *.
cut (forall z : Z, (0 <= z)%Z -> (z < 1)%Z -> z = 0%Z);
 [ intros W; apply W | auto with zarith ].
unfold Fabs in |- *; simpl in |- *; apply Zle_ZERO_Zabs.
replace 1%Z with (Fnum (Float 1 (- dExp b))); [ idtac | simpl in |- *; auto ].
apply Rlt_Fexp_eq_Zlt with (radix := radix); auto with zarith real.
try apply TwoMoreThanOne.
rewrite Fabs_correct; auto with zarith; try apply TwoMoreThanOne; apply Rlt_le_trans with (1 := H5).
right; unfold FtoR in |- *; simpl in |- *; ring.
unfold Fabs in |- *; simpl in |- *.
rewrite H4; rewrite H12; apply Zmin_le1; auto with zarith.
elim Hq; auto with zarith.
try apply TwoMoreThanOne.
intros m; case (Rle_or_lt 0 m); intros H'.
rewrite Rabs_right; auto with real.
rewrite Rabs_left1; auto with real; intros H'1.
rewrite <- (Ropp_involutive m); rewrite H'1; auto with real.
Qed.
End AxpyMisc.
Section AxpyAux.

Let radix := 2%Z.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.
Variable b : Fbound.
Variable precision : nat.
Hypothesis precisionGreaterThanOne : 1 < precision.
Hypothesis pGivesBound : Zpos (vNum b) = Zpower_nat radix precision.
Variables a1 x1 y1 : R.
Variables a x y t u r : float.
Hypothesis Fa : Fbounded b a.
Hypothesis Fx : Fbounded b x.
Hypothesis Fy : Fbounded b y.
Hypothesis Ft : Fbounded b t.
Hypothesis Fu : Fbounded b u.
Hypothesis tDef : Closest b radix (a * x) t.
Hypothesis uDef : Closest b radix (t + y) u.
Hypothesis rDef : isMin b radix (a1 * x1 + y1) r.

Theorem Axpy_aux1 :
 Fcanonic radix b u ->
 (Rabs (FtoRradix a * FtoRradix x - FtoRradix t) <=
  / 4%nat * Fulp b radix precision (FPred b radix precision u))%R ->
 (0 < u)%R ->
 (4%nat * Rabs t <= Rabs u)%R ->
 (Rabs (y1 - y) + Rabs (a1 * x1 - a * x) <
  / 4%nat * Fulp b radix precision (FPred b radix precision u))%R ->
 MinOrMax radix b (a1 * x1 + y1) u.
intros H Hblop H1 H2 H3.
cut
 (Rabs (a1 * x1 + y1 - u) <
  / 2%nat * Fulp b radix precision (FPred b radix precision u) +
  Rabs (t + y - u))%R; [ intros H4 | idtac ].
case (Rle_or_lt (t + y) u); intros H5.
apply MinOrMax1 with precision; auto with zarith arith; try apply TwoMoreThanOne.
apply Rlt_le_trans with (1 := H4).
apply
 Rplus_le_reg_l
  with (- (/ 2%nat * Fulp b radix precision (FPred b radix precision u)))%R.
ring_simplify
    (- (/ 2%nat * Fulp b radix precision (FPred b radix precision u)) +
     (/ 2%nat * Fulp b radix precision (FPred b radix precision u) +
      Rabs (FtoRradix t + FtoRradix y - FtoRradix u)))%R.
apply
 Rle_trans
  with (/ 2%nat * Fulp b radix precision (FPred b radix precision u))%R;
 [ idtac | right ].
generalize uDef; unfold Closest in |- *; intros H6; elim H6; intros H7 H8;
 clear H6 H7.
case
 (Rle_or_lt (Rabs (FtoRradix t + FtoRradix y - FtoRradix u))
    (/ 2%nat * Fulp b radix precision (FPred b radix precision u)));
 auto; intros H6.
cut
 (Rabs (FtoR radix u - (FtoRradix t + FtoRradix y)) <=
  Rabs (FtoR radix (FPred b radix precision u) - (FtoRradix t + FtoRradix y)))%R;
 [ intros H7 | apply H8 ].
2: apply FBoundedPred; auto with zarith arith.
contradict H7; apply Rlt_not_le.
apply Rlt_le_trans with (Rabs (FtoRradix t + FtoRradix y - FtoRradix u)).
2: right; rewrite <- Rabs_Ropp.
2: replace (- (FtoRradix t + FtoRradix y - FtoRradix u))%R with
    (FtoR radix u - (FtoRradix t + FtoRradix y))%R;
    auto with real; ring.
apply Rle_lt_trans with (2 := H6).
rewrite Rabs_left1.
apply
 Rplus_le_reg_l
  with
    (u - (FtoRradix t + FtoRradix y) -
     / 2%nat * Fulp b radix precision (FPred b radix precision u))%R.
ring_simplify.
fold FtoRradix in |- *; pattern (FtoRradix u) at 1 in |- *;
 replace (FtoRradix u) with
  (FtoRradix (FPred b radix precision u) +
   Fulp b radix precision (FPred b radix precision u))%R;
 [ idtac
 | unfold FtoRradix in |- *; apply FpredUlpPos; auto with zarith arith ].
apply
 Rle_trans
  with
    (Fulp b radix precision (FPred b radix precision u) -
     Fulp b radix precision (FPred b radix precision u) * / 2%nat)%R;
 [ right; ring | idtac ].
apply
 Rle_trans
  with (/ 2%nat * Fulp b radix precision (FPred b radix precision u))%R.
right;
 apply
  trans_eq
   with
     ((1 - / 2%nat) * Fulp b radix precision (FPred b radix precision u))%R;
 [ ring | auto with real arith ].
replace (1 - / 2%nat)%R with (/ 2%nat)%R; auto with real arith.
rewrite <- (Rinv_r (INR 2)); auto with real arith.
simpl; field; auto with real.
apply Rlt_le; apply Rlt_le_trans with (1 := H6).
rewrite Rabs_left1; [ right; ring | idtac ].
apply Rplus_le_reg_l with (FtoRradix u).
ring_simplify; auto.
try apply TwoMoreThanOne.
fold FtoRradix in |- *;
 apply
  Rplus_le_reg_l with (Fulp b radix precision (FPred b radix precision u)).
apply
 Rle_trans
  with
    (FtoRradix (FPred b radix precision u) +
     Fulp b radix precision (FPred b radix precision u) +
     - (FtoRradix t + FtoRradix y))%R;
 [ right; ring
 | unfold FtoRradix in |- *; rewrite FpredUlpPos; auto with zarith arith ].
ring_simplify (Fulp b radix precision (FPred b radix precision u) + 0)%R.
apply Rle_trans with (Rabs (FtoRradix u + - (FtoRradix t + FtoRradix y))).
apply RRle_abs.
rewrite <- Rabs_Ropp.
replace (- (FtoRradix u + - (FtoRradix t + FtoRradix y)))%R with
 (FtoRradix t + FtoRradix y - u)%R; [ idtac | ring ].
apply Rmult_le_reg_l with (INR 2); auto with arith real.
apply Rle_trans with (Fulp b radix precision u).
unfold FtoRradix in |- *; apply ClosestUlp; auto with arith zarith.
replace (INR 2) with (IZR radix); auto with zarith arith real.
try apply TwoMoreThanOne.
apply FulpFPredLe; auto with zarith.
try apply TwoMoreThanOne.
try apply TwoMoreThanOne.
try apply TwoMoreThanOne.
apply
 trans_eq
  with ((1 - / 2%nat) * Fulp b radix precision (FPred b radix precision u))%R;
 [ idtac | ring ].
replace (1 - / 2%nat)%R with (/ 2%nat)%R; auto with real arith.
rewrite <- (Rinv_r (INR 2)); auto with real arith.
simpl; field; auto with real.
case
 (Rle_or_lt (t + y)
    (u + / 2%nat * Fulp b radix precision (FPred b radix precision u)));
 intros H6.
apply MinOrMax1 with precision; auto with arith zarith real.
try apply TwoMoreThanOne.
apply Rlt_le_trans with (1 := H4); rewrite Rabs_right.
apply
 Rplus_le_reg_l
  with
    (u + - (/ 2%nat * Fulp b radix precision (FPred b radix precision u)))%R.
ring_simplify.
apply Rle_trans with (1 := H6).
apply
 Rle_trans
  with
    (FtoRradix u +
     (Fulp b radix precision (FPred b radix precision u) +
      - (Fulp b radix precision (FPred b radix precision u) * / 2%nat)))%R;
 [ apply Rplus_le_compat_l; right | right; ring ].
apply
 trans_eq
  with ((1 - / 2%nat) * Fulp b radix precision (FPred b radix precision u))%R;
 [ idtac | ring ].
replace (1 - / 2%nat)%R with (/ 2%nat)%R; auto with real arith.
rewrite <- (Rinv_r (INR 2)); auto with real arith.
simpl; field; auto with real.
apply Rle_ge; apply Rplus_le_reg_l with (FtoRradix u); auto with real.
ring_simplify; auto with real.
apply MinOrMax2 with precision; auto with arith zarith real.
try apply TwoMoreThanOne.
apply Rlt_le_trans with (1 := H4).
apply
 Rle_trans
  with
    (/ 2%nat * Fulp b radix precision u +
     Rabs (FtoRradix t + FtoRradix y - FtoRradix u))%R;
 [ apply Rplus_le_compat_r | idtac ].
apply Rmult_le_compat_l; auto with real arith.
apply FulpFPredGePos; auto with arith zarith.
try apply TwoMoreThanOne.
apply
 Rle_trans
  with
    (/ 2%nat * Fulp b radix precision u + / 2%nat * Fulp b radix precision u)%R;
 [ apply Rplus_le_compat_l | right ].
apply Rmult_le_reg_l with (INR 2); auto with real arith.
rewrite <- Rmult_assoc; rewrite Rinv_r; auto with real arith.
ring_simplify (1 * Fulp b radix precision u)%R.
unfold FtoRradix in |- *; apply ClosestUlp; auto with real arith zarith.
try apply TwoMoreThanOne.
apply trans_eq with ((/ 2%nat + / 2%nat) * Fulp b radix precision u)%R;
 [ ring | auto with real arith ].
replace (/ 2%nat + / 2%nat)%R with 1%R; [ ring | auto with real arith ].
rewrite <- (Rinv_r (INR 2)); auto with real arith; simpl in |- *; ring.
cut (forall z z' : R, (Rabs z <= z')%R -> (- z' <= z)%R);
 [ intros V | idtac ].
replace (a1 * x1 + y1)%R with
 (a1 * x1 - a * x + (y1 - y) + (a * x - t + (t + y)))%R;
 [ fold FtoRradix in |- * | ring ].
replace (FtoRradix u) with
 (- (/ 4%nat * Fulp b radix precision (FPred b radix precision u)) +
  (- (/ 4%nat * Fulp b radix precision (FPred b radix precision u)) +
   (FtoRradix u +
    / 2%nat * Fulp b radix precision (FPred b radix precision u))))%R.
apply Rplus_le_compat.
apply V; auto with real.
apply
 Rle_trans
  with
    (Rabs (y1 - FtoRradix y) + Rabs (a1 * x1 - FtoRradix a * FtoRradix x))%R;
 auto with real.
rewrite Rplus_comm; apply Rabs_triang.
apply Rplus_le_compat.
apply V; auto with real.
auto with real.
apply
 trans_eq
  with
    (FtoRradix u +
     (- / 4%nat + (- / 4%nat + / 2%nat)) *
     Fulp b radix precision (FPred b radix precision u))%R;
 [ ring | idtac ].
replace (- / 4%nat + (- / 4%nat + / 2%nat))%R with 0%R; [ ring | idtac ].
replace (/ 2%nat)%R with (2%nat * / 4%nat)%R; [ simpl in |- *; ring | idtac ].
replace (INR 4) with (2%nat * 2%nat)%R; auto with arith real zarith.
rewrite Rinv_mult_distr; auto with real arith.
rewrite <- Rmult_assoc; rewrite Rinv_r; auto with real arith.
rewrite <- mult_INR; auto with real arith.
intros z z'; case (Rle_or_lt 0 z); intros P1 P2.
apply Rle_trans with (-0)%R; [ apply Ropp_le_contravar | simpl in |- * ];
 auto with real.
apply Rle_trans with (Rabs z); try apply Rabs_pos; easy.
replace (-0)%R with 0%R; auto with real.
apply Ropp_le_cancel; rewrite Ropp_involutive; rewrite <- Rabs_left;
 auto with real.
replace (a1 * x1 + y1 - FtoRradix u)%R with
 (y1 - y + (a1 * x1 - a * x) + (a * x - t + (t + y - u)))%R;
 [ idtac | ring ].
apply
 Rle_lt_trans
  with
    (Rabs (y1 - FtoRradix y + (a1 * x1 - FtoRradix a * FtoRradix x)) +
     Rabs
       (FtoRradix a * FtoRradix x - FtoRradix t +
        (FtoRradix t + FtoRradix y - FtoRradix u)))%R;
 [ apply Rabs_triang | idtac ].
apply
 Rle_lt_trans
  with
    (Rabs (y1 - FtoRradix y + (a1 * x1 - FtoRradix a * FtoRradix x)) +
     (Rabs (FtoRradix a * FtoRradix x - FtoRradix t) +
      Rabs (FtoRradix t + FtoRradix y - FtoRradix u)))%R;
 [ apply Rplus_le_compat_l; apply Rabs_triang | idtac ].
apply
 Rle_lt_trans
  with
    (Rabs (y1 - FtoRradix y) + Rabs (a1 * x1 - FtoRradix a * FtoRradix x) +
     (Rabs (FtoRradix a * FtoRradix x - FtoRradix t) +
      Rabs (FtoRradix t + FtoRradix y - FtoRradix u)))%R.
apply Rplus_le_compat_r; apply Rabs_triang.
apply
 Rlt_le_trans
  with
    (/ 4%nat * Fulp b radix precision (FPred b radix precision u) +
     (Rabs (FtoRradix a * FtoRradix x - FtoRradix t) +
      Rabs (FtoRradix t + FtoRradix y - FtoRradix u)))%R;
 auto with real.
apply
 Rle_trans
  with
    (/ 4%nat * Fulp b radix precision (FPred b radix precision u) +
     (/ 4%nat * Fulp b radix precision (FPred b radix precision u) +
      Rabs (FtoRradix t + FtoRradix y - FtoRradix u)))%R;
 auto with real.
apply
 Rle_trans
  with
    ((/ 4%nat + / 4%nat) * Fulp b radix precision (FPred b radix precision u) +
     Rabs (FtoRradix t + FtoRradix y - FtoRradix u))%R;
 [ right; ring | apply Rplus_le_compat_r ].
replace (/ 4%nat + / 4%nat)%R with (/ 2%nat)%R; auto with real.
replace (/ 2%nat)%R with (2%nat * / 4%nat)%R; [ simpl in |- *; ring | idtac ].
replace (INR 4) with (2%nat * 2%nat)%R; auto with arith real zarith.
rewrite Rinv_mult_distr; auto with real arith.
rewrite <- Rmult_assoc; rewrite Rinv_r; auto with real arith.
rewrite <- mult_INR; auto with real arith.
Qed.

Theorem Axpy_aux1_aux1 :
 Fnormal radix b t ->
 Fcanonic radix b u ->
 (0 < u)%R ->
 (4%nat * Rabs t <= Rabs u)%R ->
 (Rabs (FtoRradix a * FtoRradix x - FtoRradix t) <=
  / 4%nat * Fulp b radix precision (FPred b radix precision u))%R.
intros H1 H2 H3 H4.
cut (Fcanonic radix b t); [ intros H5 | left; auto ].
apply Rle_trans with (/ 2%nat * Fulp b radix precision t)%R.
apply Rmult_le_reg_l with (INR 2); auto with real arith.
rewrite <- Rmult_assoc; rewrite Rinv_r; auto with real arith.
ring_simplify (1 * Fulp b radix precision t)%R.
unfold FtoRradix in |- *; apply ClosestUlp; auto with real arith zarith.
try apply TwoMoreThanOne.
apply Rle_trans with (/ 2%nat * (/ 4%nat * Fulp b radix precision u))%R.
apply Rmult_le_compat_l; auto with real arith.
apply Rmult_le_reg_l with (INR 4); auto with real arith.
rewrite <- Rmult_assoc; rewrite Rinv_r; auto with real arith.
ring_simplify (1 * Fulp b radix precision u)%R.
cut (Fbounded b (Float (Fnum t) (Z.succ (Z.succ (Fexp t)))));
 [ intros H6 | idtac ].
2: elim H1; intros H7 H8; elim H7; intros H9 H10.
2: repeat (split; simpl in |- *; auto with zarith).
apply
 Rle_trans
  with (Fulp b radix precision (Float (Fnum t) (Z.succ (Z.succ (Fexp t))))).
right; unfold Fulp in |- *.
replace (Fnormalize radix b precision t) with t;
 [ idtac
 | apply
    FcanonicUnique with (radix := radix) (b := b) (precision := precision);
    auto with real zarith ].
2: try apply TwoMoreThanOne.
2: apply FnormalizeCanonic; auto with zarith.
2: try apply TwoMoreThanOne.
2: apply sym_eq; apply FnormalizeCorrect; auto with real zarith.
2: try apply TwoMoreThanOne.
replace
 (Fnormalize radix b precision (Float (Fnum t) (Z.succ (Z.succ (Fexp t)))))
 with (Float (Fnum t) (Z.succ (Z.succ (Fexp t)))).
replace (Fexp (Float (Fnum t) (Z.succ (Z.succ (Fexp t))))) with
 (Z.succ (Z.succ (Fexp t))); [ idtac | simpl in |- *; auto ].
replace (Z.succ (Z.succ (Fexp t))) with (2 + Fexp t)%Z;
 [ rewrite powerRZ_add | unfold Z.succ in |- * ]; auto with zarith real.
f_equal; simpl; unfold radix; ring_simplify; easy.
apply FcanonicUnique with (radix := radix) (b := b) (precision := precision);
 auto with real arith zarith.
try apply TwoMoreThanOne.
2: apply FnormalizeCanonic; auto with zarith.
2: try apply TwoMoreThanOne.
2: apply sym_eq; apply FnormalizeCorrect; auto with real zarith.
2: try apply TwoMoreThanOne.
elim H1; intros H7 H8.
left; repeat (split; simpl in |- *; auto with zarith).
rewrite
 FulpFabs with b radix precision (Float (Fnum t) (Z.succ (Z.succ (Fexp t))));
 auto with zarith.
2: try apply TwoMoreThanOne.
rewrite FulpFabs with b radix precision u; auto with zarith.
2: try apply TwoMoreThanOne.
apply LeFulpPos; auto with zarith real.
try apply TwoMoreThanOne.
now apply absFBounded.
now apply absFBounded.
unfold FtoRradix in |- *; rewrite Fabs_correct; auto with real zarith.
apply Rabs_pos.
try apply TwoMoreThanOne.
replace (FtoR radix (Fabs (Float (Fnum t) (Z.succ (Z.succ (Fexp t)))))) with
 (Z.abs (Fnum t) * powerRZ radix (Z.succ (Z.succ (Fexp t))))%R;
 [ idtac | unfold FtoR in |- *; simpl in |- *; auto ].
replace (Z.succ (Z.succ (Fexp t))) with (2 + Fexp t)%Z;
 [ rewrite powerRZ_add | unfold Z.succ in |- * ]; auto with zarith real.
replace (powerRZ radix 2) with (INR 4);
 [ idtac | simpl; unfold radix; ring_simplify; easy ].
rewrite Rmult_comm; rewrite Rmult_assoc.
replace (powerRZ radix (Fexp t) * Z.abs (Fnum t))%R with (Rabs (FtoRradix t));
 [ unfold FtoRradix in |- *; repeat rewrite Fabs_correct;
    auto with real zarith
 | idtac ].
try apply TwoMoreThanOne.
unfold FtoRradix in |- *; rewrite <- Fabs_correct; auto with zarith;
 try apply TwoMoreThanOne; unfold FtoR in |- *; simpl in |- *; ring.
apply Rmult_le_reg_l with (INR 2); auto with arith real.
rewrite <- Rmult_assoc; rewrite Rinv_r; auto with arith real.
ring_simplify (1 * (/ 4%nat * Fulp b radix precision u))%R.
apply
 Rle_trans
  with
    (/ 4%nat * (2%nat * Fulp b radix precision (FPred b radix precision u)))%R;
 [ apply Rmult_le_compat_l; auto with real arith | right; ring ].
replace (INR 2) with (IZR radix); auto with arith zarith real.
apply FulpFPredLe; auto with arith zarith real; try apply TwoMoreThanOne.
Qed.

Theorem Axpy_aux2 :
 Fcanonic radix b u ->
 Fsubnormal radix b t ->
 (0 < u)%R ->
 FtoRradix u = (t + y)%R ->
 (Rabs (y1 - y) + Rabs (a1 * x1 - a * x) <
  / 4%nat * Fulp b radix precision (FPred b radix precision u))%R ->
 MinOrMax radix b (a1 * x1 + y1) u.
intros H1 H2 H3 H4 H5.
apply MinOrMax1 with precision; auto with zarith; fold FtoRradix in |- *.
try apply TwoMoreThanOne.
replace (a1 * x1 + y1 - FtoRradix u)%R with
 (y1 - y + (a1 * x1 - a * x) + (a * x - t + (t + y - u)))%R;
 [ idtac | ring ].
apply
 Rle_lt_trans
  with
    (Rabs (y1 - FtoRradix y + (a1 * x1 - FtoRradix a * FtoRradix x)) +
     Rabs
       (FtoRradix a * FtoRradix x - FtoRradix t +
        (FtoRradix t + FtoRradix y - FtoRradix u)))%R;
 [ apply Rabs_triang | idtac ].
apply
 Rle_lt_trans
  with
    (Rabs (y1 - FtoRradix y) + Rabs (a1 * x1 - FtoRradix a * FtoRradix x) +
     Rabs
       (FtoRradix a * FtoRradix x - FtoRradix t +
        (FtoRradix t + FtoRradix y - FtoRradix u)))%R;
 [ apply Rplus_le_compat_r; apply Rabs_triang | idtac ].
apply
 Rle_lt_trans
  with
    (Rabs (y1 - FtoRradix y) + Rabs (a1 * x1 - FtoRradix a * FtoRradix x) +
     (Rabs (FtoRradix a * FtoRradix x - FtoRradix t) +
      Rabs (FtoRradix t + FtoRradix y - FtoRradix u)))%R;
 [ apply Rplus_le_compat_l; apply Rabs_triang | idtac ].
apply
 Rlt_le_trans
  with
    (/ 4%nat * Fulp b radix precision (FPred b radix precision u) +
     (Rabs (FtoRradix a * FtoRradix x - FtoRradix t) +
      Rabs (FtoRradix t + FtoRradix y - FtoRradix u)))%R;
 [ apply Rplus_lt_compat_r; auto | idtac ].
apply
 Rle_trans
  with
    (/ 4%nat * Fulp b radix precision (FPred b radix precision u) +
     (/ 2%nat * powerRZ radix (- dExp b) + 0))%R.
apply Rplus_le_compat; auto with real.
apply Rplus_le_compat.
apply Rle_trans with (/ 2%nat * Fulp b radix precision t)%R;
 [ apply Rmult_le_reg_l with (INR 2) | apply Rmult_le_compat_l ];
 auto with real arith.
rewrite <- Rmult_assoc; rewrite Rinv_r; auto with real arith.
ring_simplify (1 * Fulp b radix precision t)%R; unfold FtoRradix in |- *;
 apply ClosestUlp; auto with zarith.
try apply TwoMoreThanOne.
unfold Fulp in |- *; replace (Fnormalize radix b precision t) with t.
elim H2; intros H6 H7; elim H7; intros H8 H9; rewrite H8;
 auto with zarith real.
apply FcanonicUnique with (radix := radix) (b := b) (precision := precision);
 auto with real zarith.
try apply TwoMoreThanOne.
right; auto.
apply FnormalizeCanonic; try apply TwoMoreThanOne; auto with zarith.
apply sym_eq; apply FnormalizeCorrect; auto with zarith.
try apply TwoMoreThanOne.
rewrite H4; right.
ring_simplify (FtoRradix t + FtoRradix y - (FtoRradix t + FtoRradix y))%R;
 apply Rabs_R0.
replace (/ 2%nat * powerRZ radix (- dExp b) + 0)%R with
 (/ 2%nat * powerRZ radix (- dExp b))%R; [ idtac | ring ].
apply
 Rle_trans
  with
    (/ 4%nat * Fulp b radix precision (FPred b radix precision u) +
     / 2%nat * Fulp b radix precision (FPred b radix precision u))%R;
 [ apply Rplus_le_compat_l; apply Rmult_le_compat_l; auto with real arith
 | idtac ].
unfold Fulp in |- *; apply Rle_powerRZ; auto with zarith real.
cut (Fbounded b (Fnormalize radix b precision (FPred b radix precision u)));
 [ intros H6; elim H6; auto | idtac ].
apply FnormalizeBounded; auto with zarith arith.
try apply TwoMoreThanOne.
apply FBoundedPred; auto with zarith arith.
try apply TwoMoreThanOne.
apply
 Rle_trans
  with
    ((/ 4%nat + / 2%nat) * Fulp b radix precision (FPred b radix precision u))%R;
 [ right; ring | idtac ].
apply
 Rle_trans with (1 * Fulp b radix precision (FPred b radix precision u))%R;
 [ apply Rmult_le_compat_r; auto with real zarith | right; ring ].
unfold Fulp in |- *; auto with real zarith.
apply powerRZ_le, Rlt_IZR; try apply TwoMoreThanOne.
apply Rmult_le_reg_l with (INR 4); auto with real arith.
apply Rle_trans with (4%nat * / 4%nat + 2%nat * (2%nat * / 2%nat))%R;
 [ right; simpl; ring | idtac ].
replace (2%nat * 2%nat)%R with (INR (2 * 2)); auto with arith real.
repeat rewrite Rinv_r; auto with real arith.
rewrite 2!Rmult_1_r.
apply Rle_trans with 3%R;[right; easy|idtac].
apply Rle_trans with 4%R; [idtac|right; simpl; ring].
auto with real.
Qed.

Theorem Axpy_aux1_aux3 :
 Fsubnormal radix b t ->
 Fcanonic radix b u ->
 (0 < u)%R ->
 (Z.succ (- dExp b) <= Fexp (FPred b radix precision u))%Z ->
 (Rabs (FtoRradix a * FtoRradix x - FtoRradix t) <=
  / 4%nat * Fulp b radix precision (FPred b radix precision u))%R.
intros H1 H2 H3 H4.
apply Rle_trans with (/ 2%nat * Fulp b radix precision t)%R.
apply Rmult_le_reg_l with (INR 2); auto with real arith;
 rewrite <- Rmult_assoc; rewrite Rinv_r; auto with arith real.
ring_simplify (1 * Fulp b radix precision t)%R; unfold FtoRradix in |- *;
 apply ClosestUlp; auto with zarith.
try apply TwoMoreThanOne.
apply Rmult_le_reg_l with (INR 4); auto with real arith;
 repeat rewrite <- Rmult_assoc.
replace (INR 4) with (2%nat * 2%nat)%R;
 [ rewrite (Rmult_assoc 2%nat 2%nat (/ 2%nat)); repeat rewrite Rinv_r;
    auto with real arith
 | rewrite <- mult_INR; auto with real arith ].
ring_simplify; unfold Fulp in |- *.
replace (Fnormalize radix b precision t) with t;
 [ elim H1; intros H9 H10; elim H10; intros H11 H12; rewrite H11 | idtac ].
2: apply
    FcanonicUnique with (radix := radix) (b := b) (precision := precision);
    auto with real zarith.
2: try apply TwoMoreThanOne.
2: right; auto.
2: apply FnormalizeCanonic; auto with zarith; try apply TwoMoreThanOne.
2: apply sym_eq; apply FnormalizeCorrect; auto with zarith; try apply TwoMoreThanOne.
replace (Fnormalize radix b precision (FPred b radix precision u)) with
 (FPred b radix precision u).
2: apply
    FcanonicUnique with (radix := radix) (b := b) (precision := precision);
    auto with real zarith.
2: try apply TwoMoreThanOne.
2: apply FPredCanonic; auto with zarith; try apply TwoMoreThanOne.
2: apply FnormalizeCanonic; auto with zarith; try apply TwoMoreThanOne.
2: apply FBoundedPred; auto with zarith; try apply TwoMoreThanOne.
2: apply sym_eq; apply FnormalizeCorrect; auto with zarith; try apply TwoMoreThanOne.
replace (2%nat * powerRZ radix (- dExp b))%R with
 (powerRZ radix (Z.succ (- dExp b)));
 [ apply Rle_powerRZ | unfold Z.succ in |- * ]; auto with zarith real.
rewrite powerRZ_add; auto with real zarith; simpl in |- *; unfold radix.
rewrite Rmult_comm, Rmult_1_r; easy.
Qed.

Theorem Axpy_aux3 :
 Fcanonic radix b u ->
 Fsubnormal radix b t ->
 (0 < u)%R ->
 Fexp (FPred b radix precision u) = (- dExp b)%Z ->
 (Z.succ (- dExp b) <= Fexp u)%Z ->
 (Rabs (y1 - y) + Rabs (a1 * x1 - a * x) <
  / 4%nat * Fulp b radix precision (FPred b radix precision u))%R ->
 MinOrMax radix b (a1 * x1 + y1) u.
intros H1 H2 H3 H4 H5 H6.
case (Rle_or_lt u (a1 * x1 + y1)); intros H7.
apply MinOrMax2 with precision; auto with zarith real;
 fold FtoRradix in |- *.
try apply TwoMoreThanOne.
replace (a1 * x1 + y1 - FtoRradix u)%R with
 (y1 - y + (a1 * x1 - a * x) + (a * x - t + (t + y - u)))%R;
 [ idtac | ring ].
apply
 Rlt_le_trans
  with
    (/ 4%nat * powerRZ radix (- dExp b) +
     (/ 2%nat * powerRZ radix (- dExp b) + / 2%nat * Fulp b radix precision u))%R.
apply
 Rle_lt_trans
  with
    (Rabs (y1 - FtoRradix y + (a1 * x1 - FtoRradix a * FtoRradix x)) +
     Rabs
       (FtoRradix a * FtoRradix x - FtoRradix t +
        (FtoRradix t + FtoRradix y - FtoRradix u)))%R;
 [ apply Rabs_triang | idtac ].
apply
 Rlt_le_trans
  with
    (/ 4%nat * powerRZ radix (- dExp b) +
     Rabs
       (FtoRradix a * FtoRradix x - FtoRradix t +
        (FtoRradix t + FtoRradix y - FtoRradix u)))%R;
 [ apply Rplus_lt_compat_r | apply Rplus_le_compat_l ].
apply
 Rle_lt_trans
  with
    (Rabs (y1 - FtoRradix y) + Rabs (a1 * x1 - FtoRradix a * FtoRradix x))%R;
 [ apply Rabs_triang | idtac ].
apply Rlt_le_trans with (1 := H6); unfold Fulp in |- *.
replace (Fexp (Fnormalize radix b precision (FPred b radix precision u)))
 with (- dExp b)%Z; auto with real zarith.
replace (Fnormalize radix b precision (FPred b radix precision u)) with
 (FPred b radix precision u); auto with zarith.
apply FcanonicUnique with (radix := radix) (b := b) (precision := precision);
 auto with real zarith.
try apply TwoMoreThanOne.
apply FPredCanonic; auto with zarith; try apply TwoMoreThanOne.
apply FnormalizeCanonic; auto with zarith; try apply TwoMoreThanOne.
apply FBoundedPred; auto with zarith; try apply TwoMoreThanOne.
apply sym_eq; apply FnormalizeCorrect; auto with zarith.
try apply TwoMoreThanOne.
apply
 Rle_trans
  with
    (Rabs (FtoRradix a * FtoRradix x - FtoRradix t) +
     Rabs (FtoRradix t + FtoRradix y - FtoRradix u))%R;
 [ apply Rabs_triang | apply Rplus_le_compat ].
apply Rmult_le_reg_l with (INR 2);
 [ idtac | rewrite <- Rmult_assoc; rewrite Rinv_r ];
 auto with arith real.
ring_simplify; apply Rle_trans with (Fulp b radix precision t);
 [ unfold FtoRradix in |- *; apply ClosestUlp; auto with zarith
 | unfold Fulp in |- * ].
try apply TwoMoreThanOne.
replace (Fexp (Fnormalize radix b precision t)) with (- dExp b)%Z;
 auto with real zarith.
replace (Fnormalize radix b precision t) with t.
elim H2; auto with zarith.
apply FcanonicUnique with (radix := radix) (b := b) (precision := precision);
 auto with real zarith.
try apply TwoMoreThanOne.
right; auto.
apply FnormalizeCanonic; auto with zarith; try apply TwoMoreThanOne.
apply sym_eq; apply FnormalizeCorrect; auto with zarith.
try apply TwoMoreThanOne.
apply Rmult_le_reg_l with (INR 2);
 [ idtac | rewrite <- Rmult_assoc; rewrite Rinv_r ];
 auto with arith real.
ring_simplify (1 * Fulp b radix precision u)%R; unfold FtoRradix in |- *;
 apply ClosestUlp; auto with zarith.
try apply TwoMoreThanOne.
apply Rmult_le_reg_l with (INR 4); auto with arith real.
apply Rle_trans with (powerRZ radix (-dExp b) * 3%nat
  + Fulp b radix precision u * 2%nat)%R.
right; simpl; field; auto with real.
apply
 Rle_trans
  with
    (Fulp b radix precision u * 2%nat + Fulp b radix precision u * 2%nat)%R;
 [ apply Rplus_le_compat_r | idtac ]; auto with real arith.
apply Rle_trans with (3%nat * powerRZ radix (- dExp b))%R; [ right;ring | idtac ].
apply Rle_trans with (4%nat * powerRZ radix (- dExp b))%R;
 auto with arith zarith real.
apply Rmult_le_compat_r.
apply powerRZ_le, Rlt_IZR; try apply TwoMoreThanOne.
rewrite 2!INR_IZR_INZ; apply Rle_IZR; simpl; omega.
unfold Fulp in |- *; replace (INR 2) with (powerRZ radix 1);
 [ idtac | simpl in |- *; auto with zarith real ].
replace (INR 4) with (powerRZ radix 2);
 [ idtac | simpl; unfold pow, radix; ring ].
repeat rewrite <- powerRZ_add; auto with zarith real.
replace (2 + - dExp b)%Z with (Z.succ (- dExp b) + 1)%Z;
 [ apply Rle_powerRZ | unfold Z.succ in |- * ]; auto with zarith real.
replace (Fnormalize radix b precision u) with u; auto with zarith arith.
apply FcanonicUnique with (radix := radix) (b := b) (precision := precision);
 auto with real zarith.
try apply TwoMoreThanOne.
apply FnormalizeCanonic; auto with zarith; try apply TwoMoreThanOne.
apply sym_eq; apply FnormalizeCorrect; auto with zarith.
try apply TwoMoreThanOne.
right; replace (INR 4) with (2%nat + 2%nat)%R;
 [ ring | rewrite <- plus_INR; auto with arith real ].
case (Rle_or_lt (t + y) u); intros H8.
apply Axpy_aux2; auto.
apply sym_eq; unfold FtoRradix in |- *; apply ExactSum_Near with b precision;
 auto with zarith.
elim H2; intuition.
cut (forall v w : R, (v <= w)%R -> v <> w -> (v < w)%R); [ intros V | idtac ].
2: intros V W H' H'1; case H'; auto with real.
2: intros H'2; absurd (V = W); auto with real.
apply V.
apply Rmult_le_reg_l with (INR 2); auto with real arith.
apply Rle_trans with (Fulp b radix precision u);
 [ unfold FtoRradix in |- *; apply ClosestUlp; auto with zarith | idtac ].
try apply TwoMoreThanOne.
replace (powerRZ 2%Z (- dExp b)) with
 (Fulp b radix precision (FPred b radix precision u)).
replace (INR 2) with (IZR radix); auto with zarith real.
apply FulpFPredLe; auto with zarith.
try apply TwoMoreThanOne.
unfold Fulp in |- *;
 replace (Fnormalize radix b precision (FPred b radix precision u)) with
  (FPred b radix precision u).
rewrite H4; auto with zarith real.
apply FcanonicUnique with (radix := radix) (b := b) (precision := precision);
 auto with real zarith.
try apply TwoMoreThanOne.
apply FPredCanonic; auto with zarith; try apply TwoMoreThanOne.
apply FnormalizeCanonic; auto with zarith; try apply TwoMoreThanOne.
apply FBoundedPred; auto with zarith; try apply TwoMoreThanOne.
apply sym_eq; apply FnormalizeCorrect; auto with zarith.
try apply TwoMoreThanOne.
replace (powerRZ 2%Z (- dExp b)) with
 (Fulp b radix precision (FPred b radix precision u)).
2: unfold Fulp in |- *;
    replace (Fnormalize radix b precision (FPred b radix precision u)) with
     (FPred b radix precision u).
2: rewrite H4; auto with zarith real.
rewrite Rabs_left1; auto with real.
2: apply Rplus_le_reg_l with (FtoRradix u); unfold FtoRradix, radix in |- *.
2: ring_simplify; auto with real zarith.
case
 (Req_dec (- (FtoR 2 t + FtoR 2 y - FtoR 2 u))
    (Fulp b radix precision (FPred b radix precision u)));
 auto.
intros W; contradict uDef.
replace (FtoRradix t + FtoRradix y)%R with
 (FtoRradix (FPred b radix precision u)); auto with real.
cut (FtoRradix u <> FPred b radix precision u); [ intros V' | idtac ].
contradict V'.
cut (ProjectorP b radix (Closest b radix));
 [ unfold ProjectorP in |- *; intros W' | idtac ].
apply sym_eq; apply W'; auto with zarith real.
apply FBoundedPred; auto with zarith; try apply TwoMoreThanOne.
apply RoundedProjector; apply ClosestRoundedModeP with precision;
 auto with zarith.
try apply TwoMoreThanOne.
apply not_eq_sym; apply Rlt_dichotomy_converse; left.
unfold FtoRradix in |- *; apply FPredLt; auto with zarith.
try apply TwoMoreThanOne.
apply
 Rplus_eq_reg_l with (Fulp b radix precision (FPred b radix precision u)).
rewrite Rplus_comm; unfold FtoRradix in |- *; rewrite FpredUlpPos;
 auto with zarith.
rewrite <- W; unfold FtoRradix, radix in |- *; ring.
try apply TwoMoreThanOne.
apply FcanonicUnique with (radix := radix) (b := b) (precision := precision);
 auto with real zarith.
try apply TwoMoreThanOne.
apply FPredCanonic; auto with zarith; try apply TwoMoreThanOne.
apply FnormalizeCanonic; auto with zarith; try apply TwoMoreThanOne.
apply FBoundedPred; auto with zarith; try apply TwoMoreThanOne.
apply sym_eq; apply FnormalizeCorrect; auto with zarith.
try apply TwoMoreThanOne.
apply MinOrMax1 with precision; auto with zarith; fold FtoRradix in |- *.
try apply TwoMoreThanOne.
rewrite Rabs_left1.
2: apply Rplus_le_reg_l with (FtoRradix u).
2: ring_simplify; auto with real.
apply Rle_lt_trans with (y - y1 + (t - a1 * x1) - (t + y - u))%R;
 [ right; ring | idtac ].
apply Rlt_le_trans with (Rabs (y - y1 + (t - a1 * x1))).
cut (forall u v : R, (0 < v)%R -> (u - v < Rabs u)%R);
 [ intros V; apply V | idtac ].
apply Rplus_lt_reg_r with (FtoRradix u).
ring_simplify; auto with real.
intros v w H'1.
apply Rlt_le_trans with v; auto with real.
unfold Rminus in |- *; apply Rlt_le_trans with (v + -0)%R;
 [ auto with real | right; ring ].
apply RRle_abs.
replace (FtoRradix y - y1 + (FtoRradix t - a1 * x1))%R with
 (- (y1 - FtoRradix y + (a1 * x1 - a * x) + (a * x - t)))%R;
 [ idtac | ring ].
rewrite Rabs_Ropp.
apply
 Rle_trans
  with
    (Rabs (y1 - FtoRradix y + (a1 * x1 - FtoRradix a * FtoRradix x)) +
     Rabs (FtoRradix a * FtoRradix x - FtoRradix t))%R;
 [ apply Rabs_triang | idtac ].
apply Rmult_le_reg_l with (INR 2); auto with real arith.
apply
 Rle_trans
  with
    (2%nat * Rabs (y1 - FtoRradix y + (a1 * x1 - FtoRradix a * FtoRradix x)) +
     2%nat * Rabs (FtoRradix a * FtoRradix x - FtoRradix t))%R;
 [ right; ring | idtac ].
apply
 Rle_trans
  with
    (Fulp b radix precision (FPred b radix precision u) +
     Fulp b radix precision (FPred b radix precision u))%R;
 [ idtac | right; simpl; ring ].
apply Rplus_le_compat.
apply
 Rle_trans
  with
    (2%nat * (/ 4%nat * Fulp b radix precision (FPred b radix precision u)))%R;
 auto with real.
apply
 Rle_trans
  with
    (2%nat *
     (Rabs (y1 - FtoRradix y) + Rabs (a1 * x1 - FtoRradix a * FtoRradix x)))%R;
 [ apply Rmult_le_compat_l; auto with real arith; apply Rabs_triang | idtac ].
apply Rmult_le_compat_l; auto with real arith.
apply
 Rle_trans with (1 * Fulp b radix precision (FPred b radix precision u))%R;
 [ rewrite <- Rmult_assoc; apply Rmult_le_compat_r | right; ring ];
 auto with real arith zarith.
unfold Fulp in |- *; auto with real zarith.
apply powerRZ_le, Rlt_IZR; try apply TwoMoreThanOne.
rewrite Rmult_comm; apply Rmult_le_reg_l with (INR 4); auto with arith real;
 rewrite <- Rmult_assoc; rewrite Rinv_r; auto with real arith.
simpl; ring_simplify; auto with real arith.
apply Rle_trans with (Fulp b radix precision t);
 [ unfold FtoRradix in |- *; apply ClosestUlp; auto with zarith
 | unfold Fulp in |- * ].
try apply TwoMoreThanOne.
apply Rle_powerRZ; auto with zarith real.
replace (Fnormalize radix b precision t) with t;
 [ elim H2; intros H9 H10; elim H10; intros H11 H12; rewrite H11 | idtac ].
2: apply sym_eq; apply FcanonicFnormalizeEq; auto with zarith; try apply TwoMoreThanOne; right; auto.
replace (Fnormalize radix b precision (FPred b radix precision u)) with
 (FPred b radix precision u); [ rewrite H4; auto with zarith | idtac ].
apply sym_eq; apply FcanonicFnormalizeEq; auto with zarith.
try apply TwoMoreThanOne.
apply FPredCanonic; auto with zarith; try apply TwoMoreThanOne.
Qed.

Theorem AxpyPos :
 Fcanonic radix b u ->
 Fcanonic radix b t ->
 (0 < u)%R ->
 (4%nat * Rabs t <= Rabs u)%R ->
 (Rabs (y1 - y) + Rabs (a1 * x1 - a * x) <
  / 4%nat * Fulp b radix precision (FPred b radix precision u))%R ->
 MinOrMax radix b (a1 * x1 + y1) u.
intros H1 H2 H3 H4 H5.
case H2; intros H6.
apply Axpy_aux1; auto; apply Axpy_aux1_aux1; auto.
cut (forall z1 z2 : Z, (z1 <= z2)%Z -> z1 = z2 \/ (Z.succ z1 <= z2)%Z);
 [ intros V | idtac ].
2: intros z1 z2 V; omega.
case (V (- dExp b)%Z (Fexp u));
 [ elim Fu; intuition | intros H7 | intros H7 ].
apply Axpy_aux2; auto.
unfold FtoRradix in |- *; apply plusExact1 with b precision;
 auto with zarith.
try apply TwoMoreThanOne.
rewrite <- H7; elim H6; intros H8 H9; elim H9; intros H10 H11; rewrite H10.
rewrite Zmin_le1; auto with zarith; elim Fy; auto.
case (V (- dExp b)%Z (Fexp (FPred b radix precision u))).
cut (Fbounded b (FPred b radix precision u)); [ intros H8 | idtac ];
 auto.
apply H8.
apply FBoundedPred; auto with zarith.
try apply TwoMoreThanOne.
intros H8; apply Axpy_aux3; auto.
intros H8; apply Axpy_aux1; auto.
apply Axpy_aux1_aux3; auto.
Qed.

Definition FLess (p : float) :=
  match Rcase_abs p with
  | left _ => FSucc b radix precision p
  | right _ => FPred b radix precision p
  end.

Theorem UlpFlessuGe_aux :
 forall p : float,
 Fbounded b p ->
 Fcanonic radix b p ->
 (Rabs p - Fulp b radix precision p <= Rabs (FLess p))%R.
intros p H H1.
cut
 (forall p : float,
  Fbounded b p ->
  Fcanonic radix b p ->
  (0 < p)%R ->
  (Rabs p - Fulp b radix precision p <= Rabs (FPred b radix precision p))%R);
 [ intros V | idtac ].
unfold FLess in |- *; case (Rcase_abs p); intros H2.
rewrite <- Rabs_Ropp; unfold FtoRradix in |- *; rewrite <- Fopp_correct.
pattern p at 3 in |- *; rewrite <- Fopp_Fopp.
rewrite <- (Rabs_Ropp (FtoR radix (FSucc b radix precision (Fopp (Fopp p)))));
 rewrite <- (Fopp_correct radix (FSucc b radix precision (Fopp (Fopp p)))).
rewrite <- FPredFopFSucc with b radix precision (Fopp p); auto with zarith.
2: try apply TwoMoreThanOne.
replace (Fulp b radix precision p) with (Fulp b radix precision (Fopp p));
 auto.
fold FtoRradix in |- *; apply V; auto.
apply oppBounded; easy.
apply FcanonicFopp; easy.
unfold FtoRradix in |- *; rewrite Fopp_correct; auto with real.
unfold Fulp in |- *; rewrite Fnormalize_Fopp; auto with zarith.
cut (0 <= p)%R; [ intros H3; case H3; intros H4 | auto with real ].
apply V; auto.
rewrite <- H4; rewrite Rabs_R0.
ring_simplify (0 - Fulp b radix precision p)%R; apply Rle_trans with 0%R;
 auto with real zarith.
rewrite <- Ropp_0; apply Ropp_le_contravar.
unfold Fulp; apply powerRZ_le, Rlt_IZR; try apply TwoMoreThanOne.
apply Rabs_pos.
intros q H2 H3 H4.
unfold FtoRradix in |- *; rewrite <- FpredUlpPos with b radix precision q;
 auto with zarith; fold FtoRradix in |- *; try apply TwoMoreThanOne.
apply
 Rle_trans
  with
    (Rabs (FtoRradix (FPred b radix precision q)) +
     Rabs (Fulp b radix precision (FPred b radix precision q)) -
     Fulp b radix precision q)%R;
 [ unfold Rminus in |- *; apply Rplus_le_compat_r; apply Rabs_triang
 | idtac ].
rewrite (Rabs_right (Fulp b radix precision (FPred b radix precision q)));
 [ idtac | apply Rle_ge; unfold Fulp in |- *; auto with real zarith ].
apply Rplus_le_reg_l with (Fulp b radix precision q).
ring_simplify
    (Fulp b radix precision q +
     (Rabs (FtoRradix (FPred b radix precision q)) +
      Fulp b radix precision (FPred b radix precision q) -
      Fulp b radix precision q))%R.
rewrite Rplus_comm; apply Rplus_le_compat_r.
apply FulpFPredGePos; auto with zarith real.
try apply TwoMoreThanOne.
apply powerRZ_le, Rlt_IZR; try apply TwoMoreThanOne.
Qed.

Theorem UlpFlessuGe :
 Fcanonic radix b u ->
 (/
  (4%nat * (powerRZ radix precision - 1) * (1 + powerRZ radix (- precision))) *
  ((1 - powerRZ radix (Z.succ (- precision))) * Rabs y) +
  -
  (/
   (4%nat * (powerRZ radix precision - 1) * (1 + powerRZ radix (- precision)) *
    (1 - powerRZ radix (- precision))) *
   ((1 - powerRZ radix (Z.succ (- precision))) * Rabs (a * x))) +
  -
  (powerRZ radix (Z.pred (- dExp b)) *
   (/ (2%nat * (powerRZ radix precision - 1)) +
    /
    (4%nat * (powerRZ radix precision - 1) *
     (1 + powerRZ radix (- precision)) * (1 - powerRZ radix (- precision))) *
    (1 - powerRZ radix (Z.succ (- precision))))) <=
  / 4%nat * Fulp b radix precision (FLess u))%R.
intros Cu.
cut (0 < 1 - powerRZ radix (- precision))%R; [ intros H'1 | idtac ].
2: apply Rplus_lt_reg_r with (powerRZ radix (- precision)).
2: ring_simplify.
2: replace 1%R with (powerRZ radix 0);
    [ auto with real zarith | simpl in |- *; auto ].
2: apply Rlt_powerRZ; auto with zarith.
2: apply Rlt_IZR; try apply TwoMoreThanOne.
cut (0 < 1 + powerRZ radix (- precision))%R; [ intros H'2 | idtac ].
2: apply Rlt_le_trans with 1%R; auto with real.
2: apply Rle_trans with (1 + 0)%R; auto with real zarith.
2: apply Rplus_le_compat_l.
2: apply powerRZ_le, Rlt_IZR; try apply TwoMoreThanOne.
cut (0 < powerRZ radix precision - 1)%R; [ intros H'3 | idtac ].
2: apply Rplus_lt_reg_r with 1%R.
2: ring_simplify.
2: replace 1%R with (powerRZ radix 0);
    [ auto with real zarith | simpl in |- *; auto ].
2: apply Rlt_powerRZ; auto with zarith; apply Rlt_IZR; try apply TwoMoreThanOne.
cut (0 < 1 - powerRZ radix (Z.succ (- precision)))%R; [ intros H'4 | idtac ].
2: apply Rplus_lt_reg_r with (powerRZ radix (Z.succ (- precision))).
2: ring_simplify.
2: replace 1%R with (powerRZ radix 0);
    [ auto with real zarith | simpl in |- *; auto ].
2: apply Rlt_powerRZ; auto with zarith; apply Rlt_IZR; try apply TwoMoreThanOne.
cut (0 < INR 4)%R; [ intros H'5 | auto with arith real ].
rewrite
 (Rinv_mult_distr
    (4%nat * (powerRZ radix precision - 1) *
     (1 + powerRZ radix (- precision))) (1 - powerRZ radix (- precision)))
 ; auto with real.
2: apply Rmult_integral_contrapositive; split; auto with real.
apply
 Rle_trans
  with
    (/
     (4%nat * (powerRZ radix precision - 1) *
      (1 + powerRZ radix (- precision))) *
     (1 - powerRZ radix (Z.succ (- precision))) *
     (Rabs y -
      (Rabs (FtoRradix a * FtoRradix x) * / (1 - powerRZ radix (- precision)) +
       powerRZ radix (Z.pred (- dExp b)) * / (1 - powerRZ radix (- precision)))) +
     -
     (powerRZ radix (Z.pred (- dExp b)) *
      / (2%nat * (powerRZ radix precision - 1))))%R;
 [ right; ring; ring | idtac ].
apply
 Rle_trans
  with
    (/
     (4%nat * (powerRZ radix precision - 1) *
      (1 + powerRZ radix (- precision))) *
     (1 - powerRZ radix (Z.succ (- precision))) *
     (Rabs (FtoRradix y) - Rabs t) +
     -
     (powerRZ radix (Z.pred (- dExp b)) *
      / (2%nat * (powerRZ radix precision - 1))))%R;
 [ apply Rplus_le_compat_r | idtac ].
apply Rmult_le_compat_l.
apply Rlt_le; apply Rmult_lt_0_compat; auto with real.
apply Rinv_0_lt_compat; apply Rmult_lt_0_compat; auto with real;
 apply Rmult_lt_0_compat; auto with real.
unfold Rminus in |- *; apply Rplus_le_compat_l; apply Ropp_le_contravar.
apply
 Rle_trans
  with
    (Rabs (FtoRradix a * FtoRradix x) * / (1 - powerRZ 2%Z (- precision)) +
     powerRZ 2%Z (Z.pred (- dExp b)) * / (1 - powerRZ 2%Z (- precision)))%R;
 [ idtac | right; unfold FtoRradix, radix, Rminus in |- *; auto with real ].
apply RoundLeGeneral; auto with zarith.
rewrite
 (Rinv_mult_distr (4%nat * (powerRZ radix precision - 1))
    (1 + powerRZ radix (- precision))); auto with real.
apply
 Rle_trans
  with
    (/ (4%nat * (powerRZ radix precision - 1)) *
     (1 - powerRZ radix (Z.succ (- precision))) *
     (/ (1 + powerRZ radix (- precision)) *
      (Rabs (FtoRradix y) - Rabs (FtoRradix t))) +
     -
     (powerRZ radix (Z.pred (- dExp b)) *
      / (2%nat * (powerRZ radix precision - 1))))%R;
 [ right; ring | idtac ].
apply
 Rle_trans
  with
    (/ (4%nat * (powerRZ radix precision - 1)) *
     (1 - powerRZ radix (Z.succ (- precision))) * Rabs u +
     -
     (powerRZ radix (Z.pred (- dExp b)) *
      / (2%nat * (powerRZ radix precision - 1))))%R.
apply Rplus_le_compat_r; apply Rmult_le_compat_l.
apply Rlt_le; apply Rmult_lt_0_compat; auto with real.
apply Rinv_0_lt_compat; apply Rmult_lt_0_compat; auto with real.
apply
 Rle_trans
  with
    (/ (1 + powerRZ radix (- precision)) * Rabs (FtoRradix y + FtoRradix t))%R.
apply Rmult_le_compat_l; auto with real.
rewrite <- (Rabs_Ropp (FtoRradix t)).
replace (FtoRradix y + FtoRradix t)%R with (FtoRradix y - - FtoRradix t)%R;
 [ apply Rabs_triang_inv | ring ].
case Cu; intros H1.
apply Rmult_le_reg_l with (1 + powerRZ radix (- precision))%R; auto.
apply
 Rle_trans
  with
    (Rabs (FtoRradix y + FtoRradix t) *
     ((1 + powerRZ radix (- precision)) * / (1 + powerRZ radix (- precision))))%R;
 [ right; ring; ring | idtac ].
rewrite Rinv_r; auto with real.
ring_simplify.
apply Rle_trans with (Rabs u + / radix * Fulp b radix precision u)%R.
apply Rplus_le_reg_l with (- Rabs u)%R.
ring_simplify.
apply Rle_trans with (Rabs (FtoRradix y + FtoRradix t - FtoRradix u)).
apply
 Rle_trans with (Rabs (FtoRradix y + FtoRradix t) - Rabs (FtoRradix u))%R;
 [ right; ring | apply Rabs_triang_inv ].
apply Rmult_le_reg_l with (IZR radix); auto with real zarith.
apply Rle_trans with (Fulp b radix precision u * (radix * / radix))%R;
 [ rewrite Rinv_r; auto with real zarith | right; ring ].
ring_simplify (Fulp b radix precision u * 1)%R.
unfold FtoRradix in |- *; apply ClosestUlp; auto with real zarith.
try apply TwoMoreThanOne.
rewrite Rplus_comm; auto.
rewrite Rplus_comm; apply Rplus_le_compat_r.
apply Rmult_le_reg_l with (IZR radix); auto with real zarith.
rewrite <- Rmult_assoc; rewrite Rinv_r; auto with real zarith.
ring_simplify (1 * Fulp b radix precision u)%R.
apply
 Rle_trans with (Rabs (FtoRradix u) * powerRZ radix (Z.succ (- precision)))%R.
unfold FtoRradix in |- *; apply FulpLe2; auto with zarith.
try apply TwoMoreThanOne.
replace (Fnormalize radix b precision u) with u;
 [ idtac
 | apply
    FcanonicUnique with (radix := radix) (precision := precision) (b := b) ];
 auto with zarith real.
try apply TwoMoreThanOne.
apply FnormalizeCanonic; auto with zarith.
try apply TwoMoreThanOne.
apply sym_eq; apply FnormalizeCorrect; auto with zarith.
try apply TwoMoreThanOne.
unfold Z.succ in |- *; rewrite powerRZ_add; auto with zarith real;
 simpl in |- *; auto with zarith real; right; ring.
replace (FtoRradix y + FtoRradix t)%R with (FtoRradix u); auto with real.
apply Rle_trans with (1 * Rabs (FtoRradix u))%R;
 [ apply Rmult_le_compat_r | right; ring ]; auto with real.
apply Rabs_pos.
pattern 1%R at 2 in |- *; replace 1%R with (/ 1)%R; auto with real zarith.
apply Rle_Rinv; auto with real zarith.
pattern 1%R at 1 in |- *; replace 1%R with (1 + 0)%R; auto with real zarith.
apply Rplus_le_compat_l.
apply powerRZ_le, Rlt_IZR; try apply TwoMoreThanOne.
unfold FtoRradix in |- *; apply plusExact1 with b precision;
 auto with real zarith.
try apply TwoMoreThanOne.
rewrite Rplus_comm; auto.
elim H1; intros H5 H6; elim H6; intros H7 H8; rewrite H7; apply Zmin_Zle;
 elim Fy; elim Ft; auto with zarith.
apply
 Rle_trans
  with
    (/ (4%nat * (powerRZ radix precision - 1)) *
     (Rabs (FtoRradix u) -
      (Rabs u * powerRZ radix (Z.succ (- precision)) +
       powerRZ radix (- dExp b))))%R.
right; simpl; ring_simplify.
unfold Z.pred, Z.succ; rewrite 2!powerRZ_add; auto with zarith real.
unfold radix; simpl; field; auto with real.
apply
 Rle_trans
  with
    (/ (4%nat * (powerRZ radix precision - 1)) *
     (Rabs (FtoRradix u) - Fulp b radix precision u))%R.
apply Rmult_le_compat_l; auto with real.
apply Rlt_le; apply Rinv_0_lt_compat; apply Rmult_lt_0_compat; auto with real.
unfold Rminus in |- *; apply Rplus_le_compat_l; apply Ropp_le_contravar.
apply FulpLeGeneral; auto.
apply
 Rle_trans
  with (/ (4%nat * (powerRZ radix precision - 1)) * Rabs (FLess u))%R.
apply Rmult_le_compat_l; auto with real.
apply Rlt_le; apply Rinv_0_lt_compat; apply Rmult_lt_0_compat; auto with real.
apply UlpFlessuGe_aux; auto.
rewrite Rinv_mult_distr; auto with real.
rewrite Rmult_assoc; apply Rmult_le_compat_l; auto with real.
apply Rmult_le_reg_l with (powerRZ radix precision - 1)%R; auto.
rewrite <- Rmult_assoc; rewrite Rinv_r; auto with real.
ring_simplify (1 * Rabs (FtoRradix (FLess u)))%R; unfold FtoRradix in |- *;
 apply FulpGe; auto with zarith.
try apply TwoMoreThanOne.
unfold FLess in |- *; case (Rcase_abs u); intros H.
apply FBoundedSuc; auto with zarith.
try apply TwoMoreThanOne.
apply FBoundedPred; auto with zarith; try apply TwoMoreThanOne.
Qed.

Theorem UlpFlessuGe2 :
 Fcanonic radix b u ->
 (powerRZ radix (Z.pred (Z.pred (- precision))) *
  (1 - powerRZ radix (Z.succ (- precision))) * Rabs y +
  - (powerRZ radix (Z.pred (Z.pred (- precision))) * Rabs (a * x)) +
  - powerRZ radix (Z.pred (Z.pred (- dExp b))) <
  / 4%nat * Fulp b radix precision (FLess u))%R.
intros H.
apply
 Rlt_le_trans
  with
    (/
     (4%nat * (powerRZ radix precision - 1) *
      (1 + powerRZ radix (- precision))) *
     ((1 - powerRZ radix (Z.succ (- precision))) * Rabs (FtoRradix y)) +
     -
     (/
      (4%nat * (powerRZ radix precision - 1) *
       (1 + powerRZ radix (- precision)) * (1 - powerRZ radix (- precision))) *
      ((1 - powerRZ radix (Z.succ (- precision))) *
       Rabs (FtoRradix a * FtoRradix x))) +
     -
     (powerRZ radix (Z.pred (- dExp b)) *
      (/ (2%nat * (powerRZ radix precision - 1)) +
       /
       (4%nat * (powerRZ radix precision - 1) *
        (1 + powerRZ radix (- precision)) * (1 - powerRZ radix (- precision))) *
       (1 - powerRZ radix (Z.succ (- precision))))))%R;
 [ idtac | apply UlpFlessuGe; auto ].
cut (0 < 1 - powerRZ radix (- precision))%R; [ intros H'1 | idtac ].
2: apply Rplus_lt_reg_r with (powerRZ radix (- precision)).
2: ring_simplify.
2: replace 1%R with (powerRZ radix 0);
    [ auto with real zarith | simpl in |- *; auto ].
2: apply Rlt_powerRZ; auto with zarith; apply Rlt_IZR, TwoMoreThanOne.
cut (0 < 1 + powerRZ radix (- precision))%R; [ intros H'2 | idtac ].
2: apply Rlt_le_trans with 1%R; auto with real.
2: apply Rle_trans with (1 + 0)%R; auto with real zarith.
2: apply Rplus_le_compat_l.
2: apply powerRZ_le, Rlt_IZR; try apply TwoMoreThanOne.
cut (0 < powerRZ radix precision - 1)%R; [ intros H'3 | idtac ].
2: apply Rplus_lt_reg_r with 1%R.
2: ring_simplify.
2: replace 1%R with (powerRZ radix 0);
    [ auto with real zarith | simpl in |- *; auto ].
2: apply Rlt_powerRZ; auto with zarith; apply Rlt_IZR, TwoMoreThanOne.
cut (0 < 1 - powerRZ radix (Z.succ (- precision)))%R; [ intros H'4 | idtac ].
2: apply Rplus_lt_reg_r with (powerRZ radix (Z.succ (- precision))).
2: ring_simplify.
2: replace 1%R with (powerRZ radix 0);
    [ auto with real zarith | simpl in |- *; auto ].
2: apply Rlt_powerRZ; auto with zarith; apply Rlt_IZR, TwoMoreThanOne.
cut (0 < INR 4)%R; [ intros H'5 | auto with arith real ].
apply
 Rle_lt_trans
  with
    (/
     (4%nat * (powerRZ radix precision - 1) *
      (1 + powerRZ radix (- precision))) *
     ((1 - powerRZ radix (Z.succ (- precision))) * Rabs (FtoRradix y)) +
     -
     (/
      (4%nat * (powerRZ radix precision - 1) *
       (1 + powerRZ radix (- precision)) * (1 - powerRZ radix (- precision))) *
      ((1 - powerRZ radix (Z.succ (- precision))) *
       Rabs (FtoRradix a * FtoRradix x))) +
     - powerRZ radix (Z.pred (Z.pred (- dExp b))))%R;
 [ apply Rplus_le_compat_r | apply Rplus_lt_compat_l ].
apply Rplus_le_compat.
repeat rewrite <- Rmult_assoc.
apply Rmult_le_compat_r; auto with real.
apply Rabs_pos.
rewrite Rmult_assoc.
ring_simplify ((powerRZ radix precision - 1) * (1 + powerRZ radix (- precision)))%R.
rewrite <- powerRZ_add; auto with real zarith.
replace (powerRZ radix (- precision + precision)) with 1%R;
 [ idtac
 | ring_simplify (- precision + precision)%Z; simpl in |- *; auto with real zarith ].
ring_simplify (-1 + (- powerRZ radix (- precision) + (powerRZ radix precision + 1)))%R.
apply Rmult_le_compat_r; auto with real.
apply Rle_trans with (/ (4%nat * powerRZ radix precision))%R;
 [ right | idtac ].
unfold Z.pred in |- *.
replace (- precision + -1 + -1)%Z with (- 2%nat + - precision)%Z;
 [ rewrite powerRZ_add | ring ]; auto with zarith real.
rewrite Rinv_mult_distr; auto with real zarith.
replace (powerRZ radix (- 2%nat)) with (/ 4%nat)%R.
replace (/ powerRZ radix precision)%R with (powerRZ radix (- precision));
 auto with real.
apply powerRZ_Zopp; auto with zarith real.
rewrite powerRZ_Zopp; auto with zarith real.
replace (powerRZ radix 2%nat) with (INR 4);
 [ auto with zarith real | unfold radix; simpl in |- *; ring ].
apply Rgt_not_eq, powerRZ_lt, Rlt_IZR; try apply TwoMoreThanOne.
apply Rle_Rinv; auto with real.
apply
 Rlt_le_trans
  with
    (4%nat *
     (powerRZ radix precision - 1 + (1 - powerRZ radix (- precision))))%R.
apply Rle_lt_trans with (4%nat * (0 + 0))%R; [ right; ring | auto with real ].
ring_simplify (precision+-precision)%Z; simpl; right;ring.
apply Rmult_le_compat_l; auto with real zarith.
ring_simplify (precision+-precision)%Z.
apply
 Rle_trans with (powerRZ radix precision + - powerRZ radix (- precision))%R;
 [ right; simpl; ring | idtac ]; auto with real zarith.
apply Rle_trans with (powerRZ radix precision + -0)%R;
 [ idtac | right; ring ]; auto with real zarith.
apply Rplus_le_compat_l, Ropp_le_contravar.
apply powerRZ_le; auto with zarith; apply Rlt_IZR; try apply TwoMoreThanOne.
apply Ropp_le_contravar.
repeat rewrite <- Rmult_assoc.
apply Rmult_le_compat_r; auto with real.
apply Rabs_pos.
unfold Z.pred in |- *.
replace (- precision + -1 + -1)%Z with (- 2%nat + - precision)%Z;
 [ rewrite powerRZ_add | ring ]; auto with zarith real.
repeat rewrite Rmult_assoc; rewrite Rinv_mult_distr.
2: auto with real.
2: apply Rmult_integral_contrapositive; split; auto with real.
replace (powerRZ radix (- 2%nat)) with (/ 4%nat)%R.
2: rewrite powerRZ_Zopp; auto with zarith real.
2: replace (powerRZ radix 2%nat) with (INR 4);
    [ auto with zarith real | unfold radix; simpl in |- *; ring ].
repeat rewrite Rmult_assoc; apply Rmult_le_compat_l; auto with real.
apply
 Rmult_le_reg_l
  with
    ((powerRZ radix precision - 1) *
     ((1 + powerRZ radix (- precision)) * (1 - powerRZ radix (- precision))))%R;
 auto with real.
apply Rmult_lt_0_compat; auto with real.
apply Rmult_lt_0_compat; auto with real.
rewrite <- Rmult_assoc; rewrite Rinv_r; auto with real.
2: apply Rmult_integral_contrapositive; split; auto with real.
unfold Rminus; repeat rewrite Rmult_plus_distr_r.
rewrite Ropp_mult_distr_l_reverse.
repeat rewrite Rmult_1_l.
repeat rewrite Rmult_plus_distr_l.
repeat rewrite Rmult_1_r.
repeat rewrite <- Ropp_mult_distr_r.
rewrite <- Ropp_mult_distr_l.
repeat rewrite <- powerRZ_add; auto with real zarith.
replace (powerRZ radix (precision + - precision)) with 1%R;
 [ idtac | ring_simplify ( precision +- precision)%Z; simpl in |- *; auto ].
apply
 Rle_trans
  with
    (1 +
     (- powerRZ radix (- precision) + (- powerRZ radix (- precision) + 0)))%R;
 [ right | idtac ].
unfold Z.succ; rewrite powerRZ_add; auto with real zarith; simpl; unfold radix; ring.
apply
 Rle_trans
  with
    (1 +
     (- powerRZ radix (- precision) +
      (- powerRZ radix (-2 * precision) + powerRZ radix (-3 * precision))))%R;
 [ idtac | right].
apply Rplus_le_compat_l; apply Rplus_le_compat_l.
apply Rplus_le_compat; auto with real zarith.
apply Ropp_le_contravar, Rle_powerRZ; auto with zarith; left; apply Rlt_IZR; try apply TwoMoreThanOne.
apply powerRZ_le; apply Rlt_IZR, TwoMoreThanOne.
apply trans_eq with
 (powerRZ radix (- precision) * powerRZ radix precision +( -
 powerRZ radix (- precision) *
 powerRZ radix (precision + (- precision + - precision)) +
 powerRZ radix (- precision) *
 powerRZ radix (- precision + - precision) -
 powerRZ radix (- precision)))%R;[idtac|ring].
repeat rewrite <- powerRZ_add; auto with real zarith.
f_equal.
ring_simplify (-precision+precision)%Z; easy.
rewrite <- Ropp_mult_distr_l.
rewrite <- powerRZ_add; auto with real zarith.
replace (- precision + (precision + (- precision + - precision)))%Z with
  (-2*precision)%Z by ring.
replace (- precision + (- precision + - precision))%Z with
  (-3*precision)%Z by ring.
ring.
apply Ropp_lt_contravar.
apply Rmult_lt_reg_l with (powerRZ radix (Z_of_nat 2 + dExp b));
 auto with real zarith.
apply powerRZ_lt, Rlt_IZR; try apply TwoMoreThanOne.
rewrite <- Rmult_assoc; unfold Z.pred in |- *.
repeat rewrite <- powerRZ_add; auto with real zarith.
replace (2%nat + dExp b + (- dExp b + -1 + -1))%Z with 0%Z by ring.
replace (2%nat + dExp b + (- dExp b + -1))%Z with 1%Z by ring.
replace (powerRZ radix 0) with 1%R;
 [ idtac | simpl in |- *; auto with real zarith ].
replace (powerRZ radix 1) with 2%R;
 [ idtac | unfold radix; simpl in |- *; ring; auto with real zarith ].
rewrite Rmult_plus_distr_l.
apply Rlt_le_trans with (/ 3%nat + 2%nat * / 3%nat)%R.
2: simpl; right;field; auto with real.
apply
 Rle_lt_trans
  with
    (/ 3%nat +
  2 *
 (/
  (4%nat * (powerRZ radix precision - 1) * (1 + powerRZ radix (- precision)) *
   (1 - powerRZ radix (- precision))) * (1 - powerRZ radix (- precision + 1))))%R;
 [ apply Rplus_le_compat_r | apply Rplus_lt_compat_l ].
rewrite Rinv_mult_distr; auto with arith real.
rewrite <- Rmult_assoc; rewrite Rinv_r; auto with arith real.
rewrite Rmult_1_l; replace (INR 3) with (powerRZ radix 2%nat - 1)%R;
 auto with real zarith arith.
apply Rle_Rinv; auto with real arith.
replace 0%R with (powerRZ radix 0 - 1)%R;
 [ unfold Rminus in |- *; auto with real arith zarith | simpl in |- *; ring ].
apply Rplus_lt_compat_r.
apply Rlt_powerRZ; auto with zarith; try apply Rlt_IZR; try apply TwoMoreThanOne.
unfold Rminus in |- *; apply Rplus_le_compat_r; apply Rle_powerRZ; auto with real arith zarith.
unfold radix; simpl in |- *; ring; auto with arith zarith real.
apply Rmult_lt_compat_l; auto with real arith.
repeat rewrite Rinv_mult_distr; auto with real.
2: apply Rmult_integral_contrapositive; split; auto with real.
apply
 Rle_lt_trans
  with
    (/ 4%nat * / (powerRZ radix 2%nat - 1) * / (1 + 0) *
     / (1 - powerRZ radix (- 2%nat)) * (1 - 0))%R.
repeat rewrite Rmult_assoc.
apply Rmult_le_compat_l; auto with real arith.
apply Rmult_le_compat; auto with real arith zarith.
apply Rmult_le_pos; auto with real.
apply Rmult_le_pos; auto with real.
apply Rle_Rinv; auto with real arith zarith.
apply Rle_lt_trans with (powerRZ radix 0 - 1)%R;
 [ right; simpl in |- *; ring | auto with real arith zarith ].
unfold Rminus in |- *; apply Rplus_lt_compat_r; auto with real arith zarith.
apply Rlt_powerRZ; auto with zarith; try (simpl; omega); apply Rlt_IZR; try apply TwoMoreThanOne.
unfold Rminus in |- *; apply Rplus_le_compat_r; apply Rle_powerRZ;
 auto with real arith zarith.
apply Rmult_le_compat; auto with real arith zarith.
apply Rmult_le_pos; auto with real.
apply Rle_Rinv.
ring_simplify; auto with real.
apply Rplus_le_compat_l.
apply powerRZ_le; apply Rlt_IZR, TwoMoreThanOne.
apply Rmult_le_compat; auto with real arith zarith.
apply Rle_Rinv; auto with real arith zarith.
apply Rle_lt_trans with (1 - powerRZ radix 0)%R;
 [ simpl in |- *; right; ring | auto with real arith zarith ].
unfold Rminus in |- *; apply Rplus_lt_compat_l; apply Ropp_lt_contravar;
 auto with real arith zarith.
apply Rlt_powerRZ; simpl; auto with zarith.
apply Rlt_IZR; try apply TwoMoreThanOne.
unfold Rminus in |- *; apply Rplus_le_compat_l; apply Ropp_le_contravar;
 apply Rle_powerRZ; auto with real arith zarith.
unfold Rminus in |- *; apply Rplus_le_compat_l; apply Ropp_le_contravar;
 auto with real arith zarith.
apply powerRZ_le; apply Rlt_IZR, TwoMoreThanOne.
ring_simplify (1 - 0)%R; rewrite Rmult_1_r.
ring_simplify (1 + 0)%R; rewrite Rinv_1; rewrite Rmult_1_r.
rewrite powerRZ_Zopp; auto with real arith zarith.
replace (powerRZ radix 2%nat) with (INR 4);
 [ idtac | simpl in |- *; unfold radix; ring; auto with arith real zarith ].
replace (4%nat - 1)%R with (INR 3);
 [ idtac | simpl in |- *; ring; auto with arith real zarith ].
replace (1 - / 4%nat)%R with (3%nat * / 4%nat)%R.
rewrite Rinv_mult_distr; auto with real arith zarith.
rewrite Rinv_involutive; auto with real arith zarith.
apply Rle_lt_trans with (4%nat * / 4%nat * (/ 3%nat * / 3%nat))%R;
 [ right; ring | idtac ].
rewrite Rinv_r; auto with real arith zarith; rewrite Rmult_1_l.
apply Rlt_le_trans with (/ 3%nat * 1)%R;
 [ apply Rmult_lt_compat_l; auto with arith real zarith | right; ring ].
apply Rmult_lt_reg_l with (INR 3); auto with arith real.
rewrite Rinv_r; auto with arith real; rewrite Rmult_1_r; auto with arith real.
simpl; field; auto with real.
Qed.
End AxpyAux.
Section Axpy.

Let radix := 2%Z.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.
Variable b : Fbound.
Variable precision : nat.
Hypothesis precisionGreaterThanOne : 1 < precision.
Hypothesis pGivesBound : Zpos (vNum b) = Zpower_nat radix precision.

Theorem Axpy_tFlessu :
 forall (a1 x1 y1 : R) (a x y t u : float),
 Fbounded b a ->
 Fbounded b x ->
 Fbounded b y ->
 Fbounded b t ->
 Fbounded b u ->
 Closest b radix (a * x) t ->
 Closest b radix (t + y) u ->
 Fcanonic radix b u ->
 Fcanonic radix b t ->
 (4%nat * Rabs t <= Rabs u)%R ->
 (Rabs (y1 - y) + Rabs (a1 * x1 - a * x) <
  / 4%nat * Fulp b radix precision (FLess b precision u))%R ->
 MinOrMax radix b (a1 * x1 + y1) u.
intros a1 x1 y1 a x y t u Fa Fx Fy Ft Fu tDef uDef Cu Ct H1.
unfold FLess in |- *; case (Rcase_abs (FtoR 2 u)); fold radix in |- *;
 intros H3 H2.
2: cut (0 <= u)%R;
    [ intros H'; case H'; intros H4; clear H' | auto with real ].
2: apply AxpyPos with precision a x y t; auto.
apply MinOrMax_Fopp.
replace (- (a1 * x1 + y1))%R with (- a1 * x1 + - y1)%R; [ idtac | ring ].
apply AxpyPos with precision (Fopp a) x (Fopp y) (Fopp t);
 fold radix FtoRradix in |- *; auto with real zarith.
now apply oppBounded.
now apply oppBounded.
now apply oppBounded.
replace (FtoRradix (Fopp a) * FtoRradix x)%R with (- (a * x))%R;
 [ apply ClosestOpp; auto
 | unfold FtoRradix in |- *; rewrite Fopp_correct; ring ].
replace (FtoRradix (Fopp t) + FtoRradix (Fopp y))%R with (- (t + y))%R;
 [ apply ClosestOpp; auto
 | unfold FtoRradix in |- *; repeat rewrite Fopp_correct; ring ].
now apply FcanonicFopp.
now apply FcanonicFopp.
unfold FtoRradix in |- *; rewrite Fopp_correct; auto with real.
unfold FtoRradix in |- *; repeat rewrite Fopp_correct;
 repeat rewrite Rabs_Ropp; auto with real.
unfold FtoRradix in |- *; repeat rewrite Fopp_correct;
 rewrite <- (Rabs_Ropp (- a1 * x1 - - FtoR radix a * FtoR radix x));
 auto with real.
replace (- (- a1 * x1 - - FtoR radix a * FtoR radix x))%R with
 (a1 * x1 - FtoRradix a * FtoRradix x)%R;
 [ auto | unfold FtoRradix in |- *; ring ].
rewrite <- Rabs_Ropp.
replace (- (- y1 - - FtoR radix y))%R with (y1 - FtoRradix y)%R;
 [ idtac | unfold FtoRradix in |- *; ring ].
rewrite FPredFopFSucc; auto with zarith.
rewrite Fopp_Fopp; auto with zarith.
replace (Fulp b radix precision (Fopp (FSucc b radix precision u))) with
 (Fulp b radix precision (FSucc b radix precision u));
 auto.
unfold Fulp in |- *; rewrite Fnormalize_Fopp; simpl in |- *;
 auto with real zarith.
try apply TwoMoreThanOne.
apply MinOrMax3 with precision; auto with zarith; fold FtoRradix in |- *.
try apply TwoMoreThanOne.
cut (FtoRradix t = 0%R); [ intros H' | idtac ].
cut (FtoRradix y = 0%R); [ intros H5 | idtac ].
replace (a1 * x1 + y1 - FtoRradix u)%R with
 (y1 - y + (a1 * x1 - a * x) + a * x)%R.
2: rewrite <- H4; rewrite H5; ring.
apply
 Rle_lt_trans
  with
    (Rabs (y1 - FtoRradix y + (a1 * x1 - FtoRradix a * FtoRradix x)) +
     Rabs (FtoRradix a * FtoRradix x))%R; [ apply Rabs_triang | idtac ].
apply
 Rlt_le_trans
  with
    (/ 4%nat * Fulp b radix precision (FPred b radix precision u) +
     Rabs (FtoRradix a * FtoRradix x))%R; auto with real.
apply Rplus_lt_compat_r; apply Rle_lt_trans with (2 := H2); apply Rabs_triang.
apply
 Rle_trans
  with
    (/ 4%nat * Fulp b radix precision (FPred b radix precision u) +
     / 2%nat * Fulp b radix precision (FPred b radix precision u))%R;
 [ apply Rplus_le_compat_l | idtac ].
replace (FtoRradix a * FtoRradix x)%R with (FtoRradix a * FtoRradix x - t)%R;
 [ idtac | rewrite H'; ring ].
apply Rmult_le_reg_l with (INR 2); auto with real arith.
rewrite <- Rmult_assoc; rewrite Rinv_r; auto with real arith.
ring_simplify (1 * Fulp b radix precision (FPred b radix precision u))%R;
 apply Rle_trans with (Fulp b radix precision t).
unfold FtoRradix in |- *; apply ClosestUlp; auto with zarith.
try apply TwoMoreThanOne.
unfold Fulp in |- *; apply Rle_powerRZ; auto with real zarith.
replace (Fexp (Fnormalize radix b precision t)) with (- dExp b)%Z.
cut (Fbounded b (Fnormalize radix b precision (FPred b radix precision u)));
 [ intros H6; elim H6; auto | apply FnormalizeBounded; auto with zarith ].
try apply TwoMoreThanOne.
apply FBoundedPred; auto with zarith.
try apply TwoMoreThanOne.
replace (Fnormalize radix b precision t) with t.
replace t with (Float 0 (- dExp b)); [ simpl in |- *; auto | idtac ].
apply FcanonicUnique with (radix := radix) (precision := precision) (b := b);
 auto with zarith.
try apply TwoMoreThanOne.
right; repeat (split; simpl in |- *; auto with zarith).
fold FtoRradix in |- *; rewrite H'; unfold FtoRradix, FtoR in |- *;
 simpl in |- *; ring.
apply FcanonicUnique with (radix := radix) (precision := precision) (b := b);
 auto with zarith.
try apply TwoMoreThanOne.
apply FnormalizeCanonic; auto with zarith.
try apply TwoMoreThanOne.
apply sym_eq; apply FnormalizeCorrect; auto with zarith.
try apply TwoMoreThanOne.
apply
 Rle_trans
  with
    ((/ 4%nat + / 2%nat) * Fulp b radix precision (FPred b radix precision u))%R;
 [ right; ring | idtac ].
apply
 Rle_trans with (1 * Fulp b radix precision (FPred b radix precision u))%R;
 [ apply Rmult_le_compat_r | right; ring ].
unfold Fulp in |- *; auto with real zarith.
apply powerRZ_le; apply Rlt_IZR; try apply TwoMoreThanOne.
apply Rmult_le_reg_l with (INR 4); auto with real arith.
apply Rle_trans with 3%R;simpl;[right; field|idtac].
apply Rle_trans with (3+1)%R; auto with real; right; ring.
cut (ProjectorP b radix (Closest b radix));
 [ unfold ProjectorP in |- *; intros V | idtac ].
replace 0%R with (FtoRradix (Float 0 (- dExp b)));
 [ idtac | unfold FtoRradix, FtoR in |- *; simpl in |- *; ring ].
unfold FtoRradix in |- *; apply V; auto.
replace (Float 0 (- dExp b)) with u; fold FtoRradix in |- *.
replace (FtoRradix y) with (FtoRradix t + FtoRradix y)%R; auto.
rewrite H'; ring.
apply FcanonicUnique with (radix := radix) (precision := precision) (b := b);
 auto with zarith.
try apply TwoMoreThanOne.
right; repeat (split; simpl in |- *; auto with zarith).
fold FtoRradix in |- *; rewrite <- H4; unfold FtoRradix, FtoR in |- *;
 simpl in |- *; ring.
apply RoundedProjector; auto.
apply ClosestRoundedModeP with precision; auto with zarith.
try apply TwoMoreThanOne.
cut (forall z : R, (Rabs z <= 0)%R -> z = 0%R); [ intros V; apply V | idtac ].
apply Rmult_le_reg_l with (INR 4); auto with real arith.
apply Rle_trans with (1 := H1); right; rewrite <- H4; auto with real.
ring_simplify (4%nat*0)%R; apply Rabs_R0.
intros z V; case (Req_dec 0 z); auto with real.
intros V'; contradict V.
apply Rlt_not_le; apply Rabs_pos_lt; auto.
Qed.

Theorem Axpy_opt :
 forall (a1 x1 y1 : R) (a x y t u : float),
 (Fbounded b a) -> (Fbounded b x) -> (Fbounded b y) ->
 (Fbounded b t) -> (Fbounded b u) ->
 (Closest b radix (a * x) t) ->
 (Closest b radix (t + y) u) ->
 (Fcanonic radix b u) -> (Fcanonic radix b t) ->
 ((5%nat + 4%nat * (powerRZ radix (- precision))) *
    / (1 - powerRZ radix (- precision)) *
    (Rabs (a * x) + (powerRZ radix (Z.pred (- dExp b))))
    <= Rabs y)%R ->
 (Rabs (y1 - y) + Rabs (a1 * x1 - a * x) <=
    (powerRZ radix (Z.pred (Z.pred (- precision)))) *
    (1 - powerRZ radix (Z.succ (- precision))) * Rabs y +
    - (powerRZ radix (Z.pred (Z.pred (- precision))) * Rabs (a * x)) +
    - powerRZ radix (Z.pred (Z.pred (- dExp b))))%R ->
         (MinOrMax radix b (a1 * x1 + y1) u).
intros a1 x1 y1 a x y t u Fa Fx Fy Ft Fu tDef uDef Cu Ct H1 H2.
apply Axpy_tFlessu with a x y t; auto.
cut (0 < 1 - powerRZ radix (- precision))%R; [ intros H'1 | idtac ].
2: apply Rplus_lt_reg_r with (powerRZ radix (- precision)).
2: ring_simplify (powerRZ radix (- precision) + 0)%R.
2: ring_simplify (powerRZ radix (- precision) + (1 - powerRZ radix (- precision)))%R.
2: replace 1%R with (powerRZ radix 0);
    [ auto with real zarith | simpl in |- *; auto ].
cut (0 < 1 + powerRZ radix (- precision))%R; [ intros H'2 | idtac ].
2: apply Rlt_le_trans with 1%R; auto with real.
2: apply Rle_trans with (1 + 0)%R; auto with real zarith.
cut ((5%nat + 4%nat * powerRZ radix (- precision)) * Rabs t <= Rabs y)%R;
 [ intros H3 | idtac ].
apply Rplus_le_reg_l with (- Rabs (FtoRradix u))%R.
ring_simplify (- Rabs (FtoRradix u) + Rabs (FtoRradix u))%R.
apply
 Rle_trans
  with
    (- ((Rabs y + - Rabs t) * / (1 + powerRZ radix (- precision))) +
     4%nat * Rabs (FtoRradix t))%R;
 [ apply Rplus_le_compat_r; apply Ropp_le_contravar | idtac ].
apply
 Rle_trans
  with
    (Rabs (FtoRradix y + FtoRradix t) * / (1 + powerRZ radix (- precision)))%R.
apply Rmult_le_compat_r; auto with real zarith.
rewrite <- (Rabs_Ropp (FtoRradix t)).
apply Rle_trans with (Rabs (FtoRradix y) - Rabs (- FtoRradix t))%R;
 [ right; ring | idtac ].
replace (FtoRradix y + FtoRradix t)%R with (FtoRradix y - - FtoRradix t)%R;
 [ apply Rabs_triang_inv | ring ].
case Cu; intros H4.
apply Rmult_le_reg_l with (1 + powerRZ radix (- precision))%R; auto.
apply
 Rle_trans
  with
    (Rabs (FtoRradix y + FtoRradix t) *
     ((1 + powerRZ radix (- precision)) * / (1 + powerRZ radix (- precision))))%R;
 [ right; ring; ring | idtac ].
rewrite Rinv_r; auto with real.
ring_simplify (Rabs (FtoRradix y + FtoRradix t) * 1)%R.
rewrite Rmult_plus_distr_r.
apply Rle_trans with (Rabs u + / radix * Fulp b radix precision u)%R.
apply Rplus_le_reg_l with (- Rabs u)%R.
apply Rle_trans with (/ radix * Fulp b radix precision u)%R;
  [idtac|right; ring].
apply Rle_trans with (Rabs (FtoRradix y + FtoRradix t - FtoRradix u)).
apply
 Rle_trans with (Rabs (FtoRradix y + FtoRradix t) - Rabs (FtoRradix u))%R;
 [ right; ring | apply Rabs_triang_inv ].
apply Rmult_le_reg_l with (IZR radix); auto with real zarith.
apply Rle_trans with (Fulp b radix precision u * (radix * / radix))%R;
 [ rewrite Rinv_r; auto with real zarith | right; ring ].
ring_simplify (Fulp b radix precision u * 1)%R.
unfold FtoRradix in |- *; apply ClosestUlp; auto with real zarith.
try apply TwoMoreThanOne.
rewrite Rplus_comm; auto.
ring_simplify (1*Rabs u)%R; apply Rplus_le_compat_l.
apply Rmult_le_reg_l with (IZR radix); auto with real zarith.
rewrite <- Rmult_assoc; rewrite Rinv_r; auto with real zarith.
ring_simplify (1 * Fulp b radix precision u)%R.
apply
 Rle_trans with (Rabs (FtoRradix u) * powerRZ radix (Z.succ (- precision)))%R.
unfold FtoRradix in |- *; apply FulpLe2; auto with zarith.
try apply TwoMoreThanOne.
replace (Fnormalize radix b precision u) with u;
 [ idtac
 | apply
    FcanonicUnique with (radix := radix) (precision := precision) (b := b) ];
 auto with zarith real.
try apply TwoMoreThanOne.
apply FnormalizeCanonic; auto with zarith.
try apply TwoMoreThanOne.
apply sym_eq; apply FnormalizeCorrect; auto with zarith.
try apply TwoMoreThanOne.
unfold Z.succ in |- *; rewrite powerRZ_add; auto with zarith real;
 simpl in |- *; auto with zarith real; right; ring.
replace (FtoRradix y + FtoRradix t)%R with (FtoRradix u); auto with real.
apply Rle_trans with (Rabs (FtoRradix u) * 1)%R;
 [ apply Rmult_le_compat_l | right; ring ]; auto with real.
apply Rabs_pos.
pattern 1%R at 2 in |- *; replace 1%R with (/ 1)%R; auto with real zarith.
apply Rle_Rinv; auto with real zarith.
pattern 1%R at 1 in |- *; replace 1%R with (1 + 0)%R; auto with real zarith.
apply Rplus_le_compat_l.
apply powerRZ_le; apply Rlt_IZR; try apply TwoMoreThanOne.
unfold FtoRradix in |- *; apply plusExact1 with b precision;
 auto with real zarith.
try apply TwoMoreThanOne.
rewrite Rplus_comm; auto.
elim H4; intros H5 H6; elim H6; intros H7 H8; rewrite H7; apply Zmin_Zle;
 elim Fy; elim Ft; auto with zarith.
apply
 Rle_trans
  with
    (((5%nat + 4%nat * powerRZ radix (- precision)) * Rabs (FtoRradix t) +
      - Rabs (FtoRradix y)) * / (1 + powerRZ radix (- precision)))%R.
right; replace (INR 5) with (1 + 4%nat)%R;
 [ idtac
 | replace 1%R with (INR 1); [ rewrite <- plus_INR | idtac ];
    auto with real arith ].
replace (4%nat * Rabs (FtoRradix t))%R with
 (4%nat * Rabs (FtoRradix t) *
  ((1 + powerRZ radix (- precision)) * / (1 + powerRZ radix (- precision))))%R.
ring; ring.
rewrite Rinv_r; auto with real; ring.
apply Rle_trans with (0 * / (1 + powerRZ radix (- precision)))%R;
 [ apply Rmult_le_compat_r; auto with real | right; ring ].
apply Rplus_le_reg_l with (Rabs (FtoRradix y)).
apply
 Rle_trans
  with ((5%nat + 4%nat * powerRZ radix (- precision)) * Rabs (FtoRradix t))%R;
 [ right; ring | ring_simplify (Rabs (FtoRradix y) + 0)%R ];
 auto.
apply Rle_trans with (2 := H1).
rewrite Rmult_assoc.
apply Rmult_le_compat_l; auto with real arith.
apply Rle_trans with (5%nat + 4%nat * 0)%R; auto with real arith.
ring_simplify (5%nat + 4%nat * 0)%R; auto with real arith.
apply Rplus_le_compat_l; apply Rmult_le_compat_l.
rewrite INR_IZR_INZ; apply Rle_IZR; omega.
apply powerRZ_le, Rlt_IZR; try apply TwoMoreThanOne.
apply
 Rle_trans
  with
    (Rabs (FtoRradix a * FtoRradix x) * / (1 - powerRZ radix (- precision)) +
     powerRZ radix (Z.pred (- dExp b)) * / (1 - powerRZ 2%Z (- precision)))%R;
 [ idtac | right; simpl; unfold radix; ring ].
unfold FtoRradix in |- *; apply RoundLeGeneral; auto with real zarith.
apply Rplus_le_compat_l, powerRZ_le, Rlt_IZR; try apply TwoMoreThanOne.
ring_simplify.
apply Rlt_powerRZ; auto with zarith; apply Rlt_IZR; try apply TwoMoreThanOne.
apply Rle_lt_trans with (1 := H2).
apply UlpFlessuGe2 with t; auto.
Qed.

End Axpy.

Section Generic.
Variable b : Fbound.
Variable radix : Z.
Variable p : nat.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.
Hypothesis radixMoreThanOne : (1 < radix)%Z.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ radixMoreThanOne).
Hypothesis pGivesBound : Zpos (vNum b) = Zpower_nat radix p.

Theorem FboundedMbound2Pos :
 (0 < p) ->
 forall z m : Z,
 (0 <= m)%Z ->
 (m <= Zpower_nat radix p)%Z ->
 (- dExp b <= z)%Z ->
 exists c : float, Fbounded b c /\ c = (m * powerRZ radix z)%R :>R /\ (z <= Fexp c)%Z.
intros C z m H' H'0 H'1; case (Zle_lt_or_eq _ _ H'0); intros H'2.
exists (Float m z); split; auto with zarith.
repeat split; simpl in |- *; auto with zarith.
exists (Float 1 (p+z)).
split;[split; simpl; auto with zarith|split].
rewrite pGivesBound; apply Z.le_lt_trans with (Zpower_nat radix 0); auto with zarith.
apply Zpower_nat_monotone_lt; auto with zarith.
unfold FtoRradix, FtoR; simpl; rewrite H'2; rewrite Zpower_nat_Z_powerRZ.
rewrite powerRZ_add; auto with real zarith.
apply IZR_neq; omega.
simpl; auto with zarith.
Qed.

Theorem FboundedMbound2 :
 (0 < p) ->
 forall z m : Z,
 (Z.abs m <= Zpower_nat radix p)%Z ->
 (- dExp b <= z)%Z ->
 exists c : float, Fbounded b c /\ c = (m * powerRZ radix z)%R :>R /\ (z <= Fexp c)%Z.
intros C z m H H0.
case (Zle_or_lt 0 m); intros H1.
case (FboundedMbound2Pos C z (Z.abs m)); auto; try rewrite Z.abs_eq; auto.
intros f (H2, H3); exists f; split; auto.
case (FboundedMbound2Pos C z (Z.abs m)); auto; try rewrite Zabs_eq_opp;
 auto with zarith.
intros f (H2, H3); elim H3; intros; exists (Fopp f); split; auto with zarith.
apply oppBounded; auto with zarith.
split;[idtac|simpl; auto].
rewrite (Fopp_correct radix); auto with arith; fold FtoRradix in |- *;
 rewrite H4.
rewrite Ropp_Ropp_IZR; ring.
Qed.

Hypothesis precisionGreaterThanOne : 1 < p.

Variable z:R.
Variable f:float.
Variable e:Z.

Hypothesis Bf: Fbounded b f.
Hypothesis Cf: Fcanonic radix b f.
Hypothesis zGe: (powerRZ radix (e+p-1) <= z)%R.
Hypothesis zLe: (z <= powerRZ radix (e+p))%R.
Hypothesis fGe: (powerRZ radix (e+p-1) <= f)%R.
Hypothesis eGe: (- dExp b <= e)%Z.

Theorem ClosestSuccPred: (Fcanonic radix b f)
 -> (Rabs(z-f) <= Rabs(z-(FSucc b radix p f)))%R
 -> (Rabs(z-f) <= Rabs(z-(FPred b radix p f)))%R
 -> Closest b radix z f.
intros G; intros; unfold Closest; split; auto.
intros g H1; fold FtoRradix.
cut ((FPred b radix p f) <= z)%R; [intros T1|idtac].
cut (z <= (FSucc b radix p f))%R; [intros T2|idtac].
case (Rle_or_lt g (FPred b radix p f)); intros.
apply Rle_trans with (Rabs (z - f)).
rewrite <- Rabs_Ropp; auto with real.
replace (- (f - z))%R with (z - f)%R; auto with real.
apply Rle_trans with (Rabs (z - FPred b radix p f)); auto with real.
rewrite Rabs_right.
rewrite Rabs_left1; auto with real.
apply Rplus_le_reg_l with (-z)%R.
ring_simplify.
auto with real.
apply Rplus_le_reg_l with z.
ring_simplify.
apply Rle_trans with (1:=H2); auto with real.
apply Rle_ge; auto with real.
apply Rplus_le_reg_l with (FPred b radix p f)%R.
apply Rle_trans with (FPred b radix p f)%R; auto with real.
apply Rle_trans with z; auto with real.
cut (f <= g)%R;[intros|idtac].
case H3; intros.
cut (FSucc b radix p f <= g)%R;[intros|idtac].
apply Rle_trans with (Rabs (z - f)).
rewrite <- Rabs_Ropp; auto with real.
replace (- (f - z))%R with (z - f)%R; auto with real.
apply Rle_trans with (1:=H).
rewrite Rabs_left1.
rewrite Rabs_right.
apply Rle_trans with ((FSucc b radix p f)-z)%R; auto with real.
unfold Rminus; auto with real.
apply Rle_ge; apply Rplus_le_reg_l with z.
apply Rle_trans with z; auto with real.
apply Rle_trans with (FSucc b radix p f)%R; auto with real.
apply Rle_trans with g; auto with real.
apply Rplus_le_reg_l with (FSucc b radix p f); apply Rle_trans with z; auto with real.
apply Rle_trans with (1:=T2); auto with real.
apply Rle_trans with (FNSucc b radix p f).
right; unfold FNSucc; rewrite FcanonicFnormalizeEq; auto with zarith.
unfold FtoRradix; apply FNSuccProp; auto with zarith.
rewrite H4; auto with real.
replace f with (FNSucc b radix p (FPred b radix p f)).
unfold FtoRradix; apply FNSuccProp; auto with zarith.
apply FBoundedPred; auto with zarith.
unfold FNSucc; rewrite FcanonicFnormalizeEq; auto with zarith.
apply FSucPred; auto with zarith.
apply FPredCanonic;auto with zarith.
case (Rle_or_lt z (FSucc b radix p f)); auto; intros.
contradict H; apply Rlt_not_le.
rewrite Rabs_right;[idtac|apply Rle_ge].
rewrite Rabs_right;[idtac|apply Rle_ge].
cut (f < (FSucc b radix p f))%R.
intros; unfold Rminus; auto with real.
unfold FtoRradix; apply FSuccLt; auto with zarith.
apply Rplus_le_reg_l with f.
apply Rle_trans with f; auto with real; apply Rle_trans with (FSucc b radix p f).
apply Rlt_le; unfold FtoRradix; apply FSuccLt; auto with zarith.
apply Rlt_le; apply Rlt_le_trans with (1:=H2); auto with real.
apply Rle_trans with (z-z)%R; auto with real; unfold Rminus; auto with real.
case (Rle_or_lt (FPred b radix p f) z); auto; intros.
contradict H0; apply Rlt_not_le.
cut ((FPred b radix p f) < f)%R.
intros; rewrite Rabs_left1.
rewrite Rabs_left1.
unfold Rminus; auto with real.
apply Rplus_le_reg_l with f.
apply Rle_trans with z; auto with real; apply Rlt_le.
apply Rlt_trans with (1:=H2); apply Rlt_le_trans with (1:=H0); auto with real.
apply Rle_trans with (z-z)%R; auto with real; unfold Rminus; auto with real.
unfold FtoRradix; apply FPredLt; auto with zarith.
Qed.

Theorem ImplyClosest: (Rabs(z-f) <= (powerRZ radix e)/2)%R
  -> Closest b radix z f.
intros; apply ClosestSuccPred; auto.
apply Rle_trans with (1:=H).
apply Rle_trans with (powerRZ radix e - (powerRZ radix e)/2)%R.
right; field; auto with real.
apply Rle_trans with (Rabs (f - FSucc b radix p f) - Rabs(z-f))%R.
unfold Rminus; apply Rplus_le_compat.
rewrite <- Rabs_Ropp.
replace (- (f + - FSucc b radix p f))%R with (FSucc b radix p f - f)%R;[idtac|ring].
unfold FtoRradix; rewrite <- Fminus_correct; auto;rewrite FSuccDiffPos; auto with real zarith.
unfold FtoR; simpl; ring_simplify (1 * powerRZ radix (Fexp f))%R; rewrite Rabs_right.
apply Rle_powerRZ; auto with real zarith.
apply Rle_IZR; omega.
replace e with (Fexp (Float (nNormMin radix p) e)); auto.
apply Fcanonic_Rle_Zle with radix b p; auto with real zarith.
apply FcanonicNnormMin; auto with zarith.
apply Rle_trans with (powerRZ radix (e + p - 1))%R;[right|fold FtoRradix].
unfold nNormMin, FtoR; simpl;rewrite Zpower_nat_Z_powerRZ.
rewrite <- powerRZ_add; auto with real zarith.
rewrite Rabs_right.
replace (pred p+e)%Z with (e+p-1)%Z; auto with real zarith.
rewrite inj_pred; unfold Z.pred; auto with zarith arith.
apply Rle_ge, powerRZ_le, Rlt_IZR; omega.
apply IZR_neq; omega.
rewrite Rabs_right; auto.
apply Rle_ge; apply Rle_trans with (2:=fGe); auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
apply Rle_ge, powerRZ_le, Rlt_IZR; omega.
apply Rle_trans with (2:=fGe); auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
auto with real.
rewrite <- Rabs_Ropp with (z-f)%R.
apply Rle_trans with (Rabs ((f - FSucc b radix p f) - (-(z - f))))%R.
apply Rabs_triang_inv.
ring_simplify ((f - FSucc b radix p f - - (z - f)))%R; auto with real.
right; unfold Rminus; auto with real.
case fGe; intros.
cut ((powerRZ radix (e + p - 1) <= FPred b radix p f))%R;[intros|idtac].
apply Rle_trans with (1:=H).
apply Rle_trans with (powerRZ radix e - (powerRZ radix e)/2)%R.
right; field; auto with real.
apply Rle_trans with (Rabs (f - FPred b radix p f) - Rabs(z-f))%R.
unfold Rminus; apply Rplus_le_compat.
replace ( (f + - FPred b radix p f))%R with (FSucc b radix p (FPred b radix p f) - (FPred b radix p f))%R;[idtac|ring_simplify].
unfold FtoRradix; rewrite <- Fminus_correct; auto;rewrite FSuccDiffPos; auto with real zarith.
unfold FtoR; simpl; ring_simplify (1 * powerRZ radix (Fexp (FPred b radix p f)))%R; rewrite Rabs_right.
apply Rle_powerRZ; auto with real zarith.
apply Rle_IZR; omega.
replace e with (Fexp (Float (nNormMin radix p) e)); auto.
apply Fcanonic_Rle_Zle with radix b p; auto with real zarith.
apply FcanonicNnormMin; auto with zarith.
apply FPredCanonic; auto with zarith.
apply Rle_trans with (powerRZ radix (e + p - 1))%R;[right|fold FtoRradix].
unfold nNormMin, FtoR; simpl;rewrite Zpower_nat_Z_powerRZ.
rewrite <- powerRZ_add; auto with real zarith.
rewrite Rabs_right.
replace (pred p+e)%Z with (e+p-1)%Z; auto with real zarith.
rewrite inj_pred; unfold Z.pred; auto with zarith arith.
apply Rle_ge; apply powerRZ_le, Rlt_IZR; omega.
apply IZR_neq; omega.
rewrite Rabs_right; auto.
apply Rle_ge; apply Rle_trans with (2:=H1); auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
apply Rle_ge; apply powerRZ_le, Rlt_IZR; omega.
apply Rle_trans with (2:=H1); auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
rewrite FSucPred; auto with zarith.
ring.
fold Rminus; auto with real.
rewrite <- Rabs_Ropp with (z-f)%R.
apply Rle_trans with (Rabs ((f - FPred b radix p f) - (-(z - f))))%R.
apply Rabs_triang_inv.
ring_simplify ((f - FPred b radix p f - - (z - f)))%R.
right; unfold Rminus; auto with real.
cut ((powerRZ radix (e + p - 1)= (Float (nNormMin radix p) e)))%R.
intros T; rewrite T.
unfold FtoRradix; apply FPredProp; auto with zarith.
apply FcanonicNnormMin; auto with zarith.
fold FtoRradix; rewrite <- T; auto.
unfold nNormMin, FtoRradix, FtoR; simpl;rewrite Zpower_nat_Z_powerRZ.
rewrite <- powerRZ_add; auto with real zarith.
replace (pred p+e)%Z with (e+p-1)%Z; auto with real zarith.
apply IZR_neq; omega.
cut (FPred b radix p f < f)%R; [intros|unfold FtoRradix; apply FPredLt; auto with zarith].
rewrite Rabs_right.
rewrite Rabs_right.
unfold Rminus;auto with real zarith.
apply Rle_ge; apply Rle_trans with (z-z)%R; auto with real.
right; ring.
apply Rle_trans with (z-f)%R; unfold Rminus; auto with real.
rewrite <- H0; auto with real.
apply Rle_ge; rewrite <- H0; apply Rle_trans with (z-z)%R; unfold Rminus; auto with real.
Qed.

Theorem ImplyClosestStrict: (Rabs(z-f) < (powerRZ radix e)/2)%R
  -> (forall g: float, Closest b radix z g -> (FtoRradix f=g)%R ).
intros.
case (Req_dec (FtoRradix f) (FtoRradix g));auto with real; intros M.
cut (Closest b radix z f);[intros|apply ImplyClosest; auto with real].
cut ((FtoRradix g=2*z-f)%R -> False);[intros Y|idtac].
cut (Rabs (g - z) <= Rabs (f - z))%R;[intros Q1|idtac].
2:elim H0; intros T1 T2; apply T2; auto.
cut (Rabs (f - z) <= Rabs (g - z))%R;[intros Q2|idtac].
2:elim H1; intros T1 T2; apply T2; auto; elim H0; auto.
cut (Rabs (f - z) = Rabs (g - z))%R;[intros Q3; clear Q1 Q2|auto with real].
generalize Q3; unfold Rabs; case (Rcase_abs (f - z)%R);case (Rcase_abs (g - z)%R); intros.
apply Rplus_eq_reg_l with (-z)%R; rewrite Rplus_comm;fold (Rminus f z); rewrite Rplus_comm;fold (Rminus g z).
rewrite <- Ropp_involutive;rewrite <- (Ropp_involutive (f-z)%R);apply Ropp_eq_compat; auto with real.
lapply Y;[intros V; contradict V; auto|idtac].
apply Rplus_eq_reg_l with (-z)%R; apply trans_eq with (g-z)%R; [ring|rewrite <- Q0; ring].
lapply Y;[intros V; contradict V; auto|idtac].
apply Rplus_eq_reg_l with (-z)%R; apply trans_eq with (g-z)%R; [ring|idtac].
rewrite <- (Ropp_involutive (g-z)%R); rewrite <- Q0; ring.
apply Rplus_eq_reg_l with (-z)%R; apply trans_eq with (f-z)%R;[ring|apply trans_eq with (1:=Q0);ring].
intros T; contradict H;apply Rle_not_lt.
replace (z-f)%R with ((g-f)/2)%R;[idtac|rewrite T; field; auto with real].
unfold Rdiv; rewrite Rabs_mult.
rewrite (Rabs_right (/2)%R); [idtac|apply Rle_ge;auto with real].
apply Rmult_le_reg_l with 2%R; auto with real.
apply Rle_trans with (Rabs (g - f))%R;[idtac|right;field; auto with real].
unfold FtoRradix; rewrite <- FnormalizeCorrect with radix b p g; auto.
rewrite <- Fminus_correct; auto.
rewrite <- Fabs_correct; auto.
apply Rle_trans with (FtoR radix (Float (S 0) (Fexp (((Fminus radix (Fnormalize radix b p g) f)))))).
unfold FtoR; simpl.
apply Rle_trans with (powerRZ radix e);[right; field; auto with real|idtac].
apply Rle_trans with (powerRZ radix (Z.min (Fexp (Fnormalize radix b p g)) (Fexp f)))%R;[idtac|right;ring].
apply Rle_powerRZ; auto with zarith real.
apply Rle_IZR; omega.
apply Zmin_Zle.
replace e with (Fexp (Float (nNormMin radix p) e)); auto.
apply Fcanonic_Rle_Zle with radix b p; auto with real zarith.
apply FcanonicNnormMin; auto with zarith.
apply FnormalizeCanonic; auto with zarith; elim H0; auto.
apply Rle_trans with (powerRZ radix (e + p - 1))%R;[right|fold FtoRradix].
unfold nNormMin, FtoR; simpl;rewrite Zpower_nat_Z_powerRZ.
rewrite <- powerRZ_add; auto with real zarith.
rewrite Rabs_right.
replace (pred p+e)%Z with (e+p-1)%Z; auto with real zarith.
rewrite inj_pred; unfold Z.pred; auto with zarith arith.
apply Rle_ge; apply powerRZ_le, Rlt_IZR; omega.
apply IZR_neq; omega.
cut (powerRZ radix (e + p - 1) <= g)%R;[intros Y|idtac].
unfold FtoRradix;rewrite FnormalizeCorrect; auto with zarith.
fold FtoRradix; rewrite Rabs_right; auto.
apply Rle_ge; apply Rle_trans with (2:=Y); auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
cut ((powerRZ radix (e + p - 1)= (Float (nNormMin radix p) e)))%R.
intros U; rewrite U.
case zGe; intros T'.
generalize ClosestMonotone; unfold MonotoneP; intros.
unfold FtoRradix; apply H with b (Float (nNormMin radix p) e) z; auto with zarith real.
rewrite <- U; auto.
unfold FtoRradix; apply RoundedModeProjectorIdem with (P:=(Closest b radix)) (b:=b).
apply ClosestRoundedModeP with p; auto.
cut (Fcanonic radix b (Float (nNormMin radix p) e));[intros G; elim G; intros G'; elim G'; auto|idtac].
apply FcanonicNnormMin; auto with zarith.
right; unfold FtoRradix; apply ClosestIdem with b; auto.
cut (Fcanonic radix b (Float (nNormMin radix p) e));[intros G; elim G; intros G'; elim G'; auto|idtac].
apply FcanonicNnormMin; auto with zarith.
fold FtoRradix; rewrite <- U; rewrite T'; auto.
unfold nNormMin, FtoRradix, FtoR; simpl;rewrite Zpower_nat_Z_powerRZ.
rewrite <- powerRZ_add; auto with real zarith.
replace (pred p+e)%Z with (e+p-1)%Z; auto with real zarith.
apply IZR_neq; omega.
replace e with (Fexp (Float (nNormMin radix p) e)); auto.
apply Fcanonic_Rle_Zle with radix b p; auto with real zarith.
apply FcanonicNnormMin; auto with zarith.
apply Rle_trans with (powerRZ radix (e + p - 1))%R;[right|fold FtoRradix].
unfold nNormMin, FtoR; simpl;rewrite Zpower_nat_Z_powerRZ.
rewrite <- powerRZ_add; auto with real zarith.
rewrite Rabs_right.
replace (pred p+e)%Z with (e+p-1)%Z; auto with real zarith.
rewrite inj_pred; unfold Z.pred; auto with zarith arith.
apply Rle_ge; apply powerRZ_le, Rlt_IZR; omega.
apply IZR_neq; omega.
rewrite Rabs_right; auto.
apply Rle_ge; apply Rle_trans with (2:=fGe); auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
apply RleFexpFabs; auto with zarith.
rewrite Fminus_correct; auto; rewrite FnormalizeCorrect; auto.
fold FtoRradix; auto with real.
Qed.

Theorem ImplyClosestStrict2: (Rabs(z-f) < (powerRZ radix e)/2)%R
  -> (Closest b radix z f) /\ (forall g: float, Closest b radix z g -> (FtoRradix f=g)%R ).
intros; split.
apply ImplyClosest; auto with real.
apply ImplyClosestStrict; auto.
Qed.

End Generic.

Section Generic2.
Variable b : Fbound.
Variable radix : Z.
Variable p : nat.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.
Hypothesis radixMoreThanOne : (1 < radix)%Z.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ radixMoreThanOne).
Hypothesis precisionGreaterThanOne : 1 < p.
Hypothesis pGivesBound : Zpos (vNum b) = Zpower_nat radix p.

Variable z m:R.
Variable f h:float.

Theorem ClosestImplyEven: (EvenClosest b radix p z f) ->
   (exists g: float, (z=g+(powerRZ radix (Fexp g))/2)%R /\ (Fcanonic radix b g) /\ (0 <= Fnum g)%Z)
       -> (FNeven b radix p f).
intros H T1; elim T1; intros g T2; elim T2; intros H0 T3; elim T3; intros H1 H2 ; clear T1 T2 T3.
cut (Fbounded b g);[intros L|apply FcanonicBound with radix; auto with zarith].
cut (g <z)%R;[intros I1|idtac].
cut (z=FSucc b radix p g - powerRZ radix (Fexp g) / 2)%R;[intros H0'|idtac].
cut (z < FSucc b radix p g)%R;[intros I2|idtac].
cut (Closest b radix z g);[intros H4|idtac].
cut (Closest b radix z (FSucc b radix p g));[intros H5|idtac].
generalize EvenClosestMinOrMax; unfold MinOrMaxP; intros T.
elim T with b radix p z f; auto; clear T; intros H6.
elim H; intros H7 H8; case H8; auto; intros.
absurd (FtoRradix f=FSucc b radix p g).
cut (f < FSucc b radix p g)%R; auto with real.
apply Rle_lt_trans with (2:=I2);elim H6; intros K1 K2; elim K2; auto with real.
unfold FtoRradix; apply sym_eq; apply H3; auto.
elim H; intros H7 H8; case H8; auto; intros.
absurd (FtoRradix f=g).
cut (g < f)%R; auto with real.
apply Rlt_le_trans with (1:=I1);elim H6; intros K1 K2; elim K2; auto with real.
unfold FtoRradix; apply sym_eq; apply H3; auto.
apply ClosestSuccPred with p; auto with zarith.
apply FBoundedSuc; auto with zarith.
apply FSuccCanonic; auto with zarith.
rewrite Rabs_left1.
rewrite Rabs_left1.
apply Ropp_le_contravar; unfold Rminus; apply Rplus_le_compat_l; apply Ropp_le_contravar.
apply Rlt_le; apply FSuccLt; auto with zarith.
apply Rplus_le_reg_l with (FtoR radix (FSucc b radix p (FSucc b radix p g))).
ring_simplify.
apply Rlt_le; apply Rlt_trans with (1:=I2).
unfold FtoRradix; apply FSuccLt; auto with zarith.
fold FtoRradix; apply Rle_trans with (z-z)%R; unfold Rminus;auto with real.
rewrite FPredSuc; auto with zarith.
fold FtoRradix; pattern z at 1 in |-*; rewrite H0'; rewrite H0.
ring_simplify (FSucc b radix p g - powerRZ radix (Fexp g) / 2 - FSucc b radix p g)%R.
ring_simplify (g + powerRZ radix (Fexp g) / 2 - g)%R.
rewrite Rabs_Ropp; auto with real.
apply ClosestSuccPred with p; auto with zarith.
fold FtoRradix; pattern z at 1 in |-*; rewrite H0; rewrite H0'.
ring_simplify (FSucc b radix p g - powerRZ radix (Fexp g) / 2 - FSucc b radix p g)%R;
    ring_simplify (g + powerRZ radix (Fexp g) / 2 - g)%R.
rewrite Rabs_Ropp; auto with real.
rewrite Rabs_right.
rewrite Rabs_right.
unfold Rminus; apply Rplus_le_compat_l; apply Ropp_le_contravar.
apply Rlt_le; apply FPredLt; auto with zarith.
apply Rle_ge; apply Rplus_le_reg_l with (FtoR radix (FPred b radix p g)).
ring_simplify.
apply Rlt_le; apply Rlt_trans with (2:=I1).
unfold FtoRradix; apply FPredLt; auto with zarith.
apply Rle_ge; fold FtoRradix; apply Rle_trans with (z-z)%R; unfold Rminus;auto with real.
rewrite H0'; apply Rlt_le_trans with (FSucc b radix p g - 0)%R; auto with real.
unfold Rminus; apply Rplus_lt_compat_l; apply Ropp_lt_contravar.
unfold Rdiv; apply Rmult_lt_0_compat; auto with real zarith.
apply powerRZ_lt, Rlt_IZR; omega.
rewrite H0.
apply Rplus_eq_reg_l with (-g+ powerRZ radix (Fexp g) / 2)%R; ring_simplify.
apply trans_eq with (powerRZ radix (Fexp g));[field; auto with real|idtac].
apply trans_eq with (FtoRradix (Float 1%nat (Fexp g)));[unfold FtoRradix, FtoR; simpl; ring|idtac].
unfold FtoRradix; rewrite <- FSuccDiff1 with b radix p g; auto with zarith.
rewrite Fminus_correct; auto with real; ring.
cut (- nNormMin radix p < Fnum g)%Z; auto with zarith.
apply Z.lt_le_trans with 0%Z; auto with zarith; apply Zplus_lt_reg_l with (nNormMin radix p).
ring_simplify.
unfold nNormMin; auto with zarith.
apply Zpower_nat_less; omega.
rewrite H0; apply Rle_lt_trans with (g+0)%R; auto with real.
apply Rplus_lt_compat_l; unfold Rdiv; apply Rmult_lt_0_compat; auto with real zarith.
apply powerRZ_lt, Rlt_IZR; omega.
Qed.

Theorem ClosestImplyEven_int: (Even radix)%Z
   -> (EvenClosest b radix p z f) -> (Fcanonic radix b f) -> (0 <= f)%R
   -> (z=(powerRZ radix (Fexp f))*(m+1/2))%R -> (exists n:Z, IZR n=m)
   -> (FNeven b radix p f).
intros I; intros.
elim H3; clear H3; intros n H4.
cut (0 <= Fnum f)%Z; [intros|apply LeR0Fnum with radix; auto with real zarith].
case (Zle_lt_or_eq _ _ H3); intros Y1.
case (Z.eq_dec (nNormMin radix p) (Fnum f)).
intros H5; unfold FNeven; rewrite FcanonicFnormalizeEq; auto with zarith.
unfold Feven; rewrite <- H5; unfold nNormMin.
replace (pred p) with (S (pred (pred p))); auto with zarith.
apply EvenExp; auto with zarith.
intros; apply ClosestImplyEven; auto.
exists (Float n (Fexp f)).
split.
rewrite H2; unfold FtoRradix, FtoR; simpl.
rewrite H4; field; auto with real.
cut (Fnum f -1 <= n)%Z;[intros I1|idtac].
cut (n <= Fnum f)%Z;[intros I2|idtac].
cut (0 <= n)%Z;[intros I3|idtac].
split;[idtac|simpl; auto].
case H0; intros.
cut (nNormMin radix p < Fnum f)%Z;[intros K|idtac].
elim H5; intros; elim H6; intros.
left; split;[split| idtac]; simpl; auto.
apply Z.le_lt_trans with (2:=H8); repeat rewrite Z.abs_eq; auto with zarith.
rewrite Z.abs_eq; auto with zarith.
rewrite PosNormMin with radix b p; auto with zarith.
cut (nNormMin radix p <= Fnum f)%Z; auto with zarith.
elim H5; intros.
apply Zmult_le_reg_r with radix; auto with zarith.
rewrite Zmult_comm; rewrite <- PosNormMin with radix b p; auto with zarith.
rewrite Z.abs_eq in H7; auto with zarith.
elim H5; intros T1 T2; elim T1; elim T2; clear T1 T2; intros.
right; split; split; simpl; auto with zarith.
apply Z.le_lt_trans with (2:=H7); rewrite Z.abs_eq; auto with zarith.
rewrite Z.abs_eq; auto with zarith.
apply Z.le_trans with (2:=I1); apply Zplus_le_reg_l with 1%Z.
ring_simplify; auto with zarith.
apply Zle_Rle.
rewrite H4; apply Rplus_le_reg_l with (1/2)%R.
rewrite Rplus_comm; apply Rmult_le_reg_l with (powerRZ radix (Fexp f)); auto with real zarith.
apply powerRZ_lt, Rlt_IZR; omega.
rewrite <- H2; apply Rplus_le_reg_l with (-f)%R.
apply Rle_trans with (z-f)%R;[right;ring|idtac].
apply Rle_trans with (Rabs (z-f))%R;[apply RRle_abs|idtac].
apply Rmult_le_reg_l with (INR 2); auto with real zarith.
apply Rle_trans with (powerRZ radix (Fexp f)).
unfold FtoRradix; apply ClosestExp with b p; auto with zarith.
elim H; auto.
unfold FtoRradix, FtoR; simpl; right; field; auto with real.
apply Zle_Rle.
rewrite H4; apply Rplus_le_reg_l with (1/2)%R.
rewrite (Rplus_comm (1/2)%R m); apply Rmult_le_reg_l with (powerRZ radix (Fexp f)); auto with real zarith.
apply powerRZ_lt, Rlt_IZR; omega.
rewrite <- H2; apply Rplus_le_reg_l with (-z+(1/2)*(powerRZ radix (Fexp f)))%R.
unfold Zminus; rewrite plus_IZR; simpl.
apply Rle_trans with (-(z-f))%R;[right;unfold FtoRradix, FtoR; field; auto with real|idtac].
apply Rle_trans with (Rabs (-(z-f)))%R;[apply RRle_abs|idtac].
rewrite Rabs_Ropp; apply Rmult_le_reg_l with (INR 2); auto with real zarith.
apply Rle_trans with (powerRZ radix (Fexp f)).
unfold FtoRradix; apply ClosestExp with b p; auto with zarith.
elim H; auto.
simpl; right; field; auto with real.
unfold FNeven; rewrite FcanonicFnormalizeEq; auto with zarith.
unfold Feven; rewrite <- Y1; unfold Even; exists 0%Z; auto with zarith.
Qed.

End Generic2.
Section Velt.

Variable radix : Z.
Variable b : Fbound.
Variables s t:nat.
Variables p x q hx: float.

Let b' := Bound
    (P_of_succ_nat (pred (Z.abs_nat (Zpower_nat radix (minus t s)))))
    (dExp b).

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.
Hypothesis radixMoreThanOne : (1 < radix)%Z.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ radixMoreThanOne).

Hypothesis pGivesBound: Zpos (vNum b)=(Zpower_nat radix t).
Hypothesis SLe: (2 <= s)%nat.
Hypothesis SGe: (s <= t-2)%nat.
Hypothesis Fx: Fbounded b x.
Hypothesis pDef: (Closest b radix (x*((powerRZ radix s)+1))%R p).
Hypothesis qDef: (Closest b radix (x-p)%R q).
Hypothesis hxDef:(Closest b radix (q+p)%R hx).
Hypothesis xPos: (0 < x)%R.
Hypothesis Np: Fnormal radix b p.
Hypothesis Nq: Fnormal radix b q.
Hypothesis Nx: Fnormal radix b x.

Lemma p'GivesBound: Zpos (vNum b')=(Zpower_nat radix (minus t s)).
unfold b' in |- *; unfold vNum in |- *.
apply
 trans_eq
  with
    (Z_of_nat
       (nat_of_P
          (P_of_succ_nat
             (pred (Z.abs_nat (Zpower_nat radix (minus t s))))))).
unfold Z_of_nat in |- *; rewrite nat_of_P_o_P_of_succ_nat_eq_succ;
 auto with zarith.
rewrite nat_of_P_o_P_of_succ_nat_eq_succ; auto with arith zarith.
cut (Z.abs (Zpower_nat radix (minus t s)) = Zpower_nat radix (minus t s)).
intros H; pattern (Zpower_nat radix (minus t s)) at 2 in |- *; rewrite <- H.
rewrite Zabs_absolu.
rewrite <- (S_pred (Z.abs_nat (Zpower_nat radix (minus t s))) 0);
 auto with arith zarith.
apply lt_Zlt_inv; simpl in |- *; auto with zarith arith.
rewrite <- Zabs_absolu; rewrite H; auto with arith zarith.
apply Zpower_nat_less; omega.
apply Z.abs_eq; auto with arith zarith.
apply Zpower_NR0; omega.
Qed.

Lemma pPos: (0 <= p)%R.
unfold FtoRradix; apply RleRoundedR0 with b t (Closest b radix) (x * (powerRZ radix s + 1))%R; auto with zarith.
apply ClosestRoundedModeP with t; auto with zarith.
apply Rmult_le_pos; auto with real.
apply Rplus_le_le_0_compat; auto with real.
apply powerRZ_le, Rlt_IZR; omega.
Qed.

Lemma qNeg: (q <= 0)%R.
unfold FtoRradix; apply RleRoundedLessR0 with b t (Closest b radix) (x -p)%R; auto with zarith.
apply ClosestRoundedModeP with t; auto with zarith.
apply Rplus_le_reg_l with (p)%R; ring_simplify.
generalize ClosestMonotone; unfold MonotoneP; intros.
unfold FtoRradix; apply H with b x (x * (powerRZ radix s + 1))%R; auto with zarith real.
apply Rplus_lt_reg_r with (-x)%R; ring_simplify.
apply Rle_lt_trans with (x*0)%R;[right;ring|apply Rmult_lt_compat_l;auto with real zarith].
apply powerRZ_lt, Rlt_IZR; omega.
unfold FtoRradix; apply RoundedModeProjectorIdem with (P:=(Closest b radix)) (b:=b); auto.
apply ClosestRoundedModeP with t; auto with zarith.
Qed.

Lemma RleRRounded: forall (f : float) (z : R),
   Fnormal radix b f -> Closest b radix z f -> (Rabs z <= (Rabs f)*(1+(powerRZ radix (1-t))/2))%R.
intros.
replace z with ((z-f)+f)%R;[idtac|ring].
apply Rle_trans with (Rabs(z-f)+Rabs(f))%R;[apply Rabs_triang|idtac].
apply Rplus_le_reg_l with (- Rabs(f))%R.
ring_simplify.
apply Rmult_le_reg_l with 2%nat; auto with real zarith.
apply Rle_trans with (Fulp b radix t f).
unfold FtoRradix; apply ClosestUlp; auto with zarith.
apply Rle_trans with (Rabs f * powerRZ radix (Z.succ (- t)))%R.
unfold FtoRradix; apply FulpLe2; auto with zarith.
elim H; auto.
rewrite FcanonicFnormalizeEq; auto with zarith; left; auto.
unfold Z.succ; replace (-t+1)%Z with (1-t)%Z;[idtac|ring].
simpl; right; field; auto with real.
Qed.

Lemma hxExact: (FtoRradix hx=p+q)%R.
replace (p+q)%R with (FtoRradix (Fminus radix p (Fopp q))).
2: unfold FtoRradix; rewrite Fminus_correct; auto; rewrite Fopp_correct;ring.
apply sym_eq; unfold FtoRradix; apply ClosestIdem with b.
2: rewrite Fminus_correct; auto; rewrite Fopp_correct; auto with real.
2: fold FtoRradix; replace (p-(-q))%R with (q+p)%R; auto with real;ring.
apply SterbenzAux; auto with zarith.
elim pDef; auto.
apply oppBounded; elim qDef; auto.
generalize ClosestMonotone; unfold MonotoneP; intros.
apply H with b (-(x-p))%R p; auto with zarith real.
apply Rplus_lt_reg_r with (x-p)%R.
ring_simplify; auto with real.
apply ClosestOpp; auto.
unfold FtoRradix; apply RoundedModeProjectorIdem with (P:=(Closest b radix)) (b:=b).
apply ClosestRoundedModeP with t; auto with zarith.
elim pDef; auto.
apply Rmult_le_reg_l with (1-(1+(powerRZ radix (1-t))/2)/(powerRZ radix s + 1))%R.
apply Rmult_lt_reg_l with (2*(powerRZ radix s + 1))%R; auto with real zarith.
apply Rmult_lt_0_compat; auto with real zarith.
apply Rplus_lt_0_compat; auto with real.
apply powerRZ_lt, Rlt_IZR; omega.
rewrite Rmult_0_r.
apply Rlt_le_trans with (2*powerRZ radix s - (powerRZ radix (1- t)))%R;[idtac|right; field].
apply Rplus_lt_reg_r with ((powerRZ radix (1-t)))%R.
ring_simplify.
apply Rle_lt_trans with (powerRZ radix s); auto with real zarith.
apply Rle_powerRZ; auto with zarith; left; apply Rlt_IZR; omega.
apply Rle_lt_trans with (powerRZ radix s + 0)%R; auto with real zarith.
apply Rlt_le_trans with (powerRZ radix s + powerRZ radix s)%R; auto with real zarith.
apply Rplus_lt_compat_l; apply powerRZ_lt, Rlt_IZR; omega.
right; ring.
cut (0 < (powerRZ radix s + 1))%R; auto with real zarith.
apply Rplus_lt_0_compat; auto with real.
apply powerRZ_lt, Rlt_IZR; omega.
apply Rle_trans with ((FtoR radix (Fopp q))*(1 + (powerRZ radix (1- t))/2))%R.
fold FtoRradix; apply Rle_trans with (p-x)%R.
apply Rle_trans with (p - (p*(1 + powerRZ radix (1 - t) / 2) / (powerRZ radix s + 1)))%R;[right|unfold Rminus;apply Rplus_le_compat_l].
field; auto with real zarith.
cut (0 < (powerRZ radix s + 1))%R; auto with real zarith.
apply Rplus_lt_0_compat; auto with real; apply powerRZ_lt, Rlt_IZR; omega.
apply Ropp_le_contravar.
apply Rmult_le_reg_l with (powerRZ radix s + 1)%R; auto with real zarith.
apply Rplus_lt_0_compat; auto with real; apply powerRZ_lt, Rlt_IZR; omega.
apply Rle_trans with ((p * (1 + powerRZ radix (1 - t)/2)))%R;[idtac|right;field].
replace ((powerRZ radix s + 1)* x)%R with (Rabs ((x * (powerRZ radix s + 1))))%R.
replace (FtoRradix p) with (Rabs p).
apply RleRRounded; auto.
apply Rabs_right; apply Rle_ge; apply pPos.
rewrite Rabs_right; auto with real; apply Rle_ge; apply Rmult_le_pos; auto with real zarith.
cut (0 < (powerRZ radix s + 1))%R; auto with real zarith.
apply Rplus_lt_0_compat; auto with real; apply powerRZ_lt, Rlt_IZR; omega.
apply Rgt_not_eq.
apply Rplus_lt_0_compat; auto with real; apply powerRZ_lt, Rlt_IZR; omega.
replace (p - x)%R with (Rabs (x-p))%R.
replace (FtoRradix (Fopp q)) with (Rabs q)%R.
apply RleRRounded; auto.
rewrite Rabs_left1;[idtac|apply qNeg].
unfold FtoRradix; rewrite Fopp_correct; auto with real.
rewrite Rabs_left1; auto with real.
apply Rplus_le_reg_l with (p)%R; ring_simplify.
generalize ClosestMonotone; unfold MonotoneP; intros.
unfold FtoRradix; apply H with b x (x * (powerRZ radix s + 1))%R; auto with zarith real.
apply Rplus_lt_reg_r with (-x)%R; ring_simplify.
apply Rle_lt_trans with (x*0)%R;[right;ring|apply Rmult_lt_compat_l;auto with real zarith].
apply powerRZ_lt, Rlt_IZR; omega.
unfold FtoRradix; apply RoundedModeProjectorIdem with (P:=(Closest b radix)) (b:=b); auto.
apply ClosestRoundedModeP with t; auto with zarith.
fold FtoRradix;apply Rle_trans with ((Fopp q)*((1 - (1 + powerRZ radix (1 - t) / 2) / (powerRZ radix s + 1)) *S 1))%R;[idtac|right;ring].
apply Rmult_le_compat_l.
generalize qNeg; unfold FtoRradix; rewrite Fopp_correct; auto with real.
apply Rle_trans with (3/2)%R.
apply Rplus_le_reg_l with (-1)%R; ring_simplify ((-1 +(1+powerRZ radix (1 - t) / 2)))%R.
apply Rmult_le_reg_l with 2%R; auto with real.
apply Rle_trans with (powerRZ radix (1 - t))%R;[right;field; auto with real|idtac].
apply Rle_trans with (powerRZ radix (0))%R;[idtac|right;simpl;field]; auto with real zarith.
apply Rle_powerRZ; auto with zarith; left; apply Rlt_IZR; omega.
apply Rmult_le_reg_l with (/2)%R; auto with real.
apply Rplus_le_reg_l with (-3/4+(1 + powerRZ radix (1 - t) / 2) / (powerRZ radix s + 1))%R.
apply Rle_trans with ((1 + powerRZ radix (1 - t) / 2) / (powerRZ radix s + 1))%R;[right; field; auto with real|idtac].
cut (0 < powerRZ radix s + 1)%R; auto with real.
apply Rplus_lt_0_compat; auto with real; apply powerRZ_lt, Rlt_IZR; omega.
apply Rmult_le_reg_l with (powerRZ radix s + 1)%R; auto with real zarith.
apply Rplus_lt_0_compat; auto with real; apply powerRZ_lt, Rlt_IZR; omega.
apply Rmult_le_reg_l with 4%R.
auto with real.
apply Rle_trans with (4+ 2*(powerRZ radix (1 - t)))%R;[right; field|idtac].
apply Rgt_not_eq, Rlt_gt.
apply Rplus_lt_0_compat; auto with real; apply powerRZ_lt, Rlt_IZR; omega.
apply Rle_trans with (powerRZ radix s + 1)%R;[idtac|right;simpl;field;auto with real].
apply Rplus_le_compat.
apply Rle_trans with (powerRZ radix 2)%R; [simpl;auto with real zarith|idtac].
ring_simplify (radix * 1)%R; apply Rmult_le_compat; auto with real zarith arith;
 replace (R1 +R1)%R with 2%R by easy; apply Rle_IZR; auto with zarith.
apply Rle_powerRZ; auto with zarith real.
left; apply Rlt_IZR; omega.
apply Rle_trans with (powerRZ radix (1+(1 - t)))%R.
rewrite powerRZ_add.
apply Rmult_le_compat_r.
apply powerRZ_le, Rlt_IZR; omega.
simpl; ring_simplify (radix*1)%R; apply Rle_trans with (IZR 2); auto with real zarith.
apply Rle_IZR; omega.
apply IZR_neq; omega.
change 1%R with (powerRZ radix 0).
apply Rle_powerRZ; auto with zarith real.
apply Rle_IZR; omega.
apply Rgt_not_eq, Rlt_gt.
apply Rplus_lt_0_compat; auto with real; apply powerRZ_lt, Rlt_IZR; omega.
Qed.

Lemma eqLeep: (Fexp q <= Fexp p)%Z.
apply Fcanonic_Rle_Zle with radix b t; auto with zarith.
left; auto.
left; auto.
rewrite Rabs_left1;[idtac|fold FtoRradix; apply qNeg].
rewrite Rabs_right;[idtac|fold FtoRradix; apply Rle_ge; apply pPos].
rewrite <- Fopp_correct.
generalize ClosestMonotone; unfold MonotoneP; intros.
unfold FtoRradix; apply H with b (-(x-p))%R p; auto with zarith real.
apply Rplus_lt_reg_r with (-p)%R; ring_simplify;auto with real.
apply ClosestOpp; auto.
unfold FtoRradix; apply RoundedModeProjectorIdem with (P:=(Closest b radix)) (b:=b); auto.
apply ClosestRoundedModeP with t; auto with zarith.
elim Np; auto.
Qed.

Lemma epLe: (Fexp p <=s+1+Fexp x)%Z.
apply Z.le_trans with (Fexp (Float (Fnum x) (s+1+Fexp x))).
2: simpl; auto with zarith.
apply Fcanonic_Rle_Zle with radix b t; auto with zarith.
left; auto.
elim Nx; intros; left; split; auto with zarith.
elim H; intros; split; simpl; auto with zarith.
rewrite Rabs_right;[idtac|fold FtoRradix; apply Rle_ge; apply pPos].
rewrite Rabs_right;[idtac|fold FtoRradix; apply Rle_ge].
generalize ClosestMonotone; unfold MonotoneP; intros.
unfold FtoRradix; apply H with b (x * (powerRZ radix s + 1))%R (x * (powerRZ radix (s + 1)))%R ; auto with zarith real.
apply Rmult_lt_compat_l; auto with real.
rewrite powerRZ_add; simpl; ring_simplify (radix*1)%R.
apply Rlt_le_trans with (powerRZ radix s * 2%Z)%R.
apply Rlt_le_trans with (powerRZ radix s+powerRZ radix s)%R.
apply Rplus_lt_compat_l; apply Rle_lt_trans with (powerRZ radix 0)%R; auto with real zarith.
apply Rlt_powerRZ; auto with zarith; apply Rlt_IZR; omega.
right; simpl; ring.
apply Rmult_le_compat_l; auto with real zarith.
apply powerRZ_le; apply Rlt_IZR; omega.
apply Rle_IZR; omega.
apply IZR_neq; omega.
replace ((x * powerRZ radix (s + 1)))%R with (FtoRradix (Float (Fnum x) (s + 1 + Fexp x)))%R.
unfold FtoRradix; apply RoundedModeProjectorIdem with (P:=(Closest b radix)) (b:=b); auto.
apply ClosestRoundedModeP with t; auto with zarith.
elim Fx; intros; split; simpl; auto with zarith.
unfold FtoRradix, FtoR; simpl; ring_simplify.
rewrite powerRZ_add. ring.
apply IZR_neq; omega.
apply Rle_trans with (x * powerRZ radix (s + 1))%R; auto with real zarith.
apply Rmult_le_pos; auto with real zarith.
apply powerRZ_le; apply Rlt_IZR; omega.
unfold FtoRradix, FtoR; simpl;right;ring_simplify.
repeat rewrite powerRZ_add; try apply IZR_neq; try omega.
ring.
Qed.

Lemma eqLe: (Fexp q <= s+ Fexp x)%Z \/
  ((FtoRradix q= - powerRZ radix (t+s+Fexp x))%R /\(Rabs (x - hx) <= (powerRZ radix (s + Fexp x))/2)%R).
cut (0 < Fnum x)%Z; [intros L|apply LtR0Fnum with radix; auto with real zarith].
cut ( (Fnum x <= Zpower_nat radix t -radix-1)%Z \/ (Zpower_nat radix t -radix <=Fnum x ))%Z.
2:case (Zle_or_lt (Zpower_nat radix t -radix)%Z (Fnum x));auto with zarith.
intros H; case H; clear H; intros H.
cut (exists g:float, (Fnormal radix b g)/\(FtoRradix g=(Fnum x+radix)*(powerRZ radix (Fexp x+s)))%R/\
   (Fexp g=Fexp x +s)%Z).
intros T; elim T; intros g T'; elim T'; intros H1 T''; elim T''; intros H2 H3; clear T T' T''.
left.
apply Z.le_trans with (Fexp g); auto with zarith.
apply Fcanonic_Rle_Zle with radix b t; auto with zarith.
left; auto.
left; auto.
fold FtoRradix; rewrite <- Rabs_Ropp.
replace (Rabs (-q))%R with (Rabs ((p-x)+((x-p)-q)))%R;[idtac|ring_simplify ((p-x)+((x-p)-q))%R; auto with real].
apply Rle_trans with (Rabs (p-x)+ Rabs((x-p)-q))%R;[apply Rabs_triang|idtac].
apply Rle_trans with ((p - x)+ /2* (powerRZ radix (Fexp q)))%R;[apply Rplus_le_compat|idtac].
rewrite Rabs_right; auto with real.
apply Rle_ge; apply Rplus_le_reg_l with (x)%R; ring_simplify.
generalize ClosestMonotone; unfold MonotoneP; intros.
unfold FtoRradix; apply H0 with b x (x * (powerRZ radix s + 1))%R; auto with zarith real.
apply Rplus_lt_reg_r with (-x)%R; ring_simplify.
apply Rle_lt_trans with (x*0)%R;[right;ring|apply Rmult_lt_compat_l;auto with real zarith].
apply powerRZ_lt;apply Rlt_IZR; omega.
unfold FtoRradix; apply RoundedModeProjectorIdem with (P:=(Closest b radix)) (b:=b); auto.
apply ClosestRoundedModeP with t; auto with zarith.
apply Rmult_le_reg_l with (2%nat); auto with real arith.
apply Rle_trans with (powerRZ radix (Fexp q)).
unfold FtoRradix; apply ClosestExp with b t; auto with zarith.
right; simpl; field; auto with real.
apply Rle_trans with ((x * (powerRZ radix s + 1)+/ 2 * powerRZ radix (Fexp p)) - x + / 2 * powerRZ radix (Fexp q))%R.
apply Rplus_le_compat_r; unfold Rminus; apply Rplus_le_compat_r.
apply Rplus_le_reg_l with (-( x * (powerRZ radix s + 1)))%R.
apply Rle_trans with (Rabs ((- (x * (powerRZ radix s + 1)) + p)))%R; [apply RRle_abs|idtac].
rewrite <- Rabs_Ropp.
replace (- (- (x * (powerRZ radix s + 1)) + p))%R with ((x * (powerRZ radix s + 1)-p))%R;[idtac|ring].
apply Rle_trans with (/ 2 * powerRZ radix (Fexp p))%R;[idtac|right;ring].
apply Rmult_le_reg_l with (2%nat); auto with real arith.
apply Rle_trans with (powerRZ radix (Fexp p)).
unfold FtoRradix; apply ClosestExp with b t; auto with zarith.
right; simpl; field; auto with real.
apply Rle_trans with (x * (powerRZ radix s)+(/ 2 * powerRZ radix (Fexp p)+/ 2 * powerRZ radix (Fexp q)))%R;
  [right;ring|idtac].
apply Rle_trans with (x * powerRZ radix s + powerRZ radix (Fexp p))%R;[apply Rplus_le_compat_l|idtac].
apply Rle_trans with (/ 2 * powerRZ radix (Fexp p) + / 2 * powerRZ radix (Fexp p))%R;
   [apply Rplus_le_compat_l|right; field; auto with real].
apply Rmult_le_compat_l; auto with real; apply Rle_powerRZ; auto with real zarith.
apply Rle_IZR; omega.
apply eqLeep.
apply Rle_trans with (x * powerRZ radix s + radix * powerRZ radix (s+Fexp x))%R;[apply Rplus_le_compat_l|idtac].
apply Rle_trans with (powerRZ radix (s+1+Fexp x))%R;[apply Rle_powerRZ; auto with real zarith; try apply epLe|idtac].
apply Rle_IZR; omega.
right; repeat rewrite powerRZ_add; try apply IZR_neq; try omega.
simpl; ring.
right; rewrite H2; rewrite Rabs_mult.
rewrite Rabs_right;[idtac|apply Rle_ge; auto with real zarith].
rewrite Rabs_right;[idtac|apply Rle_ge; auto with real zarith].
2: apply powerRZ_le; apply Rlt_IZR; omega.
unfold FtoRradix, FtoR; repeat rewrite powerRZ_add; try apply IZR_neq; try omega.
simpl; ring.
apply Rle_trans with ((Fnum x)+0)%R.
ring_simplify.
apply Rle_IZR, LeR0Fnum with radix; auto with real zarith.
apply Rplus_le_compat_l; apply Rle_IZR; omega.
exists (Float (Fnum x +radix) (Fexp x + s)).
elim Nx; elim Fx; intros.
repeat split; simpl; auto with zarith.
rewrite Z.abs_eq; auto with zarith.
rewrite Z.abs_eq in H3; auto with zarith.
apply Z.le_trans with (1:=H3); auto with zarith.
unfold FtoRradix, FtoR; simpl; rewrite plus_IZR; simpl; ring.
cut (FtoRradix p <= powerRZ radix (Fexp x+t+s) + powerRZ radix (Fexp x+t))%R;[intros J1|idtac].
cut (- (x - p) < powerRZ radix (Fexp x) * (powerRZ radix (t + s) + radix + 1))%R;[intros J2|idtac].
cut (FtoRradix (Fopp q) <= powerRZ radix (t + s + Fexp x))%R;[intros V|idtac].
case V; auto; intros V'.
left; replace (Fexp q) with (Fexp (Fopp q)); [idtac|simpl; auto].
replace (s+Fexp x)%Z with (Fexp (FPred b radix t (Float (nNormMin radix t) (s+1+Fexp x)))).
apply Fcanonic_Rle_Zle with radix b t; auto with zarith.
apply FcanonicFopp; left; auto.
apply FPredCanonic; auto with zarith.
apply FcanonicNnormMin; elim Fx; auto with zarith.
rewrite Rabs_right.
rewrite Rabs_right.
apply FPredProp; auto with zarith.
apply FcanonicFopp; left; auto.
apply FcanonicNnormMin; elim Fx; auto with zarith.
fold FtoRradix; apply Rlt_le_trans with (1:=V').
unfold FtoRradix, FtoR, nNormMin; simpl.
rewrite Zpower_nat_Z_powerRZ; rewrite <- powerRZ_add.
replace (pred t +(s+1+Fexp x))%Z with (t+s+Fexp x)%Z; auto with real.
rewrite inj_pred; unfold Z.pred; auto with zarith.
apply IZR_neq; omega.
apply Rle_ge; apply R0RltRlePred; auto with zarith.
apply LtFnumZERO; auto.
simpl; unfold nNormMin; auto with zarith.
apply Zpower_nat_less; omega.
apply Rle_ge; rewrite Fopp_correct; auto; generalize qNeg; auto with real.
rewrite FPredSimpl2; simpl; auto with zarith.
elim Fx; auto with zarith.
right; split.
unfold FtoRradix in V'; rewrite Fopp_correct in V'; auto with real.
fold FtoRradix; rewrite <- V'; ring_simplify; auto with real.
rewrite hxExact.
replace (x-(p+q))%R with ((x-p)- q)%R;[idtac|ring].
case (Rle_or_lt (x-p)%R q).
intros.
rewrite Rabs_left1.
2: apply Rplus_le_reg_l with (FtoRradix q); ring_simplify (q+0)%R.
2: apply Rle_trans with (2:=H0); right; ring.
apply Rle_trans with (q+(p+-x))%R;[right; ring|idtac].
apply Rle_trans with (-(powerRZ radix (t + s + Fexp x)) +
  ((powerRZ radix (Fexp x + t + s) + powerRZ radix (Fexp x + t))+
   -((powerRZ radix t -radix)*powerRZ radix (Fexp x))))%R.
apply Rplus_le_compat.
rewrite <- V'; unfold FtoRradix; rewrite Fopp_correct; auto with real.
apply Rplus_le_compat; auto with real.
apply Ropp_le_contravar.
unfold FtoRradix, FtoR; apply Rmult_le_compat_r; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
apply Rle_trans with (IZR (Zpower_nat radix t - radix)).
unfold Zminus; rewrite plus_IZR; rewrite Zpower_nat_Z_powerRZ;
  rewrite Ropp_Ropp_IZR; auto with real zarith.
now apply Rle_IZR.
replace (t+s+Fexp x)%Z with (Fexp x+t+s)%Z; auto with zarith.
ring_simplify.
pattern (IZR radix) at 4 in |-*; replace (IZR radix) with (powerRZ radix 1).
repeat rewrite <- powerRZ_add.
rewrite Zplus_comm.
ring_simplify.
apply Rmult_le_reg_l with 2%R; auto with real.
apply Rle_trans with (radix*powerRZ radix (1+Fexp x))%R.
apply Rmult_le_compat_r; replace 2%R with (IZR 2); auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
apply Rle_IZR; omega.
apply Rle_trans with (powerRZ radix (2+Fexp x)).
right; repeat rewrite powerRZ_add; try apply IZR_neq; try omega.
simpl; ring.
apply Rle_trans with (powerRZ radix (s+Fexp x)).
apply Rle_powerRZ; auto with real zarith.
apply Rle_IZR; omega.
right; field; auto with real.
apply IZR_neq; omega.
apply IZR_neq; omega.
simpl; ring.
intros.
apply Rmult_le_reg_l with 2%R; auto with real.
apply Rle_trans with (powerRZ radix (s + Fexp x));[idtac|right; field; auto with real].
apply Rplus_le_reg_l with (-( Rabs (x - p - q)))%R.
ring_simplify (- Rabs (x - p - q) + 2 * Rabs (x - p - q))%R.
cut (exists qplus:float, (Fbounded b qplus)/\ (qplus-q=powerRZ radix (s+Fexp x))%R
  /\ qplus=FNSucc b radix t q).
intros T; elim T; intros qplus T'; elim T'; intros H1 T''; elim T'';
  intros; clear T T' T''.
apply Rle_trans with (Rabs (x-p-qplus))%R.
elim qDef; fold FtoRradix; intros.
replace (x-p-q)%R with (-(q-(x-p)))%R;[rewrite Rabs_Ropp|ring].
replace (x-p-qplus)%R with (-(qplus-(x-p)))%R;[rewrite Rabs_Ropp|ring].
apply H5; auto.
rewrite Rabs_left1.
rewrite Rabs_right.
rewrite <- H2; right; ring.
apply Rle_ge; apply Rplus_le_reg_l with (FtoRradix q); ring_simplify (q+0)%R.
apply Rlt_le; apply Rlt_le_trans with (1:=H0); right; ring.
apply Rplus_le_reg_l with (FtoRradix qplus).
ring_simplify.
cut (isMax b radix (x-p)%R qplus).
intros H4; elim H4; intros H5 H6; elim H6; intros H7 H8; auto with real.
rewrite H3; apply MinMax; auto with zarith real.
generalize ClosestMinOrMax; unfold MinOrMaxP; intros T.
case (T b radix (x-p)%R q); auto.
clear T; intros W; elim W; intros T1 T2; elim T2; intros H4 H5; clear T1 T2 H5.
fold FtoRradix in H4; contradict H4; auto with real.
exists (FNSucc b radix t q); split.
apply FcanonicBound with radix; auto.
apply FNSuccCanonic; auto with zarith; elim Nq; auto.
split; auto.
unfold FNSucc; rewrite FcanonicFnormalizeEq; auto with zarith.
2: left; auto.
unfold FtoRradix; rewrite <- Fminus_correct; auto.
replace q with (Float (-(nNormMin radix t)) (s+1+Fexp x)).
rewrite FSuccDiff3; auto with zarith real.
unfold FtoR; simpl.
replace (Z.pred (s+1+Fexp x))%Z with (s+Fexp x)%Z; unfold Z.pred; auto with real zarith.
simpl; elim Fx; auto with zarith.
apply FnormalUnique with radix b t; auto with zarith.
replace (Float (- nNormMin radix t) (s + 1 + Fexp x)) with
   (Fopp (Float (nNormMin radix t) (s + 1 + Fexp x)));
   [idtac|unfold Fopp; auto with zarith].
apply FnormalFop; auto.
apply FnormalNnormMin; auto with zarith; elim Fx; auto with zarith.
apply trans_eq with (-(-FtoR radix q))%R; auto with real.
rewrite <- Fopp_correct; fold FtoRradix.
rewrite V'; unfold FtoRradix, FtoR, nNormMin; simpl.
rewrite Ropp_Ropp_IZR; rewrite Zpower_nat_Z_powerRZ.
apply trans_eq with (-(powerRZ radix (pred t) * powerRZ radix (s + 1 + Fexp x)))%R;
  auto with real.
rewrite <- powerRZ_add.
replace ((pred t + (s + 1 + Fexp x)))%Z with (t + s + Fexp x)%Z; auto with real.
rewrite inj_pred; auto with zarith; unfold Z.pred; ring.
apply IZR_neq; omega.
apply Rle_trans with (FtoRradix (Float 1%Z (t+s+Fexp x)));[idtac|right; unfold FtoRradix, FtoR; simpl; ring].
generalize ClosestMonotone; unfold MonotoneP; intros.
unfold FtoRradix; apply H0 with b (-(x-p))%R ((powerRZ radix (Fexp x))*(powerRZ radix (t+s)+radix+1))%R;
  auto with zarith real.
apply ClosestOpp; auto.
clear H0; generalize ClosestCompatible; unfold CompatibleP; intros T.
cut (Fbounded b (Float 1 (t + s + Fexp x)));[intros H1|idtac].
2: split; simpl; elim Fx; intros; auto with zarith.
apply T with (powerRZ radix (Fexp x) * (powerRZ radix (t + s) + radix + 1))%R
  (Fnormalize radix b t (Float 1 (t + s + Fexp x))); auto with real.
2: rewrite FnormalizeCorrect; auto with zarith.
apply ImplyClosest with t (Fexp x+s+1)%Z; auto with zarith.
apply FnormalizeBounded; auto with zarith.
apply FnormalizeCanonic; auto with zarith.
apply Rle_trans with (powerRZ radix (Fexp x) * powerRZ radix (t + s))%R.
rewrite <- powerRZ_add.
replace (Fexp x + s + 1 + t - 1)%Z with (Fexp x + (t + s))%Z; auto with real zarith.
apply IZR_neq; omega.
apply Rmult_le_compat_l.
apply powerRZ_le, IZR_lt; omega.
apply Rle_trans with (powerRZ radix (t + s) +0)%R; auto with real zarith.
rewrite Rplus_assoc; apply Rplus_le_compat_l; auto with real zarith.
rewrite <- plus_IZR; apply Rle_IZR; omega.
rewrite FnormalizeCorrect; auto with zarith; unfold FtoR; simpl; right.
replace (Fexp x + s + 1 + t - 1)%Z with (t + s + Fexp x)%Z; ring.
elim Fx; auto with zarith.
rewrite FnormalizeCorrect; auto with zarith; unfold FtoR; simpl.
replace (powerRZ radix (Fexp x) * (powerRZ radix (t + s) + radix + 1) -
       1 * powerRZ radix (t + s + Fexp x))%R
  with (powerRZ radix (Fexp x) *(radix+1))%R.
rewrite Rabs_right.
replace (Fexp x+s+1)%Z with (Fexp x+(1+s))%Z;[idtac|ring].
rewrite powerRZ_add.
apply Rle_trans with (powerRZ radix (Fexp x) * (powerRZ radix (1+s) */ 2))%R;[idtac|right;unfold Rdiv; ring].
apply Rmult_le_compat_l; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
apply Rmult_le_reg_l with 2%R; auto with real.
apply Rle_trans with (powerRZ radix (1+s));[idtac|right; field; auto with real].
rewrite powerRZ_add.
apply Rmult_le_compat; auto with real zarith.
rewrite <- plus_IZR; apply Rle_IZR; omega.
simpl; ring_simplify (radix*1)%R; apply Rle_IZR; auto with real zarith.
2: apply IZR_neq; omega.
2: apply IZR_neq; omega.
apply Rle_trans with (powerRZ radix 2); [idtac|apply Rle_powerRZ; try apply Rle_IZR; auto with real zarith].
simpl; ring_simplify (radix*1)%R.
apply Rle_trans with (radix+radix)%R.
apply Rplus_le_compat_l, Rle_IZR; omega.
apply Rle_trans with (2*radix)%R; [right;ring|idtac].
apply Rmult_le_compat_r; apply Rle_IZR; auto with real zarith.
apply Rle_ge; apply Rmult_le_pos; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
apply Rplus_le_le_0_compat; apply Rle_IZR; auto with zarith.
repeat rewrite powerRZ_add; try apply IZR_neq; try omega.
ring.
replace (-(x-p))%R with (p+-x)%R by ring.
apply Rle_lt_trans with ((powerRZ radix (Fexp x + t + s) + powerRZ radix (Fexp x + t))+
   -(powerRZ radix (Fexp x + t) - radix*powerRZ radix (Fexp x)))%R.
apply Rplus_le_compat; auto with real.
apply Ropp_le_contravar; unfold FtoRradix, FtoR; rewrite powerRZ_add.
apply Rle_trans with ((powerRZ radix t - radix)*powerRZ radix (Fexp x))%R;[right;ring|idtac].
apply Rmult_le_compat_r.
apply powerRZ_le, Rlt_IZR; omega.
rewrite <- Zpower_nat_Z_powerRZ.
apply Rle_trans with (IZR ((Zpower_nat radix t - radix))); auto with real zarith.
unfold Zminus; rewrite plus_IZR; rewrite Ropp_Ropp_IZR; simpl; auto with real.
now apply Rle_IZR.
apply IZR_neq; omega.
apply Rplus_lt_reg_r with (radix * powerRZ radix (Fexp x))%R.
ring_simplify.
rewrite Rplus_comm, Rplus_assoc; apply Rplus_lt_compat_l.
rewrite <- powerRZ_add.
replace (Fexp x+(t+s))%Z with (Fexp x +t+s)%Z; auto with zarith real.
apply Rle_lt_trans with (powerRZ radix (Fexp x + t + s)+0)%R; auto with real zarith.
apply Rplus_lt_compat_l.
apply powerRZ_lt, Rlt_IZR; omega.
apply IZR_neq; omega.
cut ( powerRZ radix (Fexp x + t + s) + powerRZ radix (Fexp x + t)=
   Float (Zpower_nat radix (pred t)+Zpower_nat radix (Z.abs_nat (t-s-1))) (Fexp x+s+1))%R.
cut (Fbounded b (Float (Zpower_nat radix (pred t)+Zpower_nat radix (Z.abs_nat (t-s-1))) (Fexp x+s+1))).
intros.
rewrite H1.
generalize ClosestMonotone; unfold MonotoneP; intros.
unfold FtoRradix; apply H2 with b (x * (powerRZ radix s + 1))%R
   (powerRZ radix (Fexp x + t + s) + powerRZ radix (Fexp x + t))%R; auto with zarith real.
unfold FtoRradix, FtoR.
apply Rlt_le_trans with ((powerRZ radix t * powerRZ radix (Fexp x) * (powerRZ radix s + 1)))%R.
apply Rmult_lt_compat_r; auto with real zarith.
apply Rplus_lt_0_compat; auto with real.
apply powerRZ_lt, Rlt_IZR; omega.
apply Rmult_lt_compat_r; auto with real zarith.
apply powerRZ_lt, Rlt_IZR; omega.
rewrite <- Zpower_nat_Z_powerRZ; rewrite <- pGivesBound; elim Fx; intros.
rewrite Z.abs_eq in H3; auto with zarith real.
now apply Rlt_IZR.
right;repeat rewrite powerRZ_add; try apply IZR_neq; try omega.
ring.
rewrite H1; unfold FtoRradix; apply RoundedModeProjectorIdem with (P:=(Closest b radix)) (b:=b); auto.
apply ClosestRoundedModeP with t; auto with zarith.
split; simpl.
rewrite pGivesBound; rewrite Z.abs_eq ; auto with zarith.
apply Z.lt_le_trans with (Zpower_nat radix (pred t) + Zpower_nat radix (pred t))%Z.
apply Zplus_lt_compat_l.
apply Zpower_nat_monotone_lt; auto with zarith.
pattern t at 3 in |-*; replace t with (1+(pred t))%nat; auto with zarith.
rewrite Zpower_nat_is_exp; replace (Zpower_nat radix 1) with radix; auto with zarith.
apply Z.le_trans with (2*Zpower_nat radix (pred t))%Z; auto with zarith.
apply Zmult_le_compat_r; try omega.
apply Zpower_NR0; omega.
unfold Zpower_nat; simpl; auto with zarith.
apply Z.add_nonneg_nonneg; apply Zpower_NR0; omega.
elim Fx; auto with zarith.
unfold FtoRradix, FtoR; simpl; rewrite plus_IZR.
repeat rewrite Zpower_nat_Z_powerRZ.
rewrite Rmult_plus_distr_r.
repeat rewrite <- powerRZ_add.
replace (Z.abs_nat (t - s - 1) + (Fexp x + s + 1))%Z with (Fexp x + t)%Z.
replace (pred t + (Fexp x + s + 1))%Z with (Fexp x + t + s)%Z; auto with real.
rewrite inj_pred; unfold Z.pred; auto with zarith.
rewrite <- Zabs_absolu; rewrite Z.abs_eq; auto with zarith.
apply IZR_neq; omega.
apply IZR_neq; omega.
Qed.

Lemma eqGe: (s+ Fexp x <= Fexp q)%Z.
case (Rle_or_lt ((powerRZ radix (Fexp x))*((powerRZ radix (t-1))+radix))%R x);intros H.
apply Z.le_trans with (Fexp (Float (nNormMin radix t) (s+Fexp x)));[simpl; auto with zarith|idtac].
apply Z.le_trans with (Fexp (Fopp q));[idtac|simpl; auto with zarith].
apply Fcanonic_Rle_Zle with radix b t; auto with zarith.
apply FcanonicNnormMin; auto with zarith.
elim Fx; auto with zarith.
apply FcanonicFopp; left; auto.
rewrite Fopp_correct; fold FtoRradix; rewrite Rabs_Ropp.
replace (FtoRradix q) with ((-(p-x))-((x-p)-q))%R;[idtac|ring].
apply Rle_trans with (2:=Rabs_triang_inv (-(p-x))%R ((x-p)-q)%R).
apply Rle_trans with ((x*(powerRZ radix s)-(powerRZ radix (Fexp p))/2)-(powerRZ radix (Fexp q))/2)%R.
apply Rle_trans with (((powerRZ radix (Fexp x) * (powerRZ radix (t - 1) + radix))) * powerRZ radix s - powerRZ radix (s+1+Fexp x) / 2 - powerRZ radix (s+1+Fexp x) / 2)%R.
unfold nNormMin, FtoRradix, FtoR; simpl;rewrite Zpower_nat_Z_powerRZ.
rewrite <- powerRZ_add.
2: apply IZR_neq; omega.
rewrite Rabs_right.
2: apply Rle_ge, powerRZ_le, Rlt_IZR; omega.
replace (pred t+(s+Fexp x))%Z with (t-1+(s+Fexp x))%Z; auto with real zarith.
apply Rle_trans with (powerRZ radix (Fexp x) * (powerRZ radix (t - 1) + radix) *
    powerRZ radix s - powerRZ radix (s + 1 + Fexp x))%R;[idtac|right;field; auto with real].
rewrite Rmult_plus_distr_l.
rewrite Rmult_plus_distr_r.
pattern (IZR radix) at 6 in |-*; replace (IZR radix) with (powerRZ radix 1)%R; [idtac|simpl; ring].
repeat rewrite <- powerRZ_add; try apply IZR_neq; try omega.
replace (Fexp x + (t - 1) + s)%Z with (t - 1 + (s + Fexp x))%Z;[idtac|ring].
replace (s + 1+ Fexp x)%Z with (Fexp x+1+s)%Z;[right|idtac];ring.
unfold Rminus; apply Rplus_le_compat.
apply Rplus_le_compat.
apply Rmult_le_compat_r; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
apply Ropp_le_contravar; unfold Rdiv; apply Rmult_le_compat_r; auto with real.
apply Rle_powerRZ; auto with real zarith.
apply Rle_IZR; omega.
apply epLe.
apply Ropp_le_contravar; unfold Rdiv; apply Rmult_le_compat_r; auto with real.
apply Rle_powerRZ; auto with real zarith.
apply Rle_IZR; omega.
apply Z.le_trans with (Fexp p);[apply eqLeep|apply epLe].
unfold Rminus; apply Rplus_le_compat.
rewrite Rabs_left1.
apply Rplus_le_reg_l with ((powerRZ radix (Fexp p) / 2)+x-p)%R.
ring_simplify.
apply Rle_trans with ((x * (powerRZ radix s + 1))-p)%R;[right;ring|idtac].
apply Rle_trans with (Rabs ((x * (powerRZ radix s + 1))-p))%R;[apply RRle_abs|idtac].
apply Rmult_le_reg_l with (INR 2); auto with real zarith.
apply Rle_trans with (powerRZ radix (Fexp p)).
unfold FtoRradix; apply ClosestExp with b t; auto with zarith.
simpl; right; field; auto with real.
apply Rplus_le_reg_l with (p)%R; ring_simplify.
generalize ClosestMonotone; unfold MonotoneP; intros.
unfold FtoRradix; apply H0 with b x (x * (powerRZ radix s + 1))%R; auto with zarith real.
apply Rplus_lt_reg_r with (-x)%R; ring_simplify.
apply Rle_lt_trans with (x*0)%R;[right;ring|apply Rmult_lt_compat_l;auto with real zarith].
apply powerRZ_lt, Rlt_IZR; omega.
unfold FtoRradix; apply RoundedModeProjectorIdem with (P:=(Closest b radix)) (b:=b); auto.
apply ClosestRoundedModeP with t; auto with zarith.
apply Ropp_le_contravar.
replace (x + - p + - q)%R with ((x-p)-q)%R;[idtac|ring].
apply Rmult_le_reg_l with (INR 2); auto with real zarith.
apply Rle_trans with (powerRZ radix (Fexp q)).
unfold FtoRradix; apply ClosestExp with b t; auto with zarith.
simpl; right; field; auto with real.
case (Rle_or_lt (powerRZ radix (Fexp x) * (powerRZ radix (t - 1) + 1))%R x); intros H'.
cut ((powerRZ radix (Fexp x) * ((powerRZ radix (s+t-1))+(powerRZ radix (t-1))+(powerRZ radix s))) <= p)%R;[intros|idtac].
apply Z.le_trans with (Fexp (Float (nNormMin radix t) (Fexp x+s)));[simpl;auto with zarith|idtac].
apply Fcanonic_Rle_Zle with radix b t; auto with real zarith.
apply FcanonicNnormMin; auto with zarith; elim Fx; auto with zarith.
left; auto.
rewrite Rabs_right.
rewrite Rabs_left1.
rewrite <- Fopp_correct.
generalize ClosestMonotone; unfold MonotoneP; intros.
unfold FtoRradix; apply H1 with b (Float (nNormMin radix t) (Fexp x + s)) (-(x-p))%R.
2: unfold FtoRradix; apply RoundedModeProjectorIdem with (P:=(Closest b radix)) (b:=b); auto.
2:apply ClosestRoundedModeP with t; auto with zarith.
2: apply FcanonicBound with radix; apply FcanonicNnormMin; auto with zarith; elim Fx; auto with zarith.
2: apply ClosestOpp; auto.
clear H1; replace (-(x-p))%R with (p+-x)%R by ring.
apply Rlt_le_trans with (((powerRZ radix (Fexp x) *
     (powerRZ radix (s + t - 1) + powerRZ radix (t - 1) + powerRZ radix s)))+
     -(powerRZ radix (Fexp x) * (powerRZ radix (t - 1) + radix)))%R; auto with real.
2: apply Rplus_le_compat; auto with real.
unfold FtoRradix, FtoR,nNormMin; simpl; rewrite Zpower_nat_Z_powerRZ.
repeat rewrite Rmult_plus_distr_l.
repeat rewrite <- powerRZ_add; try apply IZR_neq; try omega.
replace (pred t + (Fexp x + s))%Z with (Fexp x+(s + t - 1))%Z;[idtac|rewrite inj_pred; unfold Z.pred; auto with zarith].
apply Rplus_lt_reg_r with ((radix * powerRZ radix (Fexp x))- (powerRZ radix (Fexp x+(s + t - 1))))%R.
ring_simplify.
apply Rle_lt_trans with (powerRZ radix (1+Fexp x)); auto with real zarith.
rewrite powerRZ_add; try apply IZR_neq; try omega.
simpl; right;ring.
apply Rlt_powerRZ; auto with zarith; apply Rlt_IZR; omega.
apply qNeg.
apply Rle_ge; apply LeFnumZERO; simpl; auto with zarith.
unfold nNormMin; auto with zarith.
apply Zpower_NR0; omega.
cut ( (powerRZ radix (Fexp x) *
    (powerRZ radix (s + t - 1) + powerRZ radix (t - 1) + powerRZ radix s))=
    (Float ((Zpower_nat radix (pred t) + Zpower_nat radix (Z.abs_nat (t -s-1)) + 1)) ((Fexp x)+s)))%R;[intros V1|idtac].
cut (Fbounded b ( Float ((Zpower_nat radix (pred t) + Zpower_nat radix (Z.abs_nat (t -s-1)) + 1))%Z ((Fexp x)+s)));[intros V2|idtac].
rewrite V1.
generalize ClosestMonotone; unfold MonotoneP; intros.
unfold FtoRradix; apply H0 with b ( (Float
      (Zpower_nat radix (pred t) + Zpower_nat radix (Z.abs_nat (t - s - 1)) +
       1) (Fexp x + s))%R) (x * (powerRZ radix s + 1))%R; auto with zarith real.
2: unfold FtoRradix; apply RoundedModeProjectorIdem with (P:=(Closest b radix)) (b:=b); auto.
2:apply ClosestRoundedModeP with t; auto with zarith.
rewrite <- V1; clear H0 V2 V1.
apply Rlt_le_trans with ( (powerRZ radix (Fexp x) * (powerRZ radix (t - 1) + 1) *(powerRZ radix s + 1)))%R.
2: apply Rmult_le_compat_r; auto with real zarith.
2: apply Rplus_le_le_0_compat; auto with real; apply powerRZ_le, Rlt_IZR; omega.
rewrite Rmult_assoc; apply Rmult_lt_compat_l; auto with real zarith.
apply powerRZ_lt, Rlt_IZR; omega.
rewrite Rmult_plus_distr_l; rewrite Rmult_plus_distr_r.
rewrite <- powerRZ_add.
2: apply IZR_neq; omega.
apply Rlt_le_trans with ((powerRZ radix (s + t - 1))+ powerRZ radix (t - 1)+ powerRZ radix s+1)%R .
repeat rewrite Rplus_assoc; repeat apply Rplus_lt_compat_l; auto with real zarith.
replace (s+t-1)%Z with (t-1+s)%Z; [right; ring|ring].
split; simpl.
rewrite pGivesBound; rewrite Z.abs_eq; auto with zarith.
apply Z.lt_le_trans with ((Zpower_nat radix (pred t) + Zpower_nat radix (pred (pred t)) + Zpower_nat radix (pred (pred t))))%Z.
repeat rewrite <- Zplus_assoc;apply Zplus_lt_compat_l.
cut (Zpower_nat radix (Z.abs_nat (t - s - 1)) <= Zpower_nat radix (pred (pred t)))%Z;[intros|idtac].
cut (1 <Zpower_nat radix (pred (pred t)))%Z;auto with zarith.
apply Z.le_lt_trans with (Zpower_nat radix 0)%Z; auto with zarith.
apply Zpower_nat_monotone_lt; omega.
apply Zpower_nat_monotone_le; auto with zarith.
pattern t at 4 in |-*; replace t with ((pred t)+1); auto with zarith.
cut ((Zpower_nat radix 1)=radix)%Z;[intros K|unfold Zpower_nat; simpl; auto with zarith].
rewrite Zpower_nat_is_exp; rewrite K.
apply Z.le_trans with (Zpower_nat radix (pred t)+ (Zpower_nat radix (pred t)))%Z.
rewrite <- Zplus_assoc; apply Zplus_le_compat_l.
pattern (pred t) at 3 in |-*; replace (pred t) with ((pred (pred t))+1); auto with zarith.
rewrite Zpower_nat_is_exp; rewrite K.
apply Z.le_trans with (Zpower_nat radix (pred (pred t)) * 2)%Z; auto with zarith.
apply Zmult_le_compat_l; try omega.
apply Zpower_NR0; omega.
apply Z.le_trans with (Zpower_nat radix ((pred t)) * 2)%Z; auto with zarith.
apply Zmult_le_compat_l; try omega.
apply Zpower_NR0; omega.
apply Z.add_nonneg_nonneg; try apply Z.add_nonneg_nonneg; try apply Zpower_NR0; omega.
elim Fx; auto with zarith.
unfold FtoRradix, FtoR; simpl.
repeat rewrite plus_IZR; repeat rewrite Zpower_nat_Z_powerRZ.
rewrite inj_pred; auto with zarith; unfold Z.pred.
rewrite <- Zabs_absolu; rewrite Z.abs_eq; auto with zarith.
simpl; unfold Zminus; repeat rewrite powerRZ_add; try apply IZR_neq; try omega.
repeat rewrite Rmult_plus_distr_l; repeat rewrite Rmult_plus_distr_r.
repeat rewrite <- powerRZ_add; try apply IZR_neq; try omega.
replace (Fexp x + (s + t + - (1)))%Z with (t + -1 + (Fexp x + s))%Z by ring.
replace (Fexp x + (t + - (1)))%Z with (t + - s + - (1) + (Fexp x + s))%Z; ring.
cut (FtoRradix x= powerRZ radix (Fexp x+t-1))%R;[intros K|idtac].
cut (FtoRradix p= powerRZ radix (Fexp x+t-1)*((powerRZ radix s + 1)))%R;[intros K'|idtac].
replace q with (Fopp (Float ((nNormMin radix t)) (s+Fexp x)%Z)); simpl; auto with zarith.
apply FcanonicUnique with radix b t; auto with zarith.
apply FcanonicFopp; apply FcanonicNnormMin; auto with zarith.
elim Fx; auto with zarith.
left; auto.
apply ClosestIdem with b; auto.
apply FcanonicBound with radix; apply FcanonicFopp; apply FcanonicNnormMin; auto with zarith;elim Fx; auto with zarith.
replace (FtoR radix (Fopp (Float (nNormMin radix t) (s + Fexp x)))) with (x-p)%R; auto.
rewrite K'; rewrite K; rewrite Fopp_correct; unfold FtoR; simpl.
unfold nNormMin; rewrite Zpower_nat_Z_powerRZ.
rewrite inj_pred; auto with zarith.
unfold Z.pred, Zminus; repeat rewrite powerRZ_add; try apply IZR_neq; try omega.
simpl; ring.
cut ((powerRZ radix (Fexp x + t - 1) * (powerRZ radix s + 1))=
   (FtoRradix (Float (Zpower_nat radix s +1) (Fexp x+t-1))))%R;[intros L; rewrite L|idtac].
unfold FtoRradix; apply sym_eq; apply ClosestIdem with b; auto.
split;simpl;[idtac|elim Fx; auto with zarith].
rewrite pGivesBound; rewrite Z.abs_eq; auto with zarith.
replace t with ((pred t)+1); auto with zarith; rewrite Zpower_nat_is_exp.
apply Z.lt_le_trans with (Zpower_nat radix (pred t)+ Zpower_nat radix (pred t))%Z; auto with zarith.
apply Zplus_lt_compat.
apply Zpower_nat_monotone_lt; auto with zarith.
replace 1%Z with (Zpower_nat radix 0)%Z; auto with zarith.
apply Zpower_nat_monotone_lt; auto with zarith.
apply Z.le_trans with (Zpower_nat radix (pred t) * 2)%Z; auto with zarith.
apply Zmult_le_compat_l; auto with zarith.
unfold Zpower_nat; simpl; auto with zarith.
apply Zpower_NR0; auto with zarith.
apply Z.add_nonneg_nonneg; try apply Zpower_NR0; auto with zarith.
fold FtoRradix; rewrite <- L; rewrite <- K; auto.
unfold FtoRradix, FtoR; simpl; rewrite plus_IZR; rewrite Zpower_nat_Z_powerRZ; simpl; ring.
cut (Fnum x=Zpower_nat radix (pred t));[intros|idtac].
unfold FtoRradix, FtoR; rewrite H0; rewrite Zpower_nat_Z_powerRZ.
rewrite <- powerRZ_add.
replace (pred t + Fexp x)%Z with (Fexp x + t - 1)%Z; auto with real zarith.
apply IZR_neq; omega.
cut ( Zpower_nat radix (pred t) <= Fnum x)%Z;[intros P1|idtac].
cut ( Fnum x < Zpower_nat radix (pred t) +1)%Z;[intros P2; auto with zarith|idtac].
apply Zlt_Rlt; rewrite plus_IZR; rewrite Zpower_nat_Z_powerRZ.
apply Rmult_lt_reg_l with (powerRZ radix (Fexp x)); auto with real zarith.
apply powerRZ_lt, IZR_lt; omega.
apply Rle_lt_trans with (FtoRradix x);[right; unfold FtoRradix, FtoR;ring|idtac].
apply Rlt_le_trans with (1:=H'); right; simpl.
replace (t-1)%Z with (Z_of_nat (pred t));[ring|rewrite inj_pred; auto with zarith].
apply Zmult_le_reg_r with radix; auto with zarith.
apply Z.le_trans with (Zpos (vNum b)); [rewrite pGivesBound|rewrite Zmult_comm].
pattern radix at 2 in |-*; replace radix with (Zpower_nat radix 1).
rewrite <- Zpower_nat_is_exp.
replace (pred t + 1) with t; auto with zarith.
unfold Zpower_nat; simpl; auto with zarith.
elim Nx; intros.
rewrite Zabs_Zmult in H1.
rewrite Z.abs_eq in H1; auto with zarith.
rewrite Z.abs_eq in H1; auto with zarith.
apply LeR0Fnum with radix; auto with real.
Qed.

Lemma eqEqual: (Fexp q=s+Fexp x)%Z \/
  ((FtoRradix q= - powerRZ radix (t+s+Fexp x))%R /\
     (Rabs (x - hx) <= (powerRZ radix (s + Fexp x))/2)%R).
generalize eqLe; generalize eqGe; intros.
case H0; auto.
intros; left; auto with zarith.
Qed.

Lemma Veltkamp_aux_aux: forall v:float, (FtoRradix v=hx) -> Fcanonic radix b' v ->
  (Rabs (x-v) <= (powerRZ radix (s+Fexp x)) /2)%R
  -> (powerRZ radix (t-1+Fexp x) <= v)%R.
intros.
case (Rle_or_lt (powerRZ radix (t-1)+(powerRZ radix s)/2)%R (Fnum x)); intros W.
fold FtoRradix; apply Rplus_le_reg_l with (-v+x-powerRZ radix (t-1+Fexp x))%R.
ring_simplify.
apply Rle_trans with (x-v)%R; [right; ring|idtac].
apply Rle_trans with (Rabs (x-v))%R;[apply RRle_abs|idtac].
unfold FtoRradix; apply Rle_trans with (1:=H1).
unfold FtoR; rewrite powerRZ_add; unfold Rdiv.
2: apply IZR_neq; omega.
rewrite powerRZ_add.
2: apply IZR_neq; omega.
apply Rle_trans with (powerRZ radix (Fexp x) * (powerRZ radix s * / 2))%R;[right;ring|idtac].
apply Rle_trans with (powerRZ radix (Fexp x) * (- powerRZ radix (t - 1) + Fnum x))%R;[idtac|right;ring].
apply Rmult_le_compat_l; auto with real zarith.
apply powerRZ_le, IZR_lt; omega.
apply Rle_trans with ( - powerRZ radix (t - 1) + (powerRZ radix (t - 1) + powerRZ radix s / 2))%R;
   auto with real zarith.
right; unfold Rdiv; ring.
cut (exists eps:Z, (FtoRradix x=powerRZ radix (Fexp x)*(powerRZ radix (t-1) + eps))%R
    /\ (0 <= eps)%Z /\ (eps < (powerRZ radix s)/2)%R).
intros T; elim T; intros eps T'; elim T'; intros H3 T''; elim T''; intros H4 H5; clear T T' T''.
fold FtoRradix; rewrite H; rewrite hxExact.
cut (Fbounded b (Float (Zpower_nat radix (pred t)+Zpower_nat radix (Z.abs_nat (t-s-1))+eps) (s+Fexp x)));
   [intros Yp|idtac].
cut (FtoRradix (Float (Zpower_nat radix (pred t)+Zpower_nat radix (Z.abs_nat (t-s-1))+eps) (s+Fexp x))
  = powerRZ radix (Fexp x)*(powerRZ radix (t+s-1)+ powerRZ radix (t-1)+eps*powerRZ radix s))%R;
   [intros Yp'|idtac].
cut (Fbounded b (Float ((Zpower_nat radix (pred t)+eps)) (s+Fexp x))); [intros Yq|idtac].
cut (FtoRradix (Float ((Zpower_nat radix (pred t)+eps)) (s+Fexp x))
  = powerRZ radix (Fexp x)*(powerRZ radix (t+s-1)+ eps*powerRZ radix s))%R;
  [intros Yq'|idtac].
cut (FtoRradix p=(powerRZ radix (Fexp x) *
         (powerRZ radix (t + s - 1) + powerRZ radix (t - 1) +
          eps * powerRZ radix s)))%R;[intros YYp|idtac].
cut (FtoRradix (Fopp q)=(powerRZ radix (Fexp x) *
         (powerRZ radix (t + s - 1) + eps * powerRZ radix s)))%R;[intros YYq|idtac].
replace (FtoRradix q) with (-(-q))%R; [idtac|ring]; unfold FtoRradix; rewrite <- Fopp_correct.
fold FtoRradix; rewrite YYp; rewrite YYq; right.
repeat rewrite powerRZ_add; try ring.
apply IZR_neq; omega.
rewrite <- Yq'.
unfold FtoRradix; apply sym_eq.
apply ImplyClosestStrict with b t (-(x-p))%R (s+Fexp x)%Z; auto with zarith.
left; split; auto.
rewrite pGivesBound; rewrite Zabs_Zmult; rewrite Z.abs_eq; auto with zarith.
simpl; rewrite Z.abs_eq; auto with zarith.
apply Z.le_trans with (radix*((Zpower_nat radix (pred t) + 0)))%Z; auto with zarith.
pattern t at 1; replace t with (1+(pred t)); auto with zarith.
rewrite Zpower_nat_is_exp.
replace (Zpower_nat radix 1) with radix; auto with zarith.
unfold Zpower_nat; simpl; auto with zarith.
apply Z.add_nonneg_nonneg; try apply Zpower_NR0; omega.
rewrite YYp; rewrite H3.
ring_simplify.
rewrite <- powerRZ_add.
2: apply IZR_neq; omega.
replace (Fexp x+(t+s-1))%Z with (s+Fexp x+t-1)%Z;[idtac|ring].
apply Rplus_le_reg_l with ( -(powerRZ radix (s + Fexp x + t - 1))+eps * powerRZ radix (Fexp x))%R.
ring_simplify.
rewrite Rmult_assoc; apply Rmult_le_compat_l; auto with real zarith.
apply Rle_IZR; easy.
apply Rle_trans with (powerRZ radix (Fexp x)*1)%R; auto with real; apply Rmult_le_compat_l; auto with real zarith.
apply powerRZ_le, IZR_lt; omega.
apply Rle_trans with (powerRZ radix 0); auto with real zarith.
apply Rle_powerRZ; try apply Rle_IZR; omega.
fold FtoRradix; rewrite Yq'.
apply Rle_trans with (powerRZ radix (Fexp x) *(powerRZ radix (t + s - 1) + 0))%R.
ring_simplify (powerRZ radix (t + s - 1) + 0)%R.
rewrite <- powerRZ_add.
replace (s + Fexp x + t - 1)%Z with (Fexp x+(t + s - 1))%Z; auto with real zarith.
apply IZR_neq; omega.
apply Rmult_le_compat_l; auto with real zarith.
apply powerRZ_le, IZR_lt; omega.
apply Rplus_le_compat_l; apply Rmult_le_pos; auto with real zarith.
now apply Rle_IZR.
apply powerRZ_le, IZR_lt; omega.
elim Fx; auto with zarith.
fold FtoRradix; rewrite Yq'; rewrite YYp; rewrite H3.
ring_simplify ( (-
       (powerRZ radix (Fexp x) * (powerRZ radix (t - 1) + eps) -
        powerRZ radix (Fexp x) *
        (powerRZ radix (t + s - 1) + powerRZ radix (t - 1) +
         eps * powerRZ radix s)) -
       powerRZ radix (Fexp x) *
       (powerRZ radix (t + s - 1) + eps * powerRZ radix s)))%R.
rewrite Ropp_mult_distr_l_reverse; rewrite Rabs_Ropp; rewrite Rabs_mult.
rewrite Rabs_right;[idtac|apply Rle_ge; auto with real zarith].
rewrite Rabs_right;[idtac|apply Rle_ge; auto with real zarith].
unfold Rdiv; rewrite powerRZ_add.
apply Rlt_le_trans with ((powerRZ radix s*/2) * powerRZ radix (Fexp x))%R;[idtac|right;ring].
rewrite Rmult_comm; apply Rmult_lt_compat_r; auto with real zarith.
apply powerRZ_lt, IZR_lt; omega.
apply IZR_neq; omega.
apply Rle_IZR; easy.
apply powerRZ_le, IZR_lt; omega.
apply ClosestOpp; auto.
rewrite <- Yp'.
unfold FtoRradix; apply sym_eq.
apply ImplyClosestStrict with b t (x * (powerRZ radix s + 1))%R (s+Fexp x)%Z; auto with zarith.
left; split; auto.
rewrite pGivesBound; rewrite Zabs_Zmult; rewrite Z.abs_eq; auto with zarith.
simpl; rewrite Z.abs_eq; auto with zarith.
apply Z.le_trans with (radix*((Zpower_nat radix (pred t) + 0+0)))%Z; auto with zarith.
pattern t at 1; replace t with (1+(pred t)); auto with zarith.
rewrite Zpower_nat_is_exp.
rewrite Zpower_nat_1; auto with zarith.
apply Zmult_le_compat_l; auto with zarith.
apply Zplus_le_compat; auto with zarith.
apply Zplus_le_compat_l.
apply Zpower_NR0; omega.
apply Z.add_nonneg_nonneg; auto.
apply Z.add_nonneg_nonneg; apply Zpower_NR0; omega.
rewrite H3.
apply Rle_trans with (powerRZ radix (Fexp x) * (powerRZ radix (t - 1) + 0) *
    (powerRZ radix s + 0))%R; auto with real zarith.
right; ring_simplify.
unfold Zminus; repeat rewrite powerRZ_add; try apply IZR_neq; auto with real zarith; ring.
apply Rmult_le_compat; auto with real zarith.
ring_simplify (powerRZ radix (t - 1) + 0)%R; apply Rmult_le_pos; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
apply powerRZ_le, Rlt_IZR; omega.
ring_simplify; apply powerRZ_le, Rlt_IZR; omega.
apply Rmult_le_compat_l.
apply powerRZ_le, Rlt_IZR; omega.
apply Rplus_le_compat_l, Rle_IZR; omega.
fold FtoRradix; rewrite Yp'.
apply Rle_trans with (powerRZ radix (Fexp x) *(powerRZ radix (t + s - 1) + 0+0))%R.
right; ring_simplify.
rewrite <- powerRZ_add; try apply IZR_neq; auto with real zarith.
replace (s + Fexp x + t - 1)%Z with (Fexp x+(t + s - 1))%Z; auto with real zarith.
apply Rmult_le_compat_l; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
apply Rplus_le_compat.
apply Rplus_le_compat_l.
apply powerRZ_le, Rlt_IZR; omega.
apply Rmult_le_pos; try apply Rle_IZR; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
elim Fx; auto with zarith.
fold FtoRradix; rewrite Yp';rewrite H3.
ring_simplify (powerRZ radix (Fexp x) * (powerRZ radix (t - 1) + eps) *
       (powerRZ radix s + 1) -
       powerRZ radix (Fexp x) *
       (powerRZ radix (t + s - 1) + powerRZ radix (t - 1) +
        eps * powerRZ radix s))%R.
replace (t+s-1)%Z with (s+(t-1))%Z; [rewrite powerRZ_add; try apply IZR_neq|idtac]; auto with real zarith.
ring_simplify (powerRZ radix (Fexp x) * powerRZ radix (t - 1) * powerRZ radix s +
    powerRZ radix (Fexp x) * eps -
    powerRZ radix (Fexp x) * (powerRZ radix s * powerRZ radix (t - 1)))%R.
rewrite Rabs_mult.
rewrite Rabs_right;[idtac|apply Rle_ge; auto with real zarith].
2: apply powerRZ_le, Rlt_IZR; omega.
rewrite Rabs_right;[idtac|apply Rle_ge; auto with real zarith].
2: now apply IZR_le.
unfold Rdiv; rewrite powerRZ_add; try apply IZR_neq; auto with real zarith.
apply Rlt_le_trans with (powerRZ radix (Fexp x)* (powerRZ radix s*/2))%R;[idtac|right;ring].
apply Rmult_lt_compat_l; auto with real zarith.
apply powerRZ_lt, Rlt_IZR; omega.
unfold FtoRradix, FtoR; simpl.
rewrite plus_IZR; rewrite Zpower_nat_Z_powerRZ; rewrite inj_pred; auto with zarith.
unfold Z.pred, Zminus; repeat rewrite powerRZ_add; try apply IZR_neq; auto with real zarith.
simpl; ring.
split; simpl.
clear Yp'; elim Yp; simpl; intros.
rewrite Z.abs_eq; auto with zarith.
rewrite Z.abs_eq in H2; auto with zarith.
apply Z.le_lt_trans with (2:=H2).
rewrite <- Zplus_assoc; apply Zplus_le_compat_l; auto with zarith.
apply Z.le_trans with (0+eps)%Z; auto with zarith; apply Zplus_le_compat_r; auto with zarith.
apply Zpower_NR0; omega.
apply Z.add_nonneg_nonneg; auto with zarith.
apply Z.add_nonneg_nonneg; apply Zpower_NR0; omega.
apply Z.add_nonneg_nonneg; try apply Zpower_NR0; omega.
elim Fx; auto with zarith.
unfold FtoRradix, FtoR; simpl.
rewrite plus_IZR; rewrite plus_IZR.
repeat rewrite Zpower_nat_Z_powerRZ; rewrite inj_pred; auto with zarith.
rewrite <- Zabs_absolu; rewrite Z.abs_eq; auto with zarith.
unfold Z.pred, Zminus; repeat rewrite powerRZ_add; try apply IZR_neq; auto with real zarith.
rewrite powerRZ_Zopp.
2: apply IZR_neq; omega.
simpl; field; auto with real zarith.
split; apply Rgt_not_eq.
apply Rlt_IZR; omega.
apply powerRZ_lt, Rlt_IZR; omega.
split; simpl.
2: elim Fx; auto with zarith.
rewrite Z.abs_eq; auto with zarith.
rewrite pGivesBound; apply Zlt_Rlt.
rewrite plus_IZR;rewrite plus_IZR; repeat rewrite Zpower_nat_Z_powerRZ.
rewrite inj_pred; auto with zarith.
rewrite <- Zabs_absolu; rewrite Z.abs_eq; auto with zarith.
apply Rlt_le_trans with (powerRZ radix (Z.pred t) + powerRZ radix (t - s - 1) + powerRZ radix s / 2)%R;
  auto with real.
apply Rle_trans with (powerRZ radix (Z.pred t)+powerRZ radix (t-2)+powerRZ radix (t-2))%R.
apply Rplus_le_compat.
apply Rplus_le_compat_l; auto with real zarith.
apply Rle_powerRZ; try apply Rle_IZR; auto with real zarith.
apply Rle_trans with (powerRZ radix s).
apply Rmult_le_reg_l with (2%R); auto with real.
apply Rle_trans with (powerRZ radix s);[right; field; auto with real|auto with real zarith].
rewrite <- (Rmult_1_l (powerRZ _ _)) at 1; apply Rmult_le_compat_r.
apply powerRZ_le, IZR_lt; omega.
auto with real.
apply Rle_powerRZ; try apply Rle_IZR; auto with zarith real.
replace (Z.pred t) with (t-1)%Z;[idtac|unfold Z.pred; ring].
apply Rle_trans with (powerRZ radix (t-1)+powerRZ radix (t-1))%R.
rewrite Rplus_assoc; apply Rplus_le_compat_l.
apply Rle_trans with (2*powerRZ radix (t - 2))%R; [right;ring|idtac].
apply Rle_trans with (radix*powerRZ radix (t - 2))%R; [apply Rmult_le_compat_r; auto with real zarith|idtac].
apply powerRZ_le, IZR_lt; omega.
apply Rle_IZR; omega.
replace (t-1)%Z with (1+(t-2))%Z;[rewrite powerRZ_add; try apply IZR_neq; simpl|idtac]; auto with real zarith.
right; ring.
apply Rle_trans with (2*powerRZ radix (t - 1))%R; [right;ring|idtac].
apply Rle_trans with (radix*powerRZ radix (t - 1))%R; [apply Rmult_le_compat_r; auto with real zarith|idtac].
apply powerRZ_le, IZR_lt; omega.
apply Rle_IZR; omega.
pattern (Z_of_nat t)%Z at 2 in |-*; replace (Z_of_nat t)%Z with (1+(t-1))%Z;
  [rewrite powerRZ_add; simpl|ring].
right; ring.
apply IZR_neq; omega.
apply Z.add_nonneg_nonneg; auto with zarith.
apply Z.add_nonneg_nonneg; apply Zpower_NR0; omega.
exists (Fnum x- Zpower_nat radix (pred t))%Z; split.
unfold Zminus; rewrite plus_IZR; rewrite Ropp_Ropp_IZR; rewrite Zpower_nat_Z_powerRZ.
replace (Z_of_nat (pred t)) with (t+-(1))%Z; [idtac|rewrite inj_pred; auto with zarith].
unfold FtoRradix, FtoR; ring.
split.
apply Zplus_le_reg_l with (Zpower_nat radix (pred t)).
ring_simplify.
apply Zmult_le_reg_r with radix; auto with zarith.
elim Nx; intros.
rewrite Zabs_Zmult in H3.
rewrite Z.abs_eq in H3; auto with zarith.
rewrite Z.abs_eq in H3; [idtac|apply LeR0Fnum with radix; auto with zarith real].
rewrite Zmult_comm with (Fnum x) radix.
apply Z.le_trans with (2:=H3); rewrite pGivesBound.
pattern t at 2; replace t with (1+(pred t)); auto with zarith.
rewrite Zpower_nat_is_exp.
replace ( Zpower_nat radix 1) with radix;[idtac|unfold Zpower_nat; simpl]; auto with zarith.
unfold Zminus; rewrite plus_IZR; rewrite Ropp_Ropp_IZR; rewrite Zpower_nat_Z_powerRZ.
replace (Z_of_nat (pred t)) with (t-1)%Z; [idtac|rewrite inj_pred; auto with zarith].
apply Rplus_lt_reg_l with (powerRZ radix (t - 1)).
apply Rle_lt_trans with (2:=W); right;ring.
Qed.

Lemma Veltkamp_aux:
   (Rabs (x-hx) <= (powerRZ radix (s+Fexp x)) /2)%R /\
   (exists hx':float, (FtoRradix hx'=hx) /\ (Closest b' radix x hx')
    /\ (s+Fexp x <= Fexp hx')%Z).
generalize p'GivesBound;intros J.
cut (powerRZ radix (t - 1 + Fexp x) <= x)%R;[intros xGe|idtac].
2:rewrite powerRZ_add; try apply IZR_neq; auto with real zarith; unfold FtoRradix, FtoR.
2:apply Rmult_le_compat_r; auto with real zarith.
2: apply powerRZ_le, IZR_lt; omega.
2:apply Rmult_le_reg_l with radix; try apply Rlt_IZR; auto with real zarith.
2:apply Rle_trans with (powerRZ radix t).
2:unfold Zminus; rewrite powerRZ_add; try apply IZR_neq; auto with real zarith; simpl; right; field; auto with real.
2: apply IZR_neq; omega.
2:rewrite <- Zpower_nat_Z_powerRZ; rewrite <- pGivesBound; rewrite <- mult_IZR; elim Nx; intros H H0.
2:rewrite Zabs_Zmult in H0; rewrite Z.abs_eq in H0; auto with zarith.
2:apply Rle_IZR; rewrite Z.abs_eq in H0; auto with zarith real.
2:apply LeR0Fnum with radix; auto with real.
cut (Rabs (x - hx) <= (powerRZ radix (s + Fexp x))/2)%R;[intros|idtac].
2:case eqEqual; intros L.
2:fold FtoRradix; rewrite hxExact.
2:replace (x-(p+q))%R with ((x-p)-q)%R;[apply Rmult_le_reg_l with (INR 2); auto with real zarith|ring].
2:apply Rle_trans with (powerRZ radix (Fexp q)).
2:unfold FtoRradix; apply ClosestExp with b t; auto with zarith.
2:rewrite L; simpl; right; field; auto with real.
2:elim L; auto.
cut (exists v:float, (FtoRradix v=hx)/\(Fcanonic radix b' v)).
intros T; elim T; intros v T'; elim T'; intros; clear T T'.
split; auto.
exists v; split; auto.
cut (Fbounded b' v);[intros Fv|apply FcanonicBound with radix; auto].
cut (Rabs (x - FtoR radix v) <= (powerRZ radix (s + Fexp x))/2)%R;[intros|idtac].
2: fold FtoRradix; rewrite H0; auto with real.
split.
apply ImplyClosest with (minus t s) (s+Fexp x)%Z; auto with zarith real.
replace (s + Fexp x + (t - s)%nat - 1)%Z with ((t-1)+(Fexp x))%Z;[auto with real|rewrite inj_minus1; auto with zarith].
2:elim Fx; unfold b'; simpl; auto with zarith.
replace (s + Fexp x + (t - s)%nat - 1)%Z with ((t-1)+(Fexp x))%Z;[idtac|rewrite inj_minus1; auto with zarith].
fold FtoRradix; apply Veltkamp_aux_aux; auto.
assert (s+Fexp x-1 < Fexp v)%Z; auto with zarith.
assert (t-1+Fexp x < t-s+Fexp v)%Z; auto with zarith.
apply Zlt_powerRZ with radix; try apply Rle_IZR; auto with real zarith.
apply Rle_lt_trans with (FtoRradix v).
apply Veltkamp_aux_aux; auto.
apply Rle_lt_trans with (1:=RRle_abs v).
unfold FtoRradix; rewrite <- Fabs_correct; auto with zarith.
unfold FtoR, Fabs; simpl.
rewrite powerRZ_add.
2: apply IZR_neq; omega.
apply Rmult_lt_compat_r; auto with real zarith.
apply powerRZ_lt, IZR_lt; omega.
elim Fv; intros.
apply Rlt_le_trans with (IZR (Zpos (vNum b'))); try apply Rlt_IZR; auto with real zarith.
rewrite J; rewrite Zpower_nat_Z_powerRZ; auto with real zarith.
rewrite inj_minus1; auto with real zarith.
cut (exists c:float, (FtoRradix c=hx) /\ (Fbounded b' c)).
intros T; elim T; intros c H'; elim H'; intros.
exists (Fnormalize radix b' (t-s) c); split.
unfold FtoRradix; rewrite FnormalizeCorrect; auto with real zarith.
apply FnormalizeCanonic; auto with zarith.
case eqEqual; intros L.
generalize FboundedMbound; intros P.
elim P with radix b' (t-s) (s+Fexp x)%Z (Fnum (Fplus radix p q)); auto with zarith; clear P.
intros v H'; elim H'; intros ; clear H'.
exists v; split; auto.
rewrite hxExact; unfold FtoRradix; rewrite <- Fplus_correct; auto.
rewrite H1; unfold FtoR; replace (s+ Fexp x)%Z with (Fexp (Fplus radix p q)); auto with real.
unfold Fplus; simpl.
rewrite Zmin_le2;[auto|apply eqLeep].
2: elim Fx; unfold b'; simpl; auto with zarith.
cut ( (Z.abs (Fnum (Fplus radix p q)) < ((Zpower_nat radix (t - s))+1)))%Z; auto with zarith.
apply Zlt_Rlt.
apply Rmult_lt_reg_l with (powerRZ radix (Fexp (Fplus radix p q))); auto with real zarith.
apply powerRZ_lt, IZR_lt; omega.
apply Rle_lt_trans with (Rabs (Fplus radix p q)).
unfold FtoRradix; rewrite <- Fabs_correct; auto.
unfold Fabs, FtoR; simpl; auto with real.
unfold FtoRradix; rewrite Fplus_correct; auto.
fold FtoRradix; rewrite <- hxExact.
replace (FtoRradix hx) with (-(x-hx)+x)%R;[idtac|ring].
apply Rle_lt_trans with (Rabs (-(x-hx))+ Rabs(x))%R;[apply Rabs_triang|idtac].
rewrite Rabs_Ropp.
apply Rle_lt_trans with ((powerRZ radix (s + Fexp x))/2 + Rabs x)%R; auto with real.
apply Rlt_le_trans with ((powerRZ radix (s + Fexp x))/2 + (powerRZ radix (t+Fexp x)))%R.
apply Rplus_lt_compat_l.
unfold FtoRradix; rewrite <- Fabs_correct; auto; unfold FtoR; simpl.
rewrite powerRZ_add; try apply IZR_neq; auto with real zarith; apply Rmult_lt_compat_r; auto with real zarith.
apply powerRZ_lt, IZR_lt; omega.
elim Fx; intros; rewrite <- Zpower_nat_Z_powerRZ; rewrite <- pGivesBound; try apply Rlt_IZR; auto with real zarith.
replace (Fexp (Fplus radix p q)) with (s+ Fexp x)%Z.
2:unfold Fplus; simpl.
2:rewrite Zmin_le2;[auto|apply eqLeep].
rewrite plus_IZR; rewrite Zpower_nat_Z_powerRZ; simpl.
rewrite Rmult_plus_distr_l; rewrite <- powerRZ_add.
rewrite Rplus_comm; apply Rplus_le_compat.
rewrite inj_minus1; auto with real zarith.
ring_simplify ((s + Fexp x + (t - s)))%Z.
rewrite Zplus_comm; auto with real.
unfold Rdiv; apply Rmult_le_compat_l; auto with real zarith.
apply powerRZ_le, IZR_lt; omega.
apply Rle_trans with (/1)%R; auto with real.
apply IZR_neq; omega.
elim L; clear L; intros L1 L2.
cut (Fexp q=s+1+Fexp x)%Z;[intros L3|idtac].
2:cut (q=Float (-(nNormMin radix t)) (s+1+Fexp x));[intros I; rewrite I; simpl; auto|idtac].
2:apply FnormalUnique with radix b t; auto with zarith.
2:replace (Float (- nNormMin radix t) (s + 1 + Fexp x)) with
   (Fopp (Float (nNormMin radix t) (s + 1 + Fexp x)));
   [idtac|unfold Fopp; auto with zarith].
2:apply FnormalFop; auto.
2:apply FnormalNnormMin; auto with zarith; elim Fx; auto with zarith.
2:fold FtoRradix; rewrite L1; unfold FtoRradix, FtoR, nNormMin; simpl.
2:rewrite Ropp_Ropp_IZR; rewrite Zpower_nat_Z_powerRZ.
2:apply trans_eq with (-(powerRZ radix (pred t) * powerRZ radix (s + 1 + Fexp x)))%R;
  auto with real.
2:rewrite <- powerRZ_add.
2:replace ((pred t + (s + 1 + Fexp x)))%Z with (t + s + Fexp x)%Z; auto with real.
2:rewrite inj_pred; auto with zarith; unfold Z.pred; ring.
2: apply IZR_neq; omega.
generalize FboundedMbound; intros P.
elim P with radix b' (t-s) (Fexp (Fplus radix p q))%Z (Fnum (Fplus radix p q));
  auto with zarith; clear P.
intros v H'; elim H'; intros ; clear H'.
exists v; split; auto.
rewrite hxExact; unfold FtoRradix; rewrite <- Fplus_correct; auto.
cut ( (Z.abs (Fnum (Fplus radix p q)) < ((Zpower_nat radix (t - s))+1)))%Z; auto with zarith.
apply Zlt_Rlt.
apply Rmult_lt_reg_l with (powerRZ radix (Fexp (Fplus radix p q))); auto with real zarith.
apply powerRZ_lt, IZR_lt; omega.
apply Rle_lt_trans with (Rabs (Fplus radix p q)).
unfold FtoRradix; rewrite <- Fabs_correct; auto.
unfold Fabs, FtoR; simpl; auto with real.
unfold FtoRradix; rewrite Fplus_correct; auto.
fold FtoRradix; rewrite <- hxExact.
replace (FtoRradix hx) with (-(x-hx)+x)%R;[idtac|ring].
apply Rle_lt_trans with (Rabs (-(x-hx))+ Rabs(x))%R;[apply Rabs_triang|idtac].
rewrite Rabs_Ropp.
apply Rle_lt_trans with ((powerRZ radix (s + Fexp x))/2 + Rabs x)%R; auto with real.
apply Rlt_le_trans with ((powerRZ radix (s + Fexp x))/2 + (powerRZ radix (t+Fexp x)))%R.
apply Rplus_lt_compat_l.
unfold FtoRradix; rewrite <- Fabs_correct; auto; unfold FtoR; simpl.
rewrite powerRZ_add; try apply IZR_neq; auto with real zarith; apply Rmult_lt_compat_r; auto with real zarith.
apply powerRZ_lt, IZR_lt; omega.
elim Fx; intros; rewrite <- Zpower_nat_Z_powerRZ; rewrite <- pGivesBound; try apply Rlt_IZR; auto with real zarith.
replace (Fexp (Fplus radix p q)) with (s+ 1+Fexp x)%Z.
2:unfold Fplus; simpl.
2:rewrite Zmin_le2;[auto|apply eqLeep].
rewrite plus_IZR; rewrite Zpower_nat_Z_powerRZ; simpl.
rewrite Rmult_plus_distr_l; rewrite <- powerRZ_add.
2: apply IZR_neq; omega.
rewrite Rplus_comm; apply Rplus_le_compat.
rewrite inj_minus1; auto with real zarith.
apply Rle_powerRZ; try apply Rle_IZR; omega.
unfold Rdiv; apply Rmult_le_compat; auto with real zarith.
apply powerRZ_le, IZR_lt; omega.
apply Rle_powerRZ; try apply Rle_IZR; omega.
apply Rle_trans with (/1)%R; auto with real.
unfold b', Fplus; simpl.
rewrite Zmin_le2;[elim Nq; intros Fq T; elim Fq; auto|apply eqLeep].
Qed.

Hypothesis pDefEven: (EvenClosest b radix t (x*((powerRZ radix s)+1))%R p).
Hypothesis qDefEven: (EvenClosest b radix t (x-p)%R q).
Hypothesis hxDefEven:(EvenClosest b radix t (q+p)%R hx).

Lemma VeltkampEven1: (Even radix)
   ->(exists hx':float, (FtoRradix hx'=hx)
    /\ (EvenClosest b' radix (t-s) x hx')).
intros I.
generalize p'GivesBound; intros J.
cut (powerRZ radix (t - 1 + Fexp x) <= x)%R;[intros xGe|idtac].
2:rewrite powerRZ_add; try apply IZR_neq; auto with real zarith; unfold FtoRradix, FtoR.
2:apply Rmult_le_compat_r; auto with real zarith.
2: apply powerRZ_le, IZR_lt; omega.
2:apply Rmult_le_reg_l with radix; auto with real zarith.
2: apply Rlt_IZR; omega.
2:apply Rle_trans with (powerRZ radix t).
2:unfold Zminus; rewrite powerRZ_add; try apply IZR_neq; auto with real zarith; simpl; right; field; auto with real.
2: apply IZR_neq; omega.
2:rewrite <- Zpower_nat_Z_powerRZ; rewrite <- pGivesBound; rewrite <- mult_IZR; elim Nx; intros.
2:apply Rle_IZR; rewrite Zabs_Zmult in H0; rewrite Z.abs_eq in H0; auto with zarith.
2:rewrite Z.abs_eq in H0; auto with zarith real.
2:apply LeR0Fnum with radix; auto with real.
cut (Rabs (x - hx) <= (powerRZ radix (s + Fexp x))/2)%R;[intros|idtac].
2:case eqEqual; intros L.
2:fold FtoRradix; rewrite hxExact.
2:replace (x-(p+q))%R with ((x-p)-q)%R;[apply Rmult_le_reg_l with (INR 2); auto
  with real zarith|ring].
2:apply Rle_trans with (powerRZ radix (Fexp q)).
2:unfold FtoRradix; apply ClosestExp with b t; auto with zarith.
2:rewrite L; simpl; right; field; auto with real.
2:elim L; auto.
cut (exists v:float, (FtoRradix v=hx)/\(Fcanonic radix b' v) /\
   ((FNeven b' radix (t-s) v) \/ (Fexp v <= s+Fexp x)%Z)).
intros T;elim T; intros v T'; elim T'; intros H0 T''; elim T''; intros H1 L; clear T T' T''.
exists v; split; auto.
cut (Fbounded b' v);[intros Fv|apply FcanonicBound with radix; auto].
cut (Rabs (x - FtoR radix v) <= (powerRZ radix (s + Fexp x))/2)%R;[intros|idtac].
2: fold FtoRradix; rewrite H0; auto with real.
case H2; intros; clear H2.
unfold EvenClosest.
cut (Closest b' radix x v /\
       (forall g : float, Closest b' radix x g -> FtoR radix v = FtoR radix g)).
intros T; elim T; split; auto.
right; intros; apply sym_eq; auto.
apply ImplyClosestStrict2 with (minus t s) (s+Fexp x)%Z; auto with zarith real.
replace (s + Fexp x + (t - s)%nat - 1)%Z with ((t-1)+(Fexp x))%Z;[auto with real|rewrite inj_minus1; auto with zarith].
2:elim Fx; unfold b'; simpl; auto with zarith.
replace (s + Fexp x + (t - s)%nat - 1)%Z with ((t-1)+(Fexp x))%Z;[idtac|rewrite inj_minus1; auto with zarith].
fold FtoRradix; apply Veltkamp_aux_aux; auto with real.
cut (Closest b' radix x v);[intros|idtac].
2: apply ImplyClosest with (minus t s) (s+Fexp x)%Z; auto with zarith real.
2: rewrite inj_minus1; auto with zarith real.
2: replace (s + Fexp x + (t - s) - 1)%Z with (t - 1 + Fexp x)%Z; [auto with real|ring].
2: rewrite inj_minus1; auto with zarith real.
2: replace (s + Fexp x + (t - s) - 1)%Z with (t - 1 + Fexp x)%Z; [auto with real|ring].
2: fold FtoRradix; apply Veltkamp_aux_aux; auto with real.
2: elim Fx; unfold b'; simpl; auto with zarith.
split; auto.
left.
case L; clear L; intros L; auto.
unfold FNeven; rewrite FcanonicFnormalizeEq; auto with zarith.
case (Zle_lt_or_eq _ _ L); intros H4; clear L;unfold Feven.
cut (exists m:Z, (Fnum v=radix*m)%Z);[intros T; elim T; intros m H5|idtac].
rewrite H5; apply EvenMult1; auto.
exists (Fnum p*Zpower_nat radix ((Z.abs_nat (Fexp p-Fexp v-1)))+
        Fnum q*Zpower_nat radix ((Z.abs_nat (Fexp q-Fexp v-1))))%Z.
apply eq_IZR.
rewrite mult_IZR; rewrite plus_IZR; repeat rewrite mult_IZR.
repeat rewrite Zpower_nat_Z_powerRZ.
generalize eqGe; generalize eqLeep; intros.
repeat rewrite <- Zabs_absolu.
repeat rewrite Z.abs_eq; auto with zarith.
unfold Zminus; repeat rewrite powerRZ_add; try apply IZR_neq; auto with real zarith.
repeat rewrite powerRZ_Zopp; try apply IZR_neq; auto with real zarith.
rewrite powerRZ_1.
apply Rmult_eq_reg_l with (powerRZ radix (Fexp v)); auto with real zarith.
apply trans_eq with (FtoRradix v);[unfold FtoRradix, FtoR; ring|idtac].
rewrite H0; rewrite hxExact; unfold FtoRradix, FtoR; field.
split; apply Rgt_not_eq; try apply Rlt_IZR; auto with zarith.
apply powerRZ_lt, IZR_lt; omega.
apply Rgt_not_eq, powerRZ_lt, IZR_lt; omega.
replace (Fnum v) with (Fnum p*Zpower_nat radix ((Z.abs_nat (Fexp p-Fexp v)))+
        Fnum q*Zpower_nat radix ((Z.abs_nat (Fexp q-Fexp v))))%Z.
2:apply eq_IZR.
2:rewrite plus_IZR; repeat rewrite mult_IZR.
2:repeat rewrite Zpower_nat_Z_powerRZ.
2: generalize eqGe; generalize eqLeep; intros.
2:repeat rewrite <- Zabs_absolu.
2:repeat rewrite Z.abs_eq; auto with zarith.
2:unfold Zminus; repeat rewrite powerRZ_add; try apply IZR_neq; auto with real zarith.
2:repeat rewrite powerRZ_Zopp; try apply IZR_neq; auto with real zarith.
2:apply Rmult_eq_reg_l with (powerRZ radix (Fexp v)); auto with real zarith.
2:apply trans_eq with (FtoRradix v);[idtac|unfold FtoRradix, FtoR; ring].
2:rewrite H0; rewrite hxExact; unfold FtoRradix, FtoR; field.
2: apply Rgt_not_eq, powerRZ_lt, IZR_lt; omega.
2: apply Rgt_not_eq, powerRZ_lt, IZR_lt; omega.
cut (exists eps:R, ((eps=1)%R \/ (eps=-1)%R) /\ (FtoRradix x=v+ eps*(powerRZ radix (s + Fexp x))/2)%R).
intros T; elim T; intros eps T'; elim T'; intros Heps1 Heps2; clear T T'.
apply EvenPlus1.
rewrite H4.
cut ((Fexp p=1+s+Fexp x)%Z \/ (Fexp p=s+Fexp x)%Z);[intros T; case T; clear T; intros|idtac].
rewrite H5; ring_simplify (1 + s + Fexp x - (s + Fexp x))%Z.
replace (Zpower_nat radix (Z.abs_nat 1))%Z with radix%Z.
apply EvenMult2; auto.
unfold Zpower_nat; simpl; auto with zarith.
rewrite H5;ring_simplify ( (s + Fexp x - (s + Fexp x)))%Z.
unfold Zpower_nat; simpl;ring_simplify (Fnum p * 1)%Z.
cut (FNeven b radix t p).
unfold FNeven;rewrite FcanonicFnormalizeEq; auto with zarith.
left; auto.
apply ClosestImplyEven_int with (x * (powerRZ radix s + 1))%R
   ((Fnum v)*(powerRZ radix s)+(Fnum v)+eps*(powerRZ radix s)/2+(eps-1)/2)%R; auto with zarith.
left; auto.
apply pPos.
rewrite Heps2; unfold FtoRradix, FtoR; rewrite H4; rewrite H5.
unfold Zminus; repeat rewrite powerRZ_add; try apply IZR_neq; auto with real zarith; simpl.
field; auto with real.
elim I; intros rradix I'.
cut ((powerRZ radix s)/2=(rradix*Zpower_nat radix (pred s))%Z)%R;[intros K|idtac].
case Heps1; intros T; rewrite T.
exists (Fnum v*(Zpower_nat radix s)+Fnum v+rradix*Zpower_nat radix (pred s))%Z.
repeat rewrite plus_IZR; rewrite mult_IZR.
rewrite <- K.
rewrite Zpower_nat_Z_powerRZ; unfold Rdiv; ring.
exists (Fnum v*(Zpower_nat radix s)+Fnum v+-(rradix*Zpower_nat radix (pred s))-1)%Z.
unfold Zminus; repeat rewrite plus_IZR; rewrite mult_IZR; rewrite Ropp_Ropp_IZR.
rewrite <- K.
simpl; rewrite Zpower_nat_Z_powerRZ; unfold Rdiv; field; auto with real.
rewrite mult_IZR; rewrite Zpower_nat_Z_powerRZ; rewrite inj_pred; auto with zarith.
unfold Z.pred, Zminus; rewrite powerRZ_add; try apply IZR_neq; auto with real zarith; simpl.
rewrite I'; rewrite mult_IZR; simpl; field.
apply IZR_neq; auto with real zarith.
generalize eqLeep; generalize epLe; generalize eqLe; generalize eqGe; intros.
cut (s+Fexp x <= Fexp p)%Z; auto with zarith.
intros T; case (Zle_lt_or_eq _ _ T); auto with zarith.
cut ((Fexp q=1+s+Fexp x)%Z \/ (Fexp q=s+Fexp x)%Z);[intros T; case T; clear T; intros|idtac].
rewrite H4; rewrite H5; ring_simplify (1 + s + Fexp x - (s + Fexp x))%Z.
replace (Zpower_nat radix (Z.abs_nat 1))%Z with radix%Z.
apply EvenMult2; auto.
unfold Zpower_nat; simpl; auto with zarith.
rewrite H4; rewrite H5;ring_simplify ( (s + Fexp x - (s + Fexp x)))%Z.
unfold Zpower_nat; simpl;ring_simplify (Fnum q * 1)%Z.
2: generalize eqLeep; generalize epLe; generalize eqLe; generalize eqGe; intros.
2: case (Zle_lt_or_eq _ _ H5); auto with zarith.
cut (FNeven b radix t q).
unfold FNeven;rewrite FcanonicFnormalizeEq; auto with zarith.
left; auto.
replace q with (Fopp (Fopp q)).
apply FNevenFop; auto with zarith.
apply ClosestImplyEven_int with (-(x-p))%R
  ((Fnum p)*(powerRZ radix ((Fexp p)-s-(Fexp x)))-(Fnum v)-(eps+1)/2)%R; auto with zarith.
generalize EvenClosestSymmetric; unfold SymmetricP; intros; auto with zarith.
left; apply FnormalFop; auto.
rewrite Fopp_correct; auto; generalize qNeg; auto with real.
simpl; rewrite H5.
apply trans_eq with ((powerRZ radix (s + Fexp x) * powerRZ radix (Fexp p - s
 - Fexp x))*Fnum p+ (powerRZ radix (s + Fexp x) * ( - Fnum v -
  (eps + 1) / 2 + 1 / 2)))%R;[idtac|ring].
rewrite <- powerRZ_add; try apply IZR_neq; auto with real zarith; ring_simplify (s + Fexp x + (Fexp p - s - Fexp x))%Z.
rewrite Heps2; unfold FtoRradix, FtoR; rewrite H4.
unfold Zminus; repeat rewrite powerRZ_add; try apply IZR_neq; auto with real zarith; simpl.
field; auto with real.
case Heps1; intros T; rewrite T.
exists (Fnum p*(Zpower_nat radix (Z.abs_nat (Fexp p-(s+Fexp x))))-Fnum v-1)%Z.
unfold Zminus; repeat rewrite plus_IZR; rewrite mult_IZR; repeat rewrite Ropp_Ropp_IZR; simpl.
repeat rewrite Zpower_nat_Z_powerRZ; replace (Z_of_nat (Z.abs_nat (Fexp p + - (s + Fexp x)))) with
    (Fexp p + - s + - Fexp x)%Z;[unfold Rdiv; field; auto with real|idtac].
rewrite <- Zabs_absolu; rewrite Z.abs_eq; auto with zarith.
generalize eqLeep; generalize epLe; generalize eqLe; generalize eqGe; intros; auto with zarith.
exists (Fnum p*(Zpower_nat radix (Z.abs_nat (Fexp p-(s+Fexp x))))-Fnum v)%Z.
unfold Zminus; repeat rewrite plus_IZR; rewrite mult_IZR; repeat rewrite Ropp_Ropp_IZR; simpl.
repeat rewrite Zpower_nat_Z_powerRZ; replace (Z_of_nat (Z.abs_nat (Fexp p + - (s + Fexp x)))) with
    (Fexp p + - s + - Fexp x)%Z;[unfold Rdiv; field; auto with real|idtac].
rewrite <- Zabs_absolu; rewrite Z.abs_eq; auto with zarith.
generalize eqLeep; generalize epLe; generalize eqLe; generalize eqGe; intros; auto with zarith.
unfold Fopp; destruct q; simpl; auto with zarith.
ring_simplify (-(-Fnum0))%Z; auto.
fold FtoRradix in H3; case (Rcase_abs (x-v)%R); intros.
rewrite Rabs_left in H3; auto.
exists (-1)%R; split; auto with real.
apply trans_eq with (v + -1 * (powerRZ radix (s + Fexp x) / 2))%R.
rewrite <- H3; ring.
unfold Rdiv; ring.
rewrite Rabs_right in H3; auto.
exists (1)%R; split; auto with real.
apply trans_eq with (v + 1 * (powerRZ radix (s + Fexp x) / 2))%R.
rewrite <- H3; ring.
unfold Rdiv; ring.
cut (exists v : float,
     FtoRradix v = hx /\
     Fbounded b' v /\
     (FNeven b' radix (t - s) v \/ (Fexp v <= s + Fexp x)%Z)).
intros T; elim T; intros v T1; elim T1; intros H1 T2; elim T2; intros H2 H3; clear T T1 T2.
exists (Fnormalize radix b' (t-s) v).
split.
rewrite <- H1; unfold FtoRradix; rewrite FnormalizeCorrect; auto with real zarith.
split.
apply FnormalizeCanonic; auto with zarith.
case H3; intros.
left; unfold FNeven; unfold FNeven in H0.
rewrite FcanonicFnormalizeEq; auto with zarith.
apply FnormalizeCanonic; auto with zarith.
right; apply Z.le_trans with (2:=H0).
apply FcanonicLeastExp with radix b' (t-s); auto with zarith.
rewrite FnormalizeCorrect; auto with real zarith.
apply FnormalizeCanonic; auto with zarith.
cut (exists m:Z, (FtoRradix hx=m*powerRZ radix (s+Fexp x))%R /\
        ((Z.abs m) <= Zpos (vNum b'))%Z ).
intros T; elim T; intros m T'; elim T'; intros; clear T T'.
case (Zle_lt_or_eq _ _ H1); intros H2.
exists (Float m (s+Fexp x)).
split;[rewrite H0; unfold FtoRradix, FtoR; simpl; ring|split].
split; simpl; elim Fx; auto with zarith.
right; simpl; auto with zarith.
exists (Float (nNormMin radix (t-s)) (s+1+Fexp x)).
cut (Fcanonic radix b' (Float (nNormMin radix (t-s)) (s+1+Fexp x))).
2: apply FcanonicNnormMin; elim Fx; unfold b'; simpl; auto with zarith.
intros H3; split.
rewrite H0; unfold FtoRradix, FtoR, nNormMin; simpl.
rewrite Zpower_nat_Z_powerRZ; rewrite inj_pred; auto with zarith.
rewrite Z.abs_eq in H2.
rewrite H2; rewrite J;rewrite Zpower_nat_Z_powerRZ.
repeat rewrite <- powerRZ_add; try apply IZR_neq; auto with real zarith.
replace (Z.pred (t - s)%nat + (s + 1 + Fexp x))%Z with
   ((t - s)%nat + (s + Fexp x))%Z; auto with real zarith; unfold Z.pred; ring.
apply Zle_Rle.
apply Rmult_le_reg_l with (powerRZ radix (s + Fexp x)); auto with real zarith.
apply powerRZ_lt, IZR_lt; omega.
apply Rle_trans with 0%R;[simpl; right; ring|rewrite Rmult_comm].
rewrite <- H0; rewrite hxExact.
apply Rplus_le_reg_l with (-q)%R.
ring_simplify; unfold FtoRradix; rewrite <- Fopp_correct.
generalize ClosestMonotone; unfold MonotoneP; intros.
apply H4 with b (-(x-p))%R p; auto with zarith real.
apply Rplus_lt_reg_r with (x-p)%R.
ring_simplify; auto with real.
apply ClosestOpp; auto.
unfold FtoRradix; apply RoundedModeProjectorIdem with (P:=(Closest b radix)) (b:=b).
apply ClosestRoundedModeP with t; auto with zarith.
elim pDef; auto.
split;[apply FcanonicBound with radix; auto|idtac].
left; unfold FNeven; rewrite FcanonicFnormalizeEq; auto with zarith.
unfold Feven, nNormMin; simpl.
replace (pred (t-s)) with (S (pred (pred (t-s)))); auto with zarith.
apply EvenExp; auto with zarith.
exists (Fnum p*Zpower_nat radix ((Z.abs_nat (Fexp p-s-Fexp x)))+
        Fnum q*Zpower_nat radix ((Z.abs_nat (Fexp q-s-Fexp x))))%Z.
cut (FtoRradix hx =
   ((Fnum p * Zpower_nat radix (Z.abs_nat (Fexp p - s - Fexp x)) +
     Fnum q * Zpower_nat radix (Z.abs_nat (Fexp q - s - Fexp x)))%Z *
    powerRZ radix (s + Fexp x)))%R;[intros H'; split; auto|idtac].
cut (Z.abs
      (Fnum p * Zpower_nat radix (Z.abs_nat (Fexp p - s - Fexp x)) +
       Fnum q * Zpower_nat radix (Z.abs_nat (Fexp q - s - Fexp x))) <
    Zpos (vNum b')+1)%Z; auto with zarith.
apply Zlt_Rlt.
rewrite plus_IZR; simpl (IZR 1).
rewrite <- Rabs_Zabs.
apply Rmult_lt_reg_l with (powerRZ radix (s + Fexp x)); auto with real zarith.
apply powerRZ_lt, Rlt_IZR; omega.
apply Rle_lt_trans with (Rabs ((powerRZ radix (s + Fexp x))*((Fnum p *
      Zpower_nat radix (Z.abs_nat (Fexp p - s - Fexp x)) +
       Fnum q * Zpower_nat radix (Z.abs_nat (Fexp q - s - Fexp x)))%Z)))%R.
rewrite Rabs_mult; rewrite (Rabs_right (powerRZ radix (s + Fexp x))); auto with real.
apply Rle_ge; apply powerRZ_le, Rlt_IZR; omega.
rewrite Rmult_comm; rewrite <- H'.
replace (FtoRradix hx) with (x+(-(x-hx)))%R;[idtac|ring].
apply Rle_lt_trans with (Rabs x+Rabs (-(x-hx)))%R;[apply Rabs_triang|idtac].
rewrite Rabs_Ropp; apply Rlt_le_trans with (powerRZ radix (t+Fexp x)+Rabs (x-hx))%R.
apply Rplus_lt_compat_r; unfold FtoRradix; rewrite <- Fabs_correct; auto.
unfold FtoR, Fabs; simpl; rewrite powerRZ_add; try apply IZR_neq; auto with real zarith.
apply Rmult_lt_compat_r; auto with real zarith.
apply powerRZ_lt, Rlt_IZR; omega.
elim Fx; intros; rewrite <- Zpower_nat_Z_powerRZ; rewrite <- pGivesBound; try apply Rlt_IZR; auto with real zarith.
apply Rle_trans with (powerRZ radix (t + Fexp x)+ powerRZ radix (s + Fexp x) / 2)%R;
   auto with real.
rewrite J; rewrite Zpower_nat_Z_powerRZ; rewrite inj_minus1; auto with zarith.
rewrite Rmult_plus_distr_l.
apply Rplus_le_compat.
rewrite <- powerRZ_add; try apply IZR_neq; auto with real zarith.
replace (s + Fexp x + (t - s))%Z with (t + Fexp x)%Z by ring; auto with real.
unfold Rdiv; apply Rmult_le_compat_l; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
apply Rle_trans with (/1)%R; auto with real.
rewrite plus_IZR; repeat rewrite mult_IZR.
repeat rewrite Zpower_nat_Z_powerRZ.
generalize eqGe; generalize eqLeep; intros.
repeat rewrite <- Zabs_absolu.
repeat rewrite Z.abs_eq; auto with zarith.
rewrite Rmult_plus_distr_r.
repeat rewrite Rmult_assoc.
repeat rewrite <- powerRZ_add; try apply IZR_neq; auto with real zarith.
ring_simplify (Fexp p - s - Fexp x + (s + Fexp x))%Z.
ring_simplify (Fexp q - s - Fexp x + (s + Fexp x))%Z.
rewrite hxExact; unfold FtoRradix, FtoR; ring.
Qed.

Lemma VeltkampEven2: (Odd radix)
   -> (exists hx':float, (FtoRradix hx'=hx) /\ (EvenClosest b' radix (t-s) x hx')).
intros I.
generalize p'GivesBound;intros J.
cut (powerRZ radix (t - 1 + Fexp x) <= x)%R;[intros xGe|idtac].
2:rewrite powerRZ_add; try apply IZR_neq; auto with real zarith; unfold FtoRradix, FtoR.
2:apply Rmult_le_compat_r; auto with real zarith.
2: apply powerRZ_le, Rlt_IZR; omega.
2:apply Rmult_le_reg_l with radix; try apply Rlt_IZR; auto with real zarith.
2:apply Rle_trans with (powerRZ radix t).
2:unfold Zminus; rewrite powerRZ_add; try apply IZR_neq; auto with real zarith; simpl; right; field; auto with real.
2: apply Rgt_not_eq, Rlt_IZR; omega.
2:rewrite <- Zpower_nat_Z_powerRZ; rewrite <- pGivesBound; rewrite <- mult_IZR; elim Nx; intros.
2:apply Rle_IZR; rewrite Zabs_Zmult in H0; rewrite Z.abs_eq in H0; auto with zarith.
2:rewrite Z.abs_eq in H0; auto with zarith real.
2:apply LeR0Fnum with radix; auto with real.
cut (Rabs (x - hx) <= (powerRZ radix (s + Fexp x))/2)%R;[intros|idtac].
2:case eqEqual; intros L.
2:fold FtoRradix; rewrite hxExact.
2:replace (x-(p+q))%R with ((x-p)-q)%R;[apply Rmult_le_reg_l with (INR 2); auto with real zarith|ring].
2:apply Rle_trans with (powerRZ radix (Fexp q)).
2:unfold FtoRradix; apply ClosestExp with b t; auto with zarith.
2:rewrite L; simpl; right; field; auto with real.
2:elim L; auto.
cut (exists v:float, (FtoRradix v=hx)/\(Fcanonic radix b' v)).
intros T; elim T; intros v T'; elim T'; intros; clear T T'.
exists v; split; auto.
cut (Fbounded b' v);[intros Fv|apply FcanonicBound with radix; auto].
cut (Rabs (x - FtoR radix v) <= (powerRZ radix (s + Fexp x))/2)%R;[intros|idtac].
2: fold FtoRradix; rewrite H0; auto with real.
case H2; intros L.
unfold EvenClosest.
cut (Closest b' radix x v /\
       (forall g : float, Closest b' radix x g -> FtoR radix v = FtoR radix g)).
intros T; elim T; split; auto.
right; intros; apply sym_eq; auto.
apply ImplyClosestStrict2 with (minus t s) (s+Fexp x)%Z; auto with zarith real.
replace (s + Fexp x + (t - s)%nat - 1)%Z with ((t-1)+(Fexp x))%Z;[auto with real|rewrite inj_minus1; auto with zarith].
2:elim Fx; unfold b'; simpl; auto with zarith.
replace (s + Fexp x + (t - s)%nat - 1)%Z with ((t-1)+(Fexp x))%Z;[idtac|rewrite inj_minus1; auto with zarith].
fold FtoRradix; apply Veltkamp_aux_aux; auto.
absurd (Even (Zpower_nat radix s)).
apply OddNEven.
elim s.
unfold Zpower_nat; simpl; unfold Odd.
exists 0%Z; ring.
intros n Hrecn.
replace (S n)with (1+n); auto with zarith.
rewrite Zpower_nat_is_exp.
apply OddMult; auto.
unfold Zpower_nat; simpl; ring_simplify (radix*1)%Z; auto.
replace (Zpower_nat radix s) with (2*(Z.abs (Fnum x-
   Fnum v*Zpower_nat radix (Z.abs_nat (Fexp v-Fexp x)))))%Z.
apply EvenMult1; unfold Even; exists 1%Z; auto with zarith.
apply eq_IZR.
rewrite mult_IZR; rewrite <- Rabs_Zabs.
unfold Zminus; rewrite plus_IZR; rewrite Ropp_Ropp_IZR.
rewrite mult_IZR; repeat rewrite Zpower_nat_Z_powerRZ; simpl.
apply Rmult_eq_reg_l with (powerRZ radix (Fexp x)); auto with zarith real.
rewrite <- powerRZ_add; try apply IZR_neq; auto with real zarith.
apply Rmult_eq_reg_l with (/2)%R; auto with real.
apply trans_eq with (powerRZ radix (s + Fexp x) / 2)%R.
2: unfold Rdiv; rewrite Zplus_comm; ring.
2: apply Rgt_not_eq; apply powerRZ_lt, Rlt_IZR; omega.
rewrite <- L.
apply trans_eq with ((powerRZ radix (Fexp x) *
     (Rabs (Fnum x +- (Fnum v * powerRZ radix (Z.abs_nat (Fexp v +- Fexp x)))))))%R.
field; auto with real.
rewrite <- (Rabs_right (powerRZ radix (Fexp x)));[idtac|apply Rle_ge; auto with real zarith].
rewrite <- Rabs_mult.
replace (x - FtoR radix v)%R with (powerRZ radix (Fexp x) * (Fnum x +
   -(Fnum v * powerRZ radix (Z.abs_nat (Fexp v +- Fexp x)))))%R; auto with real.
unfold FtoRradix, FtoR; rewrite Rmult_plus_distr_l.
rewrite <- Zabs_absolu; rewrite Z.abs_eq.
rewrite powerRZ_add; try apply IZR_neq; auto with real zarith.
rewrite powerRZ_Zopp; try apply IZR_neq; auto with real zarith.
field.
apply Rgt_not_eq; apply powerRZ_lt, Rlt_IZR; omega.
apply Zplus_le_reg_l with (Fexp x).
ring_simplify.
apply Z.le_trans with (Fexp (Float (nNormMin radix (t-s)) (Fexp x)));
   [simpl; auto with zarith|idtac].
apply Fcanonic_Rle_Zle with radix b' (t-s); auto with zarith.
apply FcanonicNnormMin; auto with zarith.
unfold b'; simpl; elim Fx; auto.
cut (powerRZ radix (t - 1 + Fexp x) <= v)%R;[intros H3|idtac].
2: apply Veltkamp_aux_aux; auto.
fold (FtoRradix v);unfold FtoR, nNormMin; simpl.
rewrite Zpower_nat_Z_powerRZ; rewrite <- powerRZ_add.
rewrite Rabs_right;[idtac|apply Rle_ge; auto with real zarith].
rewrite Rabs_right;[idtac|apply Rle_ge;
   apply Rle_trans with (2:=H3); auto with real zarith].
apply Rle_trans with (2:=H3); apply Rle_powerRZ; try apply Rle_IZR; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
apply powerRZ_le, Rlt_IZR; omega.
apply IZR_neq; omega.
apply powerRZ_le, Rlt_IZR; omega.
cut (exists c:float, (FtoRradix c=hx) /\ (Fbounded b' c)).
intros T; elim T; intros c H'; elim H'; intros.
exists (Fnormalize radix b' (t-s) c); split.
unfold FtoRradix; rewrite FnormalizeCorrect; auto with real zarith.
apply FnormalizeCanonic; auto with zarith.
case eqEqual; intros L.
generalize FboundedMbound; intros P.
elim P with radix b' (t-s) (s+Fexp x)%Z (Fnum (Fplus radix p q)); auto with zarith; clear P.
intros v H'; elim H'; intros ; clear H'.
exists v; split; auto.
rewrite hxExact; unfold FtoRradix; rewrite <- Fplus_correct; auto.
rewrite H1; unfold FtoR; replace (s+ Fexp x)%Z with (Fexp (Fplus radix p q)); auto with real.
unfold Fplus; simpl.
rewrite Zmin_le2;[auto|apply eqLeep].
2: elim Fx; unfold b'; simpl; auto with zarith.
cut ( (Z.abs (Fnum (Fplus radix p q)) < ((Zpower_nat radix (t - s))+1)))%Z; auto with zarith.
apply Zlt_Rlt.
apply Rmult_lt_reg_l with (powerRZ radix (Fexp (Fplus radix p q))); auto with real zarith.
apply powerRZ_lt, Rlt_IZR; omega.
apply Rle_lt_trans with (Rabs (Fplus radix p q)).
unfold FtoRradix; rewrite <- Fabs_correct; auto.
unfold Fabs, FtoR; simpl; auto with real.
unfold FtoRradix; rewrite Fplus_correct; auto.
fold FtoRradix; rewrite <- hxExact.
replace (FtoRradix hx) with (-(x-hx)+x)%R;[idtac|ring].
apply Rle_lt_trans with (Rabs (-(x-hx))+ Rabs(x))%R;[apply Rabs_triang|idtac].
rewrite Rabs_Ropp.
apply Rle_lt_trans with ((powerRZ radix (s + Fexp x))/2 + Rabs x)%R; auto with real.
apply Rlt_le_trans with ((powerRZ radix (s + Fexp x))/2 + (powerRZ radix (t+Fexp x)))%R.
apply Rplus_lt_compat_l.
unfold FtoRradix; rewrite <- Fabs_correct; auto; unfold FtoR; simpl.
rewrite powerRZ_add; try apply IZR_neq; auto with real zarith; apply Rmult_lt_compat_r; auto with real zarith.
apply powerRZ_lt, Rlt_IZR; omega.
elim Fx; intros; rewrite <- Zpower_nat_Z_powerRZ; rewrite <- pGivesBound; auto with real zarith.
apply IZR_lt; easy.
replace (Fexp (Fplus radix p q)) with (s+ Fexp x)%Z.
2:unfold Fplus; simpl.
2:rewrite Zmin_le2;[auto|apply eqLeep].
rewrite plus_IZR; rewrite Zpower_nat_Z_powerRZ; simpl.
rewrite Rmult_plus_distr_l; rewrite <- powerRZ_add.
rewrite Rplus_comm; apply Rplus_le_compat.
rewrite inj_minus1; auto with real zarith.
replace ((s + Fexp x + (t - s)))%Z with (t+Fexp x)%Z; auto with real; ring.
unfold Rdiv; apply Rmult_le_compat_l; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
apply Rle_trans with (/1)%R; auto with real.
apply IZR_neq; omega.
elim L; clear L; intros L1 L2.
cut (Fexp q=s+1+Fexp x)%Z;[intros L3|idtac].
2:cut (q=Float (-(nNormMin radix t)) (s+1+Fexp x));[intros I'; rewrite I'; simpl; auto|idtac].
2:apply FnormalUnique with radix b t; auto with zarith.
2:replace (Float (- nNormMin radix t) (s + 1 + Fexp x)) with
   (Fopp (Float (nNormMin radix t) (s + 1 + Fexp x)));
   [idtac|unfold Fopp; auto with zarith].
2:apply FnormalFop; auto.
2:apply FnormalNnormMin; auto with zarith; elim Fx; auto with zarith.
2:fold FtoRradix; rewrite L1; unfold FtoRradix, FtoR, nNormMin; simpl.
2:rewrite Ropp_Ropp_IZR; rewrite Zpower_nat_Z_powerRZ.
2:apply trans_eq with (-(powerRZ radix (pred t) * powerRZ radix (s + 1 + Fexp x)))%R;
  auto with real.
2:rewrite <- powerRZ_add.
2:replace ((pred t + (s + 1 + Fexp x)))%Z with (t + s + Fexp x)%Z; auto with real.
2:rewrite inj_pred; auto with zarith; unfold Z.pred; ring.
2: apply IZR_neq; omega.
generalize FboundedMbound; intros P.
elim P with radix b' (t-s) (Fexp (Fplus radix p q))%Z (Fnum (Fplus radix p q));
  auto with zarith; clear P.
intros v H'; elim H'; intros ; clear H'.
exists v; split; auto.
rewrite hxExact; unfold FtoRradix; rewrite <- Fplus_correct; auto.
cut ( (Z.abs (Fnum (Fplus radix p q)) < ((Zpower_nat radix (t - s))+1)))%Z; auto with zarith.
apply Zlt_Rlt.
apply Rmult_lt_reg_l with (powerRZ radix (Fexp (Fplus radix p q))); auto with real zarith.
apply powerRZ_lt, Rlt_IZR; omega.
apply Rle_lt_trans with (Rabs (Fplus radix p q)).
unfold FtoRradix; rewrite <- Fabs_correct; auto.
unfold Fabs, FtoR; simpl; auto with real.
unfold FtoRradix; rewrite Fplus_correct; auto.
fold FtoRradix; rewrite <- hxExact.
replace (FtoRradix hx) with (-(x-hx)+x)%R;[idtac|ring].
apply Rle_lt_trans with (Rabs (-(x-hx))+ Rabs(x))%R;[apply Rabs_triang|idtac].
rewrite Rabs_Ropp.
apply Rle_lt_trans with ((powerRZ radix (s + Fexp x))/2 + Rabs x)%R; auto with real.
apply Rlt_le_trans with ((powerRZ radix (s + Fexp x))/2 + (powerRZ radix (t+Fexp x)))%R.
apply Rplus_lt_compat_l.
unfold FtoRradix; rewrite <- Fabs_correct; auto; unfold FtoR; simpl.
rewrite powerRZ_add; try apply IZR_neq; auto with real zarith; apply Rmult_lt_compat_r; auto with real zarith.
apply powerRZ_lt, Rlt_IZR; omega.
elim Fx; intros; rewrite <- Zpower_nat_Z_powerRZ; rewrite <- pGivesBound; auto with real zarith.
now apply IZR_lt.
replace (Fexp (Fplus radix p q)) with (s+ 1+Fexp x)%Z.
2:unfold Fplus; simpl.
2:rewrite Zmin_le2;[auto|apply eqLeep].
rewrite plus_IZR; rewrite Zpower_nat_Z_powerRZ; simpl.
rewrite Rmult_plus_distr_l; rewrite <- powerRZ_add; try apply IZR_neq; auto with real zarith.
rewrite Rplus_comm; apply Rplus_le_compat.
apply Rle_powerRZ; try apply Rle_IZR; auto with zarith.
unfold Rdiv; apply Rmult_le_compat; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
apply Rle_powerRZ; try apply Rle_IZR; auto with zarith.
apply Rle_trans with (/1)%R; auto with real.
unfold b', Fplus; simpl.
rewrite Zmin_le2;[elim Nq; intros Fq T; elim Fq; auto|apply eqLeep].
Qed.

End Velt.
Section VeltN.

Variable radix : Z.
Variable b : Fbound.
Variables s t:nat.

Let b' := Bound
    (P_of_succ_nat (pred (Z.abs_nat (Zpower_nat radix (minus t s)))))
    (dExp b).

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.
Hypothesis radixMoreThanOne : (1 < radix)%Z.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ radixMoreThanOne).

Hypothesis pGivesBound: Zpos (vNum b)=(Zpower_nat radix t).
Hypothesis SLe: (2 <= s)%nat.
Hypothesis SGe: (s <= t-2)%nat.

Lemma Veltkamp_pos: forall x p q hx:float,
     Fnormal radix b x -> Fcanonic radix b p -> Fcanonic radix b q
  -> (0 < x)%R
  -> (Closest b radix (x*((powerRZ radix s)+1))%R p)
  -> (Closest b radix (x-p)%R q)
  -> (Closest b radix (q+p)%R hx)
  -> (Rabs (x-hx) <= (powerRZ radix (s+Fexp x)) /2)%R /\
     (exists hx':float, (FtoRradix hx'=hx) /\ (Closest b' radix x hx')
       /\ (s+Fexp x <= Fexp hx')%Z).
intros x p q hx Nx Cp Cq; intros.
unfold FtoRradix, b'; apply Veltkamp_aux with p q; auto.
elim Nx; auto.
case Cp; auto; intros T.
absurd (p < (firstNormalPos radix b t))%R.
apply Rle_not_lt; generalize ClosestMonotone; unfold MonotoneP; intros H3.
unfold FtoRradix; apply H3 with b (firstNormalPos radix b t)
   (x * (powerRZ radix s + 1))%R; auto.
apply Rle_lt_trans with x.
unfold FtoRradix; apply FnormalLtFirstNormalPos; auto with zarith real.
apply Rle_lt_trans with (x*1)%R; auto with real.
apply Rmult_lt_compat_l; auto with real zarith.
apply Rle_lt_trans with (0+1)%R; auto with real zarith.
apply Rplus_lt_compat_r.
apply powerRZ_lt; try apply Rlt_IZR; auto with zarith.
unfold FtoRradix; apply RoundedModeProjectorIdem with (P:=(Closest b radix)) (b:=b).
apply ClosestRoundedModeP with t; auto with zarith.
generalize firstNormalPosNormal; intros H4.
elim H4 with radix b t; auto with zarith.
unfold FtoRradix; apply FsubnormalLtFirstNormalPos; auto with zarith.
apply pPos with b s t x; auto.
rewrite <- Fopp_Fopp; apply FnormalFop.
cut (Fcanonic radix b (Fopp q));[intros T'|apply FcanonicFopp; auto].
case T'; auto; intros T.
absurd (Fopp q < (firstNormalPos radix b t))%R.
apply Rle_not_lt; generalize ClosestMonotone; unfold MonotoneP; intros H3.
unfold FtoRradix; apply H3 with b (firstNormalPos radix b t)
   (-(x-p))%R; auto.
apply Rle_lt_trans with x.
unfold FtoRradix; apply FnormalLtFirstNormalPos; auto with zarith real.
apply Rplus_lt_reg_r with (FtoRradix x).
apply Rle_lt_trans with ((IZR 2)*x)%R;[right; simpl; ring| idtac].
apply Rle_lt_trans with (radix*x)%R;auto with real zarith.
apply Rmult_le_compat_r; try apply Rle_IZR; auto with real zarith.
apply Rlt_le_trans with (radix*(radix*x))%R.
apply Rle_lt_trans with (1*(radix*x))%R; auto with real zarith.
apply Rmult_lt_compat_r; auto with real zarith.
apply Rmult_lt_0_compat; auto with real zarith.
apply Rlt_IZR; omega.
apply Rlt_IZR; omega.
apply Rle_trans with (FtoRradix p);[idtac|right; ring].
apply Rle_trans with (FtoRradix (Float (Fnum x) (Fexp x+2))).
unfold FtoRradix, FtoR; simpl; rewrite powerRZ_add; simpl.
right; ring.
apply IZR_neq; omega.
unfold FtoRradix; apply H3 with b (Float (Fnum x) (Fexp x + 2))
   (x * (powerRZ radix s + 1))%R; auto.
apply Rle_lt_trans with (x * (powerRZ radix 2))%R.
unfold FtoRradix, FtoR; simpl; rewrite powerRZ_add. auto with real.
apply IZR_neq; omega.
apply Rmult_lt_compat_l; auto with real zarith.
apply Rle_lt_trans with (powerRZ radix s+0)%R; auto with real zarith.
apply Rle_trans with (powerRZ radix s)%R; auto with real zarith.
apply Rle_powerRZ; try apply Rle_IZR; auto with zarith real.
unfold FtoRradix; apply RoundedModeProjectorIdem with (P:=(Closest b radix)) (b:=b).
apply ClosestRoundedModeP with t; auto with zarith.
elim Nx; intros T1 T2; elim T1; intros.
split; simpl; auto with zarith.
unfold FtoRradix; apply RoundedModeProjectorIdem with (P:=(Closest b radix)) (b:=b).
apply ClosestRoundedModeP with t; auto with zarith.
generalize firstNormalPosNormal; intros H4.
elim H4 with radix b t; auto with zarith.
apply ClosestOpp; auto.
unfold FtoRradix; apply FsubnormalLtFirstNormalPos; auto with zarith.
rewrite Fopp_correct; cut (q <= 0)%R; auto with real.
unfold FtoRradix; apply qNeg with b s t p x; auto.
elim Nx; auto.
Qed.

Lemma VeltkampN_aux: forall x p q hx:float,
      Fnormal radix b x -> Fcanonic radix b p -> Fcanonic radix b q
  -> (Closest b radix (x*((powerRZ radix s)+1))%R p)
  -> (Closest b radix (x-p)%R q)
  -> (Closest b radix (q+p)%R hx)
  -> (Rabs (x-hx) <= (powerRZ radix (s+Fexp x)) /2)%R /\
     (exists hx':float, (FtoRradix hx'=hx) /\ (Closest b' radix x hx')
       /\ (s+Fexp x <= Fexp hx')%Z).
intros x p q hx Nx Cp Cq; intros.
case (Rle_or_lt 0%R x); intros H2.
case H2; clear H2; intros H2.
apply Veltkamp_pos with p q; auto.
absurd (is_Fzero x).
apply FnormalNotZero with radix b; auto.
apply is_Fzero_rep2 with radix; auto with real.
elim Veltkamp_pos with (Fopp x) (Fopp p) (Fopp q) (Fopp hx).
intros H3 T; elim T; intros v T'; elim T'; intros H4 T''; elim T''; intros ; clear T T' T''.
split.
unfold FtoRradix in H3; repeat rewrite Fopp_correct in H3.
rewrite <- Rabs_Ropp.
replace (-(x-hx))%R with (-x-(-hx))%R;[unfold FtoRradix; apply Rle_trans with (1:=H3)|ring].
unfold Fopp; auto with real.
exists (Fopp v); split.
unfold FtoRradix; rewrite Fopp_correct; fold FtoRradix; rewrite H4.
unfold FtoRradix; rewrite Fopp_correct; ring.
split.
replace (FtoRradix x) with (-(Fopp x))%R.
apply ClosestOpp; auto.
unfold FtoRradix; rewrite Fopp_correct; ring.
unfold Fopp; unfold Fopp in H6; auto with zarith.
apply FnormalFop; auto.
apply FcanonicFopp; auto.
apply FcanonicFopp; auto.
unfold FtoRradix; rewrite Fopp_correct; auto with real.
replace (Fopp x * (powerRZ radix s + 1))%R with (-(x * (powerRZ radix s + 1)))%R.
apply ClosestOpp; auto.
unfold FtoRradix; rewrite Fopp_correct; ring.
replace (Fopp x - Fopp p)%R with (-(x-p))%R;[apply ClosestOpp; auto|idtac].
unfold FtoRradix; repeat rewrite Fopp_correct; ring.
replace (Fopp q + Fopp p)%R with (-(q+p))%R;[apply ClosestOpp; auto|idtac].
unfold FtoRradix; repeat rewrite Fopp_correct; ring.
Qed.

Lemma VeltkampN: forall x p q hx:float,
      Fnormal radix b x
  -> (Closest b radix (x*((powerRZ radix s)+1))%R p)
  -> (Closest b radix (x-p)%R q)
  -> (Closest b radix (q+p)%R hx)
  -> (Rabs (x-hx) <= (powerRZ radix (s+Fexp x)) /2)%R /\
     (exists hx':float, (FtoRradix hx'=hx) /\ (Closest b' radix x hx')
       /\ (s+Fexp x <= Fexp hx')%Z).
intros.
generalize VeltkampN_aux; intros T.
elim T with x (Fnormalize radix b t p) (Fnormalize radix b t q) hx; auto; clear T.
apply FnormalizeCanonic; auto with zarith; elim H0; auto.
apply FnormalizeCanonic; auto with zarith; elim H1; auto.
apply ClosestCompatible with (1 := H0); auto.
rewrite FnormalizeCorrect; auto with real zarith.
apply FnormalizeBounded; auto with zarith; elim H0; auto.
apply ClosestCompatible with (1 := H1); auto.
unfold FtoRradix; rewrite FnormalizeCorrect; auto with real zarith.
rewrite FnormalizeCorrect; auto with real zarith.
apply FnormalizeBounded; auto with zarith; elim H1; auto.
unfold FtoRradix; repeat rewrite FnormalizeCorrect; auto with real zarith.
Qed.

Lemma VeltkampEven_pos: forall x p q hx:float,
      Fnormal radix b x -> Fcanonic radix b p -> Fcanonic radix b q
  -> (0 < x)%R
  -> (EvenClosest b radix t (x*((powerRZ radix s)+1))%R p)
  -> (EvenClosest b radix t (x-p)%R q)
  -> (EvenClosest b radix t (q+p)%R hx)
  -> (exists hx':float, (FtoRradix hx'=hx) /\ (EvenClosest b' radix (t-s) x hx')).
intros x p q hx Nx Cp Cq; intros.
cut (Fnormal radix b q);[intros Nq|idtac].
cut (Fnormal radix b p);[intros Np|idtac].
case (OddEvenDec radix); intros I.
elim Nx; elim H0; elim H1; elim H2; intros.
unfold FtoRradix, b'; apply VeltkampEven2 with p q; auto with zarith real.
elim Nx; elim H0; elim H1; elim H2; intros.
unfold FtoRradix, b'; apply VeltkampEven1 with p q; auto with zarith real.
case Cp; auto; intros T.
absurd (p < (firstNormalPos radix b t))%R.
apply Rle_not_lt; generalize EvenClosestMonotone; unfold MonotoneP; intros H3.
unfold FtoRradix; apply H3 with b t (firstNormalPos radix b t)
   (x * (powerRZ radix s + 1))%R; auto.
apply Rle_lt_trans with x.
unfold FtoRradix; apply FnormalLtFirstNormalPos; auto with zarith real.
apply Rle_lt_trans with (x*1)%R; auto with real.
apply Rmult_lt_compat_l; auto with real zarith.
apply Rle_lt_trans with (0+1)%R; auto with real zarith.
apply Rplus_lt_compat_r.
apply powerRZ_lt, Rlt_IZR; omega.
unfold FtoRradix; apply RoundedModeProjectorIdem with (P:=(EvenClosest b radix t)) (b:=b).
apply EvenClosestRoundedModeP; auto with zarith.
generalize firstNormalPosNormal; intros H4.
elim H4 with radix b t; auto with zarith.
unfold FtoRradix; apply FsubnormalLtFirstNormalPos; auto with zarith.
apply pPos with b s t x; auto.
elim H0; auto.
rewrite <- Fopp_Fopp; apply FnormalFop.
cut (Fcanonic radix b (Fopp q));[intros T'|apply FcanonicFopp; auto].
case T'; auto; intros T.
absurd (Fopp q < (firstNormalPos radix b t))%R.
apply Rle_not_lt; generalize EvenClosestMonotone; unfold MonotoneP; intros H3.
unfold FtoRradix; apply H3 with b t (firstNormalPos radix b t)
   (-(x-p))%R; auto.
apply Rle_lt_trans with x.
unfold FtoRradix; apply FnormalLtFirstNormalPos; auto with zarith real.
apply Rplus_lt_reg_r with (FtoRradix x).
apply Rle_lt_trans with ((IZR 2)*x)%R;[right; simpl; ring| idtac].
apply Rle_lt_trans with (radix*x)%R;auto with real zarith.
apply Rmult_le_compat_r; auto with real.
apply Rle_IZR; omega.
apply Rlt_le_trans with (radix*(radix*x))%R.
apply Rle_lt_trans with (1*(radix*x))%R; auto with real zarith.
apply Rmult_lt_compat_r; auto with real zarith.
apply Rmult_lt_0_compat; try apply Rlt_IZR; auto with real zarith.
apply Rlt_IZR; omega.
apply Rle_trans with (FtoRradix p);[idtac|right; ring].
apply Rle_trans with (FtoRradix (Float (Fnum x) (Fexp x+2))).
unfold FtoRradix, FtoR; simpl; rewrite powerRZ_add; simpl; auto with real zarith.
right; ring.
apply IZR_neq; omega.
unfold FtoRradix; apply H3 with b t (Float (Fnum x) (Fexp x + 2))
   (x * (powerRZ radix s + 1))%R; auto.
apply Rle_lt_trans with (x * (powerRZ radix 2))%R.
unfold FtoRradix, FtoR; simpl; rewrite powerRZ_add. auto with real zarith.
apply IZR_neq; omega.
apply Rmult_lt_compat_l; auto with real zarith.
apply Rle_lt_trans with (powerRZ radix s+0)%R; auto with real zarith.
apply Rle_trans with (powerRZ radix s)%R; auto with real zarith.
apply Rle_powerRZ; try apply Rle_IZR; auto with zarith real.
unfold FtoRradix; apply RoundedModeProjectorIdem with (P:=(EvenClosest b radix t)) (b:=b).
apply EvenClosestRoundedModeP; auto with zarith.
elim Nx; intros T1 T2; elim T1; intros.
split; simpl; auto with zarith.
unfold FtoRradix; apply RoundedModeProjectorIdem with (P:=(EvenClosest b radix t)) (b:=b).
apply EvenClosestRoundedModeP; auto with zarith.
generalize firstNormalPosNormal; intros H4.
elim H4 with radix b t; auto with zarith.
generalize EvenClosestSymmetric; unfold SymmetricP; intros H4.
apply H4; auto with zarith.
unfold FtoRradix; apply FsubnormalLtFirstNormalPos; auto with zarith.
rewrite Fopp_correct; cut (q <= 0)%R; auto with real.
unfold FtoRradix; apply qNeg with b s t p x; auto.
elim Nx; auto.
elim H0; auto.
elim H1; auto.
Qed.

Lemma VeltkampEvenN_aux: forall x p q hx:float,
      Fnormal radix b x -> Fcanonic radix b p -> Fcanonic radix b q
  -> (EvenClosest b radix t (x*((powerRZ radix s)+1))%R p)
  -> (EvenClosest b radix t (x-p)%R q)
  -> (EvenClosest b radix t (q+p)%R hx)
  -> (exists hx':float, (FtoRradix hx'=hx) /\ (EvenClosest b' radix (t-s) x hx')).
intros x p q hx Nx Cp Cq; intros.
case (Rle_or_lt 0%R x); intros H2.
case H2; clear H2; intros H2.
apply VeltkampEven_pos with p q; auto.
exists (Fzero (-(dExp b'))).
split.
cut (FtoR radix p=(Fzero (-(dExp b))))%R; [intros I1|idtac].
cut (FtoR radix q=(Fzero (-(dExp b))))%R; [intros I2|idtac].
unfold FtoRradix; apply RoundedModeProjectorIdemEq with b t (EvenClosest b radix t); auto with zarith.
apply EvenClosestRoundedModeP; auto with zarith.
unfold b'; simpl; apply FboundedFzero.
replace (FtoR radix (Fzero (- dExp b'))) with (q+p)%R; auto.
unfold FtoRradix; rewrite I1; rewrite I2; unfold FtoRradix.
repeat rewrite FzeroisZero; ring.
apply sym_eq; unfold FtoRradix; apply RoundedModeProjectorIdemEq with b t (EvenClosest b radix t); auto with zarith.
apply EvenClosestRoundedModeP; auto with zarith.
apply FboundedFzero.
replace (FtoR radix (Fzero (- dExp b))) with (x -p)%R; auto.
rewrite <- H2; unfold FtoRradix; rewrite I1; unfold FtoRradix.
repeat rewrite FzeroisZero; ring.
apply sym_eq; unfold FtoRradix; apply RoundedModeProjectorIdemEq with b t (EvenClosest b radix t); auto with zarith.
apply EvenClosestRoundedModeP; auto with zarith.
apply FboundedFzero.
replace (FtoR radix (Fzero (- dExp b))) with (x * (powerRZ radix s + 1))%R; auto.
rewrite <- H2; rewrite FzeroisZero; ring.
rewrite <- H2; rewrite <- FzeroisZero with radix b'.
apply RoundedModeProjectorIdem with (P:=(EvenClosest b' radix (t-s))) (b:=b').
apply EvenClosestRoundedModeP; auto with zarith.
unfold b'; apply p'GivesBound; auto.
apply FboundedFzero.
elim VeltkampEven_pos with (Fopp x) (Fopp p) (Fopp q) (Fopp hx).
intros v T; elim T; intros; clear T.
exists (Fopp v); split.
unfold FtoRradix; rewrite Fopp_correct; fold FtoRradix; rewrite H3.
unfold FtoRradix; rewrite Fopp_correct; ring.
replace (FtoRradix x) with (-(Fopp x))%R.
apply EvenClosestSymmetric; auto with zarith.
unfold FtoRradix; rewrite Fopp_correct; ring.
apply FnormalFop; auto.
apply FcanonicFopp; auto.
apply FcanonicFopp; auto.
unfold FtoRradix; rewrite Fopp_correct; auto with real.
replace (Fopp x * (powerRZ radix s + 1))%R with (-(x * (powerRZ radix s + 1)))%R.
apply EvenClosestSymmetric; auto with zarith.
unfold FtoRradix; rewrite Fopp_correct; ring.
replace (Fopp x - Fopp p)%R with (-(x-p))%R;[apply EvenClosestSymmetric; auto with zarith|idtac].
unfold FtoRradix; repeat rewrite Fopp_correct; ring.
replace (Fopp q + Fopp p)%R with (-(q+p))%R;[apply EvenClosestSymmetric; auto with zarith|idtac].
unfold FtoRradix; repeat rewrite Fopp_correct; ring.
Qed.

Lemma VeltkampEvenN: forall x p q hx:float,
      Fnormal radix b x
  -> (EvenClosest b radix t (x*((powerRZ radix s)+1))%R p)
  -> (EvenClosest b radix t (x-p)%R q)
  -> (EvenClosest b radix t (q+p)%R hx)
  -> (exists hx':float, (FtoRradix hx'=hx) /\ (EvenClosest b' radix (t-s) x hx')).
intros.
generalize VeltkampEvenN_aux; intros T.
elim T with x (Fnormalize radix b t p) (Fnormalize radix b t q) hx; auto; clear T.
intros x' T; elim T; intros; exists x'; auto.
apply FnormalizeCanonic; auto with zarith; elim H0;intros J1 J2; elim J1; auto.
apply FnormalizeCanonic; auto with zarith; elim H1;intros J1 J2; elim J1; auto.
apply EvenClosestCompatible with (4 := H0); auto with zarith.
rewrite FnormalizeCorrect; auto with real zarith.
apply FnormalizeBounded; auto with zarith; elim H0;intros J1 J2; elim J1; auto.
apply EvenClosestCompatible with (4 := H1); auto with zarith.
unfold FtoRradix; rewrite FnormalizeCorrect; auto with real zarith.
rewrite FnormalizeCorrect; auto with real zarith.
apply FnormalizeBounded; auto with zarith; elim H1;intros J1 J2; elim J1; auto.
unfold FtoRradix; repeat rewrite FnormalizeCorrect; auto with real zarith.
Qed.

End VeltN.
Section VeltS.

Variable radix : Z.
Variable b : Fbound.
Variables s t:nat.

Let b' := Bound
    (P_of_succ_nat (pred (Z.abs_nat (Zpower_nat radix (minus t s)))))
    (dExp b).

Definition plusExp (b:Fbound):=
   Bound
     (vNum b)
     (Nplus (dExp b) (Npos (P_of_succ_nat (pred (pred t))))).

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.
Hypothesis radixMoreThanOne : (1 < radix)%Z.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ radixMoreThanOne).

Hypothesis pGivesBound: Zpos (vNum b)=(Zpower_nat radix t).
Hypothesis SLe: (2 <= s)%nat.
Hypothesis SGe: (s <= t-2)%nat.

Lemma bimplybplusNorm: forall f:float,
   Fbounded b f -> (FtoRradix f <> 0)%R ->
    (exists g:float, (FtoRradix g=f)%R /\
      Fnormal radix (plusExp b) g).
intros.
exists (Fnormalize radix (plusExp b) t f); split.
unfold FtoRradix; rewrite FnormalizeCorrect; auto with zarith.
cut (Fcanonic radix (plusExp b) (Fnormalize radix (plusExp b) t f)).
intros H1; case H1; auto;intros H2.
absurd (Rabs f < (firstNormalPos radix (plusExp b) t))%R.
apply Rle_not_lt.
unfold firstNormalPos.
apply Rle_trans with (powerRZ radix (-(dExp b))).
unfold FtoRradix, FtoR, plusExp, nNormMin; simpl.
rewrite Zpower_nat_Z_powerRZ; rewrite <- powerRZ_add.
replace (pred t + - ((dExp b + Npos (P_of_succ_nat (pred (pred t)))))%N)%Z with (-(dExp b))%Z; auto with real.
apply trans_eq with (pred t + - (dExp b + (Zpos (P_of_succ_nat (pred (pred t))))))%Z.
replace (Zpos (P_of_succ_nat (pred (pred t)))) with
    (Z_of_nat
       (nat_of_P
          (P_of_succ_nat
             (pred (pred t))))); auto with zarith.
cut (forall (x:N) (y:positive), (x+(Zpos y)=(x +Npos y)%N)%Z).
intros T; rewrite <- T; auto with zarith.
intros;unfold Nplus.
case x; auto with zarith.
apply IZR_neq; omega.
unfold FtoRradix; rewrite <- Fabs_correct; auto.
apply Rle_trans with (1*(powerRZ radix (- dExp b)))%R; auto with real.
unfold FtoR; apply Rmult_le_compat; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
unfold Fabs; simpl.
cut ((Fnum f=0)%Z \/ (1 <= Z.abs (Fnum f))%Z).
intros H3; case H3; auto with real zarith.
intros H4; absurd (FtoRradix f=0)%R; auto with real.
unfold FtoRradix, FtoR; rewrite H4; simpl; ring.
apply Rle_IZR.
case (Zle_or_lt 0%Z (Fnum f)); intros H3.
case (Zle_lt_or_eq _ _ H3); auto with zarith; intros H4.
right; rewrite <- Zabs_Zopp; rewrite Z.abs_eq; auto with zarith.
apply Rle_powerRZ; auto with real zarith.
apply Rle_IZR; omega.
unfold Fabs; simpl; elim H; auto.
unfold FtoRradix; rewrite <- FnormalizeCorrect with radix (plusExp b) t f; auto.
rewrite <- Fabs_correct; auto.
apply FsubnormalLtFirstNormalPos; auto with zarith.
unfold plusExp; simpl; auto.
apply FsubnormFabs; auto.
rewrite Fabs_correct; auto with real.
apply Rabs_pos.
apply FnormalizeCanonic; auto with zarith.
elim H; split; unfold plusExp; simpl; auto with zarith.
Qed.

Lemma Closestbplusb: forall b0:Fbound, forall z:R, forall f:float,
   (Closest (plusExp b0) radix z f) -> (Fbounded b0 f) -> (Closest b0 radix z f).
intros.
split; auto.
intros g Fg; elim H; intros.
apply H2; auto.
elim Fg; intros; split; unfold plusExp; auto.
cut (forall (x:N) (y:positive), (x+(Zpos y)=(x +Npos y)%N)%Z).
intros T; simpl; rewrite <- T; auto with zarith.
intros;unfold Nplus.
case x; auto with zarith.
Qed.

Lemma Closestbbplus: forall b0:Fbound, forall n:nat, forall fext f:float,
    Zpos (vNum b0)=(Zpower_nat radix n) -> (1 < n) ->
   (-dExp b0 <= Fexp fext)%Z ->
    (Closest b0 radix fext f) -> (Closest (plusExp b0) radix fext f).
intros b0 n fext f K1 K2; intros.
elim H0; intros.
split.
elim H1; intros; split; auto.
cut (forall (x:N) (y:positive), (x+(Zpos y)=(x +Npos y)%N)%Z).
intros T; simpl; rewrite <- T; auto with zarith.
intros;unfold Nplus.
case x; auto with zarith.
intros g Hg.
case (Zle_or_lt (-(dExp b0)) (Fexp g)); intros.
apply H2.
elim Hg; split; auto with zarith.
case (Zle_lt_or_eq (-(dExp b0)) (Fexp (Fnormalize radix b0 n f))).
cut (Fbounded b0 (Fnormalize radix b0 n f));[intros T; elim T; auto|idtac].
apply FnormalizeBounded; auto with zarith.
intros; apply Rle_trans with ((Fulp b0 radix n f)/2)%R.
apply Rmult_le_reg_l with (INR 2); auto with zarith real.
apply Rle_trans with (Fulp b0 radix n f);[idtac|simpl; right; field; auto with real].
rewrite <- Rabs_Ropp.
replace (- (FtoR radix f - fext))%R with (fext - FtoR radix f)%R;[idtac|ring].
apply ClosestUlp; auto with zarith.
rewrite <- Rabs_Ropp.
replace (- (FtoR radix g - fext))%R with (fext - FtoR radix g)%R;[idtac|ring].
apply Rle_trans with (Rabs fext - Rabs (FtoR radix g))%R;[idtac|apply Rabs_triang_inv].
apply Rle_trans with ((powerRZ radix (n-1+Fexp (Fnormalize radix b0 n f))
      - powerRZ radix (-1+ Fexp (Fnormalize radix b0 n f)))
      - powerRZ radix (n-1-dExp b0))%R; [idtac|unfold Rminus; apply Rplus_le_compat].
apply Rplus_le_reg_l with (powerRZ radix (-1 + Fexp (Fnormalize radix b0 n f))).
ring_simplify.
apply Rle_trans with (powerRZ radix (Fexp (Fnormalize radix b0 n f))).
unfold Fulp, Rdiv; apply Rle_trans with
  ((/2+/radix)* powerRZ radix (Fexp (Fnormalize radix b0 n f)))%R.
rewrite powerRZ_add; try apply IZR_neq; auto with real zarith; simpl; right; field.
apply IZR_neq; omega.
apply Rle_trans with (1 * powerRZ radix (Fexp (Fnormalize radix b0 n f)))%R;
  [apply Rmult_le_compat_r; auto with real zarith|right; ring].
apply powerRZ_le, Rlt_IZR; omega.
apply Rmult_le_reg_l with (2*radix)%R;
  [apply Rmult_lt_0_compat; auto with real zarith|idtac].
apply Rlt_IZR; omega.
apply Rle_trans with (2+radix)%R;
  [right; field; auto with real zarith | ring_simplify (2*radix*1)%R].
apply IZR_neq; omega.
apply Rle_trans with (radix+radix)%R.
apply Rplus_le_compat_r, Rle_IZR; omega.
right; ring.
apply Rle_trans with (powerRZ radix (n-2+Fexp (Fnormalize radix b0 n f)));
  [apply Rle_powerRZ; try apply Rle_IZR; auto with real zarith|idtac].
apply Rle_trans with (1*(powerRZ radix (n - 2 + Fexp (Fnormalize radix b0 n f))))%R;
   auto with real.
apply Rle_trans with ((radix -1)*(powerRZ radix (n - 2 + Fexp
  (Fnormalize radix b0 n f))))%R;[apply Rmult_le_compat_r; auto with real zarith|idtac].
apply powerRZ_le, Rlt_IZR; omega.
apply Rplus_le_reg_l with 1%R.
ring_simplify (1+(radix-1))%R; apply Rle_trans with (IZR 2); try apply Rle_IZR; auto with real zarith.
apply Rle_trans with ( - powerRZ radix (n - 2+ Fexp (Fnormalize radix b0 n f)) +
    powerRZ radix (n - 1 + Fexp (Fnormalize radix b0 n f)))%R.
right; unfold Zminus; repeat rewrite powerRZ_add; try apply IZR_neq; auto with real zarith.
simpl; field.
apply IZR_neq; omega.
rewrite Rplus_comm; unfold Rminus;apply Rplus_le_compat_l; apply Ropp_le_contravar;
   apply Rle_powerRZ; auto with real zarith.
apply Rle_IZR; omega.
cut (powerRZ radix (n - 1 + Fexp (Fnormalize radix b0 n f)) +
    - powerRZ radix (-1 + Fexp (Fnormalize radix b0 n f))=
   (Float (pPred (vNum b0)) (-1+Fexp (Fnormalize radix b0 n f))))%R.
intros W; rewrite W.
2: unfold FtoRradix, FtoR, pPred.
2: apply trans_eq with (Z.pred (Zpos (vNum b0))*powerRZ radix
  (-1+Fexp (Fnormalize radix b0 n f)))%R;[idtac|simpl; auto with real].
2: unfold Z.pred, Zminus; rewrite plus_IZR.
2: rewrite K1; rewrite Zpower_nat_Z_powerRZ.
2: repeat rewrite powerRZ_add; try apply IZR_neq; auto with real zarith; simpl; field.
2: apply IZR_neq; omega.
case (Rle_or_lt (Float (pPred (vNum b0)) (-1 + Fexp (Fnormalize radix b0 n f)))
    (Rabs fext)); auto with real; intros V.
absurd ( Rabs f <= Float (pPred (vNum b0)) (-1 + Fexp (Fnormalize radix b0 n f)))%R.
apply Rlt_not_le.
apply Rlt_le_trans with (powerRZ radix (n-1+Fexp (Fnormalize radix b0 n f))).
rewrite <- W; apply Rlt_le_trans with (powerRZ radix (n - 1 +
   Fexp (Fnormalize radix b0 n f))+-0)%R; auto with real zarith.
apply Rplus_lt_compat_l, Ropp_lt_contravar.
apply powerRZ_lt, Rlt_IZR; omega.
right; ring.
unfold FtoRradix; rewrite <- FnormalizeCorrect with radix b0 n f; auto with zarith.
rewrite <- Fabs_correct; auto.
rewrite powerRZ_add; unfold FtoRradix, FtoR, Fabs; simpl.
apply Rmult_le_compat_r; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
apply Rmult_le_reg_l with radix; auto with real zarith.
now apply Rlt_IZR.
apply Rle_trans with (powerRZ radix n).
unfold Zminus; rewrite powerRZ_add; try apply IZR_neq; auto with real zarith;
  simpl; right; field ; auto with real.
apply IZR_neq; omega.
2: apply IZR_neq; omega.
cut (Fnormal radix b0 (Fnormalize radix b0 n f));[intros Nf|idtac].
rewrite <- Zpower_nat_Z_powerRZ; rewrite <- K1; rewrite <- mult_IZR;
  elim Nf; intros.
apply Rle_IZR; rewrite Zabs_Zmult in H6; rewrite Z.abs_eq in H6; auto with zarith real.
cut (Fcanonic radix b0 (Fnormalize radix b0 n f));[intros X|apply FnormalizeCanonic; auto with zarith].
case X; auto; intros X'.
elim X'; intros H5 H6; elim H6; intros.
absurd (-dExp b0 < dExp b0)%Z; auto with zarith.
unfold FtoRradix; apply RoundAbsMonotoner with b0 n (Closest b0 radix) fext;
  auto with real zarith.
apply ClosestRoundedModeP with n; auto with zarith.
split.
apply Z.le_lt_trans with (pPred (vNum b0)); auto with zarith.
simpl; rewrite Z.abs_eq; auto with zarith.
apply Zlt_le_weak; apply pPredMoreThanOne with radix n; auto with zarith.
unfold pPred; auto with zarith.
apply Z.le_trans with (Z.pred (Fexp (Fnormalize radix b0 n f))); auto with zarith.
unfold Z.pred; apply Z.le_trans with (-1+Fexp (Fnormalize radix b0 n f))%Z;auto with zarith.
apply Ropp_le_contravar; rewrite <- Fabs_correct; auto.
unfold FtoR, Fabs; simpl.
apply Rle_trans with ((powerRZ radix n)*(powerRZ radix (-1-dExp b0)))%R.
apply Rmult_le_compat; auto with real zarith.
apply Rle_IZR; auto with zarith.
apply powerRZ_le, Rlt_IZR; omega.
elim Hg; intros; rewrite <- Zpower_nat_Z_powerRZ;
   rewrite <- K1; apply Rle_IZR; auto with real zarith.
unfold plusExp in H5; simpl in H5; auto with zarith.
apply Rle_powerRZ; try apply Rle_IZR; auto with real zarith.
unfold Zminus; repeat rewrite powerRZ_add; try apply IZR_neq; auto with real zarith; right; ring.
intros H4.
apply Rle_trans with 0%R.
2: apply Rabs_pos.
right.
rewrite <- FnormalizeCorrect with radix b0 n f; auto with zarith.
unfold FtoRradix; rewrite <- Fminus_correct; auto.
rewrite <- Fabs_correct; auto.
unfold FtoR.
replace (Fnum (Fabs (Fminus radix (Fnormalize radix b0 n f) fext))) with 0%Z;
   [simpl; ring|idtac].
apply sym_eq; apply trans_eq with (Z.abs (Fnum (Fminus radix
   (Fnormalize radix b0 n f) fext)));[simpl; auto with zarith|idtac].
cut ( 0 <= Z.abs (Fnum (Fminus radix (Fnormalize radix b0 n f) fext)))%Z;
  auto with real zarith.
cut (Z.abs (Fnum (Fminus radix (Fnormalize radix b0 n f) fext)) < 1)%Z;
  auto with real zarith.
apply Zlt_Rlt.
apply Rmult_lt_reg_l with (powerRZ radix (-(dExp b0))); auto with real zarith.
apply powerRZ_lt, Rlt_IZR; omega.
apply Rle_lt_trans with (Rabs (f-fext))%R.
unfold FtoRradix; rewrite <- FnormalizeCorrect with radix b0 n f; auto with zarith.
rewrite <- Fminus_correct; auto; rewrite <- Fabs_correct; auto.
unfold FtoR; simpl.
replace (Z.min (Fexp (Fnormalize radix b0 n f)) (Fexp fext)) with (-(dExp b0))%Z;
  [right; ring|idtac].
rewrite Zmin_le1; auto with zarith.
apply Rlt_le_trans with (Fulp b0 radix n f);
   [idtac|unfold Fulp; simpl; rewrite H4; auto with real zarith].
rewrite <- Rabs_Ropp.
replace (- (f - fext))%R with (fext -f)%R;[idtac|ring].
unfold FtoRradix; apply RoundedModeUlp with (Closest b0 radix); auto with zarith real.
apply ClosestRoundedModeP with n; auto with zarith.
Qed.

Lemma EvenClosestbplusb: forall b0:Fbound, forall n:nat, forall fext f:float,
    Zpos (vNum b0)=(Zpower_nat radix n) -> (1 < n) ->
   (-dExp b0 <= Fexp fext)%Z ->
   (EvenClosest (plusExp b0) radix n fext f) -> (Fbounded b0 f)
      -> (EvenClosest b0 radix n fext f).
intros b0 n fext f nGivesB nGe H H0 H1.
elim H0; intros.
cut (Closest b0 radix fext f);[intros|apply Closestbplusb; auto].
split; auto.
cut (Fcanonic radix b0 (Fnormalize radix b0 n f));
   [idtac|apply FnormalizeCanonic; auto with zarith].
intros V; case V; clear V; intros H5.
case H3; intros H6.
left; generalize H6; unfold FNeven.
replace (Fnormalize radix (plusExp b0) n f) with (Fnormalize radix b0 n f); auto.
apply FcanonicUnique with radix (plusExp b0) n; auto with zarith.
elim H5; intros J1 J2; elim J1; intros J3 J4.
unfold plusExp; left; split;[split|idtac];simpl; auto with zarith.
apply FnormalizeCanonic; auto with zarith.
elim H0; intros J1 J2; elim J1; auto.
repeat rewrite FnormalizeCorrect; auto with real.
right; intros; apply H6.
apply Closestbbplus with n; auto.
right; intros;apply sym_eq.
apply RoundedModeProjectorIdemEq with b0 n (Closest b0 radix); auto with zarith.
apply ClosestRoundedModeP with n; auto with zarith.
replace (FtoR radix f) with (FtoR radix fext); auto with real.
apply Rplus_eq_reg_l with (-(FtoR radix f))%R.
ring_simplify (- FtoR radix f + FtoR radix f)%R.
rewrite <- FnormalizeCorrect with radix b0 n f; auto.
apply trans_eq with ((-Fnum (Fnormalize radix b0 n f) +
   (Fnum fext)*Zpower_nat radix (Z.abs_nat (Fexp fext+dExp b0)))%Z
   * (powerRZ radix (-(dExp b0))))%R.
rewrite plus_IZR; rewrite mult_IZR; rewrite Zpower_nat_Z_powerRZ.
rewrite Ropp_Ropp_IZR; unfold FtoR.
replace (Fexp (Fnormalize radix b0 n f)) with (-(dExp b0))%Z.
rewrite Rmult_plus_distr_r; rewrite Rmult_assoc.
rewrite <- powerRZ_add.
replace (Z.abs_nat (Fexp fext + dExp b0)+-dExp b0)%Z with (Fexp fext);[ring|idtac].
rewrite <- Zabs_absolu; rewrite Z.abs_eq; auto with zarith.
apply IZR_neq; omega.
elim H5; intros J1 J2; elim J2; auto.
replace (- Fnum (Fnormalize radix b0 n f) +
     Fnum fext * Zpower_nat radix (Z.abs_nat (Fexp fext + dExp b0)))%Z with 0%Z;
    [simpl; ring|idtac].
cut (Z.abs (- Fnum (Fnormalize radix b0 n f) +
    Fnum fext * Zpower_nat radix (Z.abs_nat (Fexp fext + dExp b0))) = Z.abs 0)%Z;
  auto with zarith.
cut (0 <= (Z.abs
     (- Fnum (Fnormalize radix b0 n f) + Fnum fext *
   Zpower_nat radix (Z.abs_nat (Fexp fext + dExp b0)))))%Z; auto with zarith.
cut ((Z.abs
     (- Fnum (Fnormalize radix b0 n f) + Fnum fext *
   Zpower_nat radix (Z.abs_nat (Fexp fext + dExp b0)))) < 1)%Z; auto with zarith.
apply Zlt_Rlt.
rewrite <- Rabs_Zabs; rewrite plus_IZR; rewrite Ropp_Ropp_IZR.
rewrite mult_IZR; rewrite Zpower_nat_Z_powerRZ.
apply Rmult_lt_reg_l with (Fulp b0 radix n f);
   [unfold Fulp; auto with real zarith|idtac].
apply powerRZ_lt, Rlt_IZR; omega.
pattern (Fulp b0 radix n f) at 1; rewrite <- (Rabs_right (Fulp b0 radix n f)).
2: apply Rle_ge; unfold Fulp; apply powerRZ_le, Rlt_IZR; omega.
rewrite <- Rabs_mult.
replace (Fulp b0 radix n f *
       (- Fnum (Fnormalize radix b0 n f) +
        Fnum fext * powerRZ radix (Z.abs_nat (Fexp fext + dExp b0))))%R
  with (fext -FtoR radix f)%R.
apply Rlt_le_trans with ( Fulp b0 radix n f);[idtac|simpl; right; ring].
apply RoundedModeUlp with (Closest b0 radix); auto with zarith.
apply ClosestRoundedModeP with n; auto with zarith.
rewrite <- FnormalizeCorrect with radix b0 n f; auto.
apply Rplus_eq_reg_l with (FtoR radix (Fnormalize radix b0 n f)).
unfold Fulp, FtoRradix, FtoR;ring_simplify.
apply trans_eq with (Fnum fext *
  (powerRZ radix (Fexp (Fnormalize radix b0 n f))*
   powerRZ radix (Z.abs_nat (Fexp fext + dExp b0))))%R;[idtac|ring].
rewrite <- powerRZ_add; auto with real zarith.
replace (Fexp (Fnormalize radix b0 n f) + Z.abs_nat (Fexp fext + dExp b0))%Z
  with (Fexp fext); auto.
rewrite <- Zabs_absolu; rewrite Z.abs_eq; auto with zarith.
elim H5; intros J1 J2; elim J2; intros; auto with zarith.
apply IZR_neq; omega.
Qed.

Lemma ClosestClosest: forall b0:Fbound, forall n:nat, forall z:R, forall f1 f2:float,
    Zpos (vNum b0)=(Zpower_nat radix n) -> (1 < n) ->
    (Closest b0 radix z f1) -> (Closest b0 radix z f2)
    -> Fnormal radix b0 f2 -> (Fexp f1 <= Fexp f2 -2)%Z
    -> False.
intros.
cut (FtoRradix (Fabs f1) < Fabs f2)%R;[intros|idtac].
absurd (FtoRradix (Fabs f2) = (FNSucc b0 radix n (Fabs f1)))%R.
cut (FNSucc b0 radix n (Fabs f1) < (Fabs f2))%R; auto with real.
unfold FtoRradix; apply FcanonicPosFexpRlt with b0 n; auto with zarith.
apply Rle_trans with (FtoRradix (Fabs f1)).
unfold FtoRradix; rewrite Fabs_correct; auto with real.
apply Rabs_pos.
unfold FtoRradix; apply Rlt_le; apply FNSuccLt; auto with zarith.
rewrite Fabs_correct; auto with real.
apply Rabs_pos.
apply FNSuccCanonic; auto with zarith.
apply absFBounded; elim H1; auto.
apply FcanonicFabs; auto; left; auto.
cut (Fexp (Fnormalize radix b0 n (Fabs f1)) <= Fexp (Fabs f2) - 2)%Z;[intros|idtac].
unfold FNSucc, FSucc.
case (Z_eq_bool (Fnum (Fnormalize radix b0 n (Fabs f1)))); auto with zarith.
apply Z.le_lt_trans with
  (Z.succ (Fexp (Fnormalize radix b0 n (Fabs f1)))); auto with zarith.
case (Z_eq_bool (Fnum (Fnormalize radix b0 n (Fabs f1)))
            (- nNormMin radix n)).
case (Z_eq_bool (Fexp (Fnormalize radix b0 n (Fabs f1))) (- dExp b0)).
apply Z.le_lt_trans with (Fexp (Fnormalize radix b0 n (Fabs f1))); auto with zarith.
apply Z.le_lt_trans with (Z.pred (Fexp (Fnormalize radix b0 n (Fabs f1))));
  auto with zarith.
apply Z.le_lt_trans with (Fexp (Fnormalize radix b0 n (Fabs f1))); auto with zarith.
apply Z.le_trans with (Fexp (Fabs f1));[idtac|unfold Fabs; simpl; auto with zarith].
apply FcanonicLeastExp with radix b0 n; auto with zarith.
rewrite FnormalizeCorrect; auto with real.
apply absFBounded; elim H1; auto.
apply FnormalizeCanonic; auto with zarith.
apply absFBounded; elim H1; auto.
cut (isMin b0 radix (Rabs z) (Fabs f1));[intros K|idtac].
cut (isMax b0 radix (Rabs z) (Fabs f2));[intros K'|idtac].
apply (MaxUniqueP b0 radix (Rabs z)); auto.
apply MinMax; auto with zarith.
case (Req_dec (Rabs z) (Fabs f1)); auto with real.
intros V; absurd (FtoRradix (Fabs f1) = Fabs f2)%R; auto with real.
unfold FtoRradix; apply RoundedModeProjectorIdemEq with b0 n (isMax b0 radix);
  auto with real zarith.
apply MaxRoundedModeP with n; auto with zarith.
apply absFBounded; elim H1; auto.
fold FtoRradix; rewrite <- V; auto.
case (ClosestMinOrMax b0 radix (Rabs z) (Fabs f2)); auto.
apply ClosestFabs with n; auto.
intros H6.
absurd (FtoRradix (Fabs f1)=Fabs f2); auto with real.
apply (MinUniqueP b0 radix (Rabs z)); auto.
case (ClosestMinOrMax b0 radix (Rabs z) (Fabs f1)); auto.
apply ClosestFabs with n; auto.
intros H6.
case (ClosestMinOrMax b0 radix (Rabs z) (Fabs f2)); auto.
apply ClosestFabs with n; auto.
intros H7; elim H6; elim H7; intros.
elim H9; elim H11; intros.
absurd ( (Fabs f2) <= (Fabs f1))%R; auto with real.
apply Rle_trans with (Rabs z); auto with real.
intros H7; absurd (FtoRradix (Fabs f1)=Fabs f2); auto with real.
apply (MaxUniqueP b0 radix (Rabs z)); auto.
unfold FtoRradix; rewrite <- FnormalizeCorrect with radix b0 n (Fabs f1); auto.
apply FcanonicPosFexpRlt with b0 n; auto with zarith.
rewrite FnormalizeCorrect; auto; rewrite Fabs_correct; auto with real.
apply Rabs_pos.
rewrite Fabs_correct; auto with real.
apply Rabs_pos.
apply FnormalizeCanonic; auto with zarith.
apply absFBounded; elim H1; auto.
apply FcanonicFabs; auto; left; auto.
apply Z.le_lt_trans with (Fexp (Fabs f1)); [idtac|unfold Fabs; simpl; auto with zarith].
apply FcanonicLeastExp with radix b0 n; auto with zarith.
rewrite FnormalizeCorrect; auto with real.
apply absFBounded; elim H1; auto.
apply FnormalizeCanonic; auto with zarith.
apply absFBounded; elim H1; auto.
Qed.

Lemma EvenClosestbbplus: forall b0:Fbound, forall n:nat, forall fext f:float,
    Zpos (vNum b0)=(Zpower_nat radix n) -> (1 < n) ->
   (-dExp b0 <= Fexp fext)%Z ->
    (EvenClosest b0 radix n fext f) -> (EvenClosest (plusExp b0) radix n fext f).
intros.
elim H2; intros.
cut (Closest (plusExp b0) radix fext f);
  [intros|apply Closestbbplus with n; auto].
split; auto.
cut (Fbounded b0 f);[intros K|elim H2; intros J1 J2; elim J1; auto].
case (Zle_lt_or_eq (-(dExp b0)) (Fexp (Fnormalize radix b0 n f))).
cut (Fbounded b0 (Fnormalize radix b0 n f));[intros T; elim T; auto|idtac].
apply FnormalizeBounded; auto with zarith.
intros K'.
cut (Fcanonic radix b0 (Fnormalize radix b0 n f));
   [idtac|apply FnormalizeCanonic; auto with zarith].
intros V; case V; clear V; intros H6.
2: elim H6; intros J1 J2; elim J2; intros.
2: absurd (-(dExp b0) < -(dExp b0))%Z; auto with zarith.
case H4; intros H7.
left; generalize H7; unfold FNeven.
replace (Fnormalize radix (plusExp b0) n f) with (Fnormalize radix b0 n f); auto.
apply FcanonicUnique with radix (plusExp b0) n; auto with zarith.
elim H6; intros J1 J2; elim J1; intros J3 J4.
unfold plusExp; left; split;[split|idtac];simpl; auto with zarith.
apply FnormalizeCanonic; auto with zarith.
elim H5; auto.
repeat rewrite FnormalizeCorrect; auto with real.
right; intros.
case (Zle_or_lt (-(dExp b0)) (Fexp q)); intros.
apply H7.
apply Closestbplusb; auto.
elim H8; intros J1 J2; elim J1; intros; split; auto.
absurd (1=1)%R; auto with real;intros Y; clear Y.
apply ClosestClosest with (plusExp b0) n fext q (Fnormalize radix b0 n f); auto.
apply ClosestCompatible with (1 := H5); auto.
rewrite FnormalizeCorrect; auto with real.
elim H6; intros J1 J2; elim J1; intros.
split; unfold plusExp; simpl; auto with zarith.
elim H6; intros J1 J2; elim J1; intros.
split; [split|idtac]; unfold plusExp; simpl; auto with zarith.
apply Z.le_trans with (-(dExp b0)-1)%Z; auto with zarith.
intros H6.
right; intros;apply sym_eq.
apply RoundedModeProjectorIdemEq with (plusExp b0) n (Closest (plusExp b0) radix);
  auto with zarith.
apply ClosestRoundedModeP with n; auto with zarith.
elim H5; auto.
replace (FtoR radix f) with (FtoR radix fext); auto with real.
apply Rplus_eq_reg_l with (-(FtoR radix f))%R.
ring_simplify (- FtoR radix f + FtoR radix f)%R.
rewrite <- FnormalizeCorrect with radix b0 n f; auto.
apply trans_eq with ((-Fnum (Fnormalize radix b0 n f) +
   (Fnum fext)*Zpower_nat radix (Z.abs_nat (Fexp fext+dExp b0)))%Z
   * (powerRZ radix (-(dExp b0))))%R.
rewrite plus_IZR; rewrite mult_IZR; rewrite Zpower_nat_Z_powerRZ.
rewrite Ropp_Ropp_IZR; unfold FtoR.
replace (Fexp (Fnormalize radix b0 n f)) with (-(dExp b0))%Z.
rewrite <- Zabs_absolu; rewrite Z.abs_eq; auto with zarith.
rewrite powerRZ_add.
rewrite powerRZ_Zopp.
field.
apply Rgt_not_eq, powerRZ_lt, IZR_lt; omega.
apply IZR_neq; omega.
apply IZR_neq; omega.
replace (- Fnum (Fnormalize radix b0 n f) +
     Fnum fext * Zpower_nat radix (Z.abs_nat (Fexp fext + dExp b0)))%Z with 0%Z;
    [simpl; ring|idtac].
cut (Z.abs (- Fnum (Fnormalize radix b0 n f) +
    Fnum fext * Zpower_nat radix (Z.abs_nat (Fexp fext + dExp b0))) = Z.abs 0)%Z;
  auto with zarith.
rewrite (Z.abs_eq 0%Z); auto with zarith.
cut (0 <= (Z.abs
     (- Fnum (Fnormalize radix b0 n f) + Fnum fext *
   Zpower_nat radix (Z.abs_nat (Fexp fext + dExp b0)))))%Z; auto with zarith.
cut ((Z.abs
     (- Fnum (Fnormalize radix b0 n f) + Fnum fext *
   Zpower_nat radix (Z.abs_nat (Fexp fext + dExp b0)))) < 1)%Z; auto with zarith.
apply Zlt_Rlt.
rewrite <- Rabs_Zabs; rewrite plus_IZR; rewrite Ropp_Ropp_IZR.
rewrite mult_IZR; rewrite Zpower_nat_Z_powerRZ.
apply Rmult_lt_reg_l with (Fulp b0 radix n f);
   [unfold Fulp; auto with real zarith|idtac].
apply powerRZ_lt, Rlt_IZR; omega.
pattern (Fulp b0 radix n f) at 1; rewrite <- (Rabs_right (Fulp b0 radix n f)).
2: apply Rle_ge; unfold Fulp; auto with real zarith.
2: apply powerRZ_le, Rlt_IZR; omega.
rewrite <- Rabs_mult.
replace (Fulp b0 radix n f *
       (- Fnum (Fnormalize radix b0 n f) +
        Fnum fext * powerRZ radix (Z.abs_nat (Fexp fext + dExp b0))))%R
  with (fext -FtoR radix f)%R.
apply Rlt_le_trans with ( Fulp b0 radix n f);[idtac|simpl; right; ring].
apply RoundedModeUlp with (Closest b0 radix); auto with zarith.
apply ClosestRoundedModeP with n; auto with zarith.
rewrite <- FnormalizeCorrect with radix b0 n f; auto.
apply Rplus_eq_reg_l with (FtoR radix (Fnormalize radix b0 n f)).
unfold Fulp, FtoRradix, FtoR; ring_simplify.
apply trans_eq with (Fnum fext * (powerRZ radix (Fexp (Fnormalize radix b0 n f))
  *(powerRZ radix (Z.abs_nat (Fexp fext + dExp b0)))))%R;[idtac|ring].
rewrite <- powerRZ_add.
replace (Fexp (Fnormalize radix b0 n f) + Z.abs_nat (Fexp fext + dExp b0))%Z
  with (Fexp fext); auto.
rewrite <- Zabs_absolu; rewrite Z.abs_eq; auto with zarith.
apply IZR_neq; omega.
Qed.

Lemma VeltkampS: forall x p q hx:float,
     Fsubnormal radix b x
  -> (Closest b radix (x*((powerRZ radix s)+1))%R p)
  -> (Closest b radix (x-p)%R q)
  -> (Closest b radix (q+p)%R hx)
  -> (Rabs (x-hx) <= (powerRZ radix (s+Fexp x)) /2)%R /\
     (exists hx':float, (FtoRradix hx'=hx) /\ (Closest b' radix x hx')).
intros x p q hx Sx pDef qDef hxDef.
case (Req_dec 0%R x); intros Y.
assert ((exists hx' : float,
      FtoRradix hx' = hx /\ Closest b' radix x hx' /\ (s + Fexp x <= Fexp hx')%Z /\ (FtoRradix hx'=0)%R)).
exists (Fzero (s+Fexp x)).
cut (Fbounded b (Fzero (s+Fexp x)));[intros KK|idtac].
split.
cut (FtoR radix p=(Fzero (-(dExp b))))%R; [intros I1|idtac].
cut (FtoR radix q=(Fzero (-(dExp b))))%R; [intros I2|idtac].
unfold FtoRradix; apply RoundedModeProjectorIdemEq with b t (Closest b radix); auto
 with zarith.
apply ClosestRoundedModeP with t; auto with zarith.
replace (FtoR radix (Fzero (s+Fexp x))) with (q+p)%R; auto.
unfold FtoRradix; rewrite I1; rewrite I2; unfold FtoRradix.
repeat rewrite FzeroisZero.
unfold Fzero, FtoR; simpl; ring.
apply sym_eq; unfold FtoRradix; apply RoundedModeProjectorIdemEq
  with b t (Closest b radix); auto with zarith.
apply ClosestRoundedModeP with t; auto with zarith.
apply FboundedFzero.
replace (FtoR radix (Fzero (- dExp b))) with (x -p)%R; auto.
rewrite <- Y; unfold FtoRradix; rewrite I1; unfold FtoRradix.
repeat rewrite FzeroisZero; ring.
apply sym_eq; unfold FtoRradix; apply RoundedModeProjectorIdemEq
  with b t (Closest b radix); auto with zarith.
apply ClosestRoundedModeP with t; auto with zarith.
apply FboundedFzero.
replace (FtoR radix (Fzero (- dExp b))) with (x * (powerRZ radix s + 1))%R; auto.
rewrite <- Y; rewrite FzeroisZero; ring.
split.
rewrite <- Y; replace 0%R with (FtoR radix (Fzero (s + Fexp x))).
apply RoundedModeProjectorIdem with (P:=(Closest b' radix)) (b:=b').
apply ClosestRoundedModeP with (t-s); auto with zarith.
unfold b'; apply p'GivesBound; auto.
unfold Fzero; split; auto with zarith.
unfold b'; simpl; auto with zarith.
elim Sx; intros T1 T2; elim T1; auto with zarith.
unfold Fzero, FtoR; simpl; ring.
split;[unfold Fzero; simpl; auto with zarith|idtac].
unfold Fzero, FtoRradix, FtoR; simpl; ring.
unfold Fzero; split; auto with zarith.
elim Sx; intros T1 T2; elim T1; simpl; auto with zarith.
elim H; intros f T; elim T; intros H1 T'; elim T'; intros H2 T''; elim T''; intros; clear T T' T''.
split.
rewrite <- Y; rewrite <- H1; rewrite H3; ring_simplify (0-0)%R; rewrite Rabs_R0.
unfold Rdiv; apply Rmult_le_pos; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
exists f; split; auto; split; auto.
lapply (bimplybplusNorm x);[intros T|elim Sx; auto].
lapply T; clear T; [intros T; elim T;
   intros x' T'; elim T'; intros x'Eq Nx'; clear T T'|auto with real].
generalize VeltkampN; intros.
elim H with radix (plusExp b) s t x' p q hx; auto with zarith; clear H.
intros C T; elim T; intros f H; elim H; intros; clear H T.
elim H1; clear H1; intros H1 C'.
cut (Closest (plusExp b') radix x f);[clear H1; intros H1|idtac].
case (Zle_or_lt (-(dExp b)) (Fexp f)); intros H2.
cut (Fbounded b' f);[intros H3|idtac].
split.
rewrite <- x'Eq; unfold FtoRradix; apply Rle_trans with (1:=C).
unfold Rdiv; apply Rmult_le_compat_r; auto with real.
apply Rle_powerRZ; try apply Rle_IZR; auto with real zarith.
apply Zplus_le_compat_l.
apply FcanonicLeastExp with radix (plusExp b) t; auto with zarith.
elim Sx; intros T1 T2; elim T1; intros.
split; unfold plusExp; auto with zarith.
cut (forall (x:N) (y:positive), (x+(Zpos y)=(x +Npos y)%N)%Z).
intros T; simpl; rewrite <- T; auto with zarith.
intros;unfold Nplus.
case x0; auto with zarith.
left; auto.
exists f; split;auto with real.
apply Closestbplusb; auto.
split; [idtac|unfold b'; simpl; auto].
elim H1; intros J1 J2; elim J1; intros; auto with zarith.
split.
rewrite <- x'Eq; unfold FtoRradix; apply Rle_trans with (1:=C).
unfold Rdiv; apply Rmult_le_compat_r; auto with real.
apply Rle_powerRZ; try apply Rle_IZR; auto with real zarith.
apply Zplus_le_compat_l.
apply FcanonicLeastExp with radix (plusExp b) t; auto with zarith.
elim Sx; intros T1 T2; elim T1; intros.
split; unfold plusExp; auto with zarith.
cut (forall (x:N) (y:positive), (x+(Zpos y)=(x +Npos y)%N)%Z).
intros T; simpl; rewrite <- T; auto with zarith.
intros;unfold Nplus.
case x0; auto with zarith.
left; auto.
generalize RoundedModeRep; intros T.
elim T with (plusExp b') radix (t-s) (Closest (plusExp b') radix) x f;
   auto with zarith.
clear T;intros m H3.
cut (Fbounded b' (Float m (Fexp x)));[intros H4|idtac].
exists (Float m (Fexp x)); split.
unfold FtoRradix; rewrite <- H3; rewrite H0; auto with real.
apply Closestbplusb; auto.
apply (ClosestCompatible (plusExp b') radix x x f (Float m (Fexp x)));
   auto with real zarith.
elim H4; intros; split; auto with zarith.
cut (forall (x:N) (y:positive), (x+(Zpos y)=(x +Npos y)%N)%Z).
intros T; simpl; rewrite <- T; auto with zarith.
apply Z.le_trans with (-(dExp b))%Z; auto with zarith.
intros;unfold Nplus.
case x0; auto with zarith.
unfold b'; simpl; auto with zarith.
split.
apply Z.le_lt_trans with (Z.abs (Fnum f)).
apply Z.le_trans with ((Z.abs m)*1)%Z; auto with zarith.
simpl; auto with zarith.
apply Z.le_trans with ((Z.abs m)*(Zpower_nat radix (Z.abs_nat (Fexp x-Fexp f))))%Z.
apply Zmult_le_compat_l; auto with zarith.
apply Zpower_NR1; auto with zarith.
replace (Fnum f) with (m*Zpower_nat radix (Z.abs_nat (Fexp x - Fexp f)))%Z.
rewrite Zabs_Zmult; rewrite (Z.abs_eq (Zpower_nat radix (Z.abs_nat (Fexp x - Fexp f))));
    auto with zarith.
apply Zpower_NR0; omega.
apply eq_IZR; rewrite mult_IZR; rewrite Zpower_nat_Z_powerRZ.
apply Rmult_eq_reg_l with (powerRZ radix (Fexp f)); auto with real zarith.
apply trans_eq with (FtoR radix f);[rewrite H3|unfold FtoR; ring].
unfold FtoR; simpl.
apply trans_eq with (m*(powerRZ radix (Fexp f)*
  powerRZ radix (Z.abs_nat (Fexp x - Fexp f))))%R;[ring|idtac].
rewrite <- powerRZ_add.
replace (Fexp f + Z.abs_nat (Fexp x - Fexp f))%Z with (Fexp x);[ring|idtac].
rewrite <- Zabs_absolu; rewrite Z.abs_eq; auto with zarith.
elim Sx; intros J1 J2; elim J1; intros; auto with zarith.
apply IZR_neq; omega.
apply Rgt_not_eq, powerRZ_lt, Rlt_IZR; omega.
elim H1; intros J1 J2; elim J1; unfold plusExp; simpl; auto with zarith.
elim Sx; intros J1 J2; elim J1; intros ; unfold b'; simpl; auto.
unfold plusExp; simpl.
rewrite <- p'GivesBound with radix b s t; auto with zarith.
apply ClosestRoundedModeP with (t-s); auto with zarith.
unfold plusExp; simpl.
rewrite <- p'GivesBound with radix b s t; auto with zarith.
rewrite <- x'Eq; unfold FtoRradix;auto with zarith.
replace (FtoR radix x' * (powerRZ radix s + 1))%R
  with (FtoRradix (Fplus radix x (Float (Fnum x) (s+Fexp x)%Z))).
apply Closestbbplus with t; auto with zarith.
unfold Fplus; simpl;apply Zmin_Zle.
elim Sx; intros J1 J2; elim J1; auto.
elim Sx; intros J1 J2; elim J1; auto with zarith.
replace (FtoRradix (Fplus radix x (Float (Fnum x) (s + Fexp x))))
  with (x * (powerRZ radix s + 1))%R; auto with real.
unfold FtoRradix; rewrite Fplus_correct; auto.
unfold FtoR; simpl; rewrite powerRZ_add; try apply IZR_neq; auto with real zarith; ring.
fold FtoRradix; rewrite x'Eq; unfold FtoRradix; rewrite Fplus_correct; auto.
unfold FtoR; simpl; rewrite powerRZ_add; try apply IZR_neq; auto with real zarith; ring.
unfold FtoRradix in x'Eq; rewrite x'Eq; rewrite <- Fminus_correct; auto.
apply Closestbbplus with t; auto with zarith.
unfold Fplus; simpl;apply Zmin_Zle.
elim Sx; intros J1 J2; elim J1; auto.
elim pDef; intros J1 J2; elim J1; auto with zarith.
replace (FtoRradix (Fminus radix x p))
  with (x -p)%R; auto with real.
unfold FtoRradix; rewrite Fminus_correct; auto with real.
rewrite <- Fplus_correct; auto.
apply Closestbbplus with t; auto with zarith.
unfold Fplus; simpl;apply Zmin_Zle.
elim qDef; intros J1 J2; elim J1; auto.
elim pDef; intros J1 J2; elim J1; auto with zarith.
replace (FtoRradix (Fplus radix q p))
  with (q +p)%R; auto with real.
unfold FtoRradix; rewrite Fplus_correct; auto with real.
Qed.

Lemma VeltkampEvenS: forall x p q hx:float,
      Fsubnormal radix b x
  -> (EvenClosest b radix t (x*((powerRZ radix s)+1))%R p)
  -> (EvenClosest b radix t (x-p)%R q)
  -> (EvenClosest b radix t (q+p)%R hx)
  -> (exists hx':float, (FtoRradix hx'=hx) /\ (EvenClosest b' radix (t-s) x hx')).
intros x p q hx Sx pDef qDef hxDef.
case (Req_dec 0%R x); intros Y.
exists (Fzero (-(dExp b'))).
split.
cut (FtoR radix p=(Fzero (-(dExp b))))%R; [intros I1|idtac].
cut (FtoR radix q=(Fzero (-(dExp b))))%R; [intros I2|idtac].
unfold FtoRradix; apply RoundedModeProjectorIdemEq with b t (EvenClosest b radix t); auto with zarith.
apply EvenClosestRoundedModeP; auto with zarith.
unfold b'; simpl; apply FboundedFzero.
replace (FtoR radix (Fzero (- dExp b'))) with (q+p)%R; auto.
unfold FtoRradix; rewrite I1; rewrite I2; unfold FtoRradix.
repeat rewrite FzeroisZero; ring.
apply sym_eq; unfold FtoRradix; apply RoundedModeProjectorIdemEq with b t (EvenClosest b radix t); auto with zarith.
apply EvenClosestRoundedModeP; auto with zarith.
apply FboundedFzero.
replace (FtoR radix (Fzero (- dExp b))) with (x -p)%R; auto.
rewrite <- Y; unfold FtoRradix; rewrite I1; unfold FtoRradix.
repeat rewrite FzeroisZero; ring.
apply sym_eq; unfold FtoRradix; apply RoundedModeProjectorIdemEq with b t (EvenClosest b radix t); auto with zarith.
apply EvenClosestRoundedModeP; auto with zarith.
apply FboundedFzero.
replace (FtoR radix (Fzero (- dExp b))) with (x * (powerRZ radix s + 1))%R; auto.
rewrite <- Y; rewrite FzeroisZero; ring.
rewrite <- Y; rewrite <- FzeroisZero with radix b'.
apply RoundedModeProjectorIdem with (P:=(EvenClosest b' radix (t-s))) (b:=b').
apply EvenClosestRoundedModeP; auto with zarith.
unfold b'; apply p'GivesBound; auto.
apply FboundedFzero.
lapply (bimplybplusNorm x);[intros T|elim Sx; auto].
lapply T; clear T; [intros T; elim T;
   intros x' T'; elim T'; intros x'Eq Nx'; clear T T'|auto with real].
generalize VeltkampEvenN; intros.
elim H with radix (plusExp b) s t x' p q hx; auto with zarith; clear H.
intros f H; elim H; intros; clear H.
cut (EvenClosest (plusExp b') radix (t-s) x f);[clear H1; intros H1|idtac].
case (Zle_or_lt (-(dExp b)) (Fexp f)); intros H2.
cut (Fbounded b' f);[intros H3|idtac].
exists f; split;auto with real.
apply EvenClosestbplusb; auto with zarith.
unfold b'; apply p'GivesBound; auto.
unfold b'; simpl; elim Sx; intros J1 J2; elim J1; auto.
split; [idtac|unfold b'; simpl; auto].
elim H1; intros J1 J2; elim J1; intros J3 J4; elim J3; auto with zarith.
generalize RoundedModeRep; intros T.
elim T with (plusExp b') radix (t-s) (Closest (plusExp b') radix) x f;
   auto with zarith.
clear T;intros m H3.
cut (Fbounded b' (Float m (Fexp x)));[intros H4|idtac].
exists (Float m (Fexp x)); split.
unfold FtoRradix; rewrite <- H3; rewrite H0; auto with real.
apply EvenClosestbplusb; auto with zarith.
unfold b'; apply p'GivesBound; auto.
unfold b'; simpl; elim Sx; intros J1 J2; elim J1; auto.
generalize EvenClosestCompatible; unfold CompatibleP; intros C.
apply C with x f; auto with real zarith; clear C.
rewrite <- p'GivesBound with radix b s t; auto; unfold plusExp, b'; simpl; auto.
elim H4; intros; split; auto with zarith.
cut (forall (x:N) (y:positive), (x+(Zpos y)=(x +Npos y)%N)%Z).
intros T; simpl; rewrite <- T; auto with zarith.
apply Z.le_trans with (-(dExp b'))%Z; auto with zarith.
apply Z.le_trans with (-(dExp b') + Zneg (P_of_succ_nat (pred (pred t))))%Z; auto with zarith.
apply Zeq_le; ring_simplify; auto with zarith.
intros;unfold Nplus.
case x0; auto with zarith.
unfold b'; simpl; auto with zarith.
split.
apply Z.le_lt_trans with (Z.abs (Fnum f)).
apply Z.le_trans with ((Z.abs m)*1)%Z; auto with zarith.
simpl; auto with zarith.
apply Z.le_trans with ((Z.abs m)*(Zpower_nat radix (Z.abs_nat (Fexp x-Fexp f))))%Z.
apply Zmult_le_compat_l; auto with zarith.
apply Zpower_NR1; omega.
replace (Fnum f) with (m*Zpower_nat radix (Z.abs_nat (Fexp x - Fexp f)))%Z.
rewrite Zabs_Zmult; rewrite (Z.abs_eq (Zpower_nat radix (Z.abs_nat (Fexp x - Fexp f))));
    auto with zarith.
apply Zpower_NR0; omega.
apply eq_IZR; rewrite mult_IZR; rewrite Zpower_nat_Z_powerRZ.
apply Rmult_eq_reg_l with (powerRZ radix (Fexp f)); auto with real zarith.
apply trans_eq with (FtoR radix f);[rewrite H3|unfold FtoR; ring].
unfold FtoR; simpl.
apply trans_eq with (m*(powerRZ radix (Fexp f)*powerRZ radix (Z.abs_nat (Fexp x - Fexp f))))%R;[ring|idtac].
rewrite <- powerRZ_add.
replace (Fexp f + Z.abs_nat (Fexp x - Fexp f))%Z with (Fexp x);[ring|idtac].
rewrite <- Zabs_absolu; rewrite Z.abs_eq; auto with zarith.
elim Sx; intros J1 J2; elim J1; intros; auto with zarith.
apply IZR_neq; omega.
apply Rgt_not_eq, powerRZ_lt, Rlt_IZR; omega.
elim H1; intros J1 J2; elim J1; intros J3 J4; elim J3;
  unfold plusExp; simpl; auto with zarith.
elim Sx; intros J1 J2; elim J1; intros ; unfold b'; simpl; auto.
rewrite <- p'GivesBound with radix b s t; unfold plusExp, b'; simpl; auto with zarith.
apply ClosestRoundedModeP with (t-s); auto with zarith.
rewrite <- p'GivesBound with radix b s t; unfold plusExp, b'; simpl; auto with zarith.
elim H1; auto.
rewrite <- x'Eq; unfold FtoRradix;auto with zarith.
replace (FtoR radix x' * (powerRZ radix s + 1))%R
  with (FtoRradix (Fplus radix x (Float (Fnum x) (s+Fexp x)%Z))).
apply EvenClosestbbplus; auto with zarith.
unfold Fplus; simpl;apply Zmin_Zle.
elim Sx; intros J1 J2; elim J1; auto.
elim Sx; intros J1 J2; elim J1; auto with zarith.
replace (FtoRradix (Fplus radix x (Float (Fnum x) (s + Fexp x))))
  with (x * (powerRZ radix s + 1))%R; auto with real.
unfold FtoRradix; rewrite Fplus_correct; auto.
unfold FtoR; simpl; rewrite powerRZ_add; try apply IZR_neq; auto with real zarith; ring.
fold FtoRradix; rewrite x'Eq; unfold FtoRradix; rewrite Fplus_correct; auto.
unfold FtoR; simpl; rewrite powerRZ_add; try apply IZR_neq; auto with real zarith; ring.
unfold FtoRradix in x'Eq; rewrite x'Eq; rewrite <- Fminus_correct; auto.
apply EvenClosestbbplus; auto with zarith.
unfold Fplus; simpl;apply Zmin_Zle.
elim Sx; intros J1 J2; elim J1; auto.
elim pDef; intros J1 J2; elim J1; intros J3 J4; elim J3; auto with zarith.
replace (FtoRradix (Fminus radix x p))
  with (x -p)%R; auto with real.
unfold FtoRradix; rewrite Fminus_correct; auto with real.
rewrite <- Fplus_correct; auto.
apply EvenClosestbbplus; auto with zarith.
unfold Fplus; simpl;apply Zmin_Zle.
elim qDef; intros J1 J2; elim J1; intros J3 J4; elim J3; auto.
elim pDef; intros J1 J2; elim J1; intros J3 J4; elim J3; auto with zarith.
replace (FtoRradix (Fplus radix q p))
  with (q +p)%R; auto with real.
unfold FtoRradix; rewrite Fplus_correct; auto with real.
Qed.

End VeltS.
Section VeltUlt.

Variable radix : Z.
Variable b : Fbound.
Variables s t:nat.

Let b' := Bound
    (P_of_succ_nat (pred (Z.abs_nat (Zpower_nat radix (minus t s)))))
    (dExp b).

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.
Hypothesis radixMoreThanOne : (1 < radix)%Z.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ radixMoreThanOne).

Hypothesis pGivesBound: Zpos (vNum b)=(Zpower_nat radix t).
Hypothesis SLe: (2 <= s)%nat.
Hypothesis SGe: (s <= t-2)%nat.

Theorem Veltkamp: forall x p q hx:float,
     (Fbounded b x)
  -> (Closest b radix (x*((powerRZ radix s)+1))%R p)
  -> (Closest b radix (x-p)%R q)
  -> (Closest b radix (q+p)%R hx)
  -> (Rabs (x-hx) <= (powerRZ radix (s+Fexp x)) /2)%R /\
      (exists hx':float, (FtoRradix hx'=hx) /\ (Closest b' radix x hx')
         /\ ((Fnormal radix b x) -> (s+Fexp x <= Fexp hx')%Z)).
intros.
cut (Fcanonic radix b (Fnormalize radix b t x));
  [intros C|apply FnormalizeCanonic; auto with zarith].
case C; clear C; intros.
generalize VeltkampN; intros T.
elim T with radix b s t (Fnormalize radix b t x) p q hx; auto.
intros C TT; elim TT; intros v T'; elim T'; intros ; clear T T' TT.
rewrite FnormalizeCorrect in H5; auto.
rewrite FnormalizeCorrect in C; auto.
split.
unfold FtoRradix; apply Rle_trans with (1:=C).
unfold Rdiv; apply Rmult_le_compat_r; auto with real.
apply Rle_powerRZ; try apply Rle_IZR; auto with real zarith.
apply Zplus_le_compat_l.
apply FcanonicLeastExp with radix b t; auto with zarith.
rewrite FnormalizeCorrect; auto with zarith real.
left; auto.
elim H5; intros.
exists v; split; auto with zarith.
split; auto with zarith.
intros; replace x with (Fnormalize radix b t x); auto with zarith.
rewrite FcanonicFnormalizeEq; auto with zarith.
left; auto.
rewrite FnormalizeCorrect; auto with real.
rewrite FnormalizeCorrect; auto with real.
generalize VeltkampS; intros T.
elim T with radix b s t (Fnormalize radix b t x) p q hx; auto; clear T.
intros C TT; elim TT; intros v T'; elim T'; intros ; clear T' TT.
rewrite FnormalizeCorrect in H5; auto.
rewrite FnormalizeCorrect in C; auto.
split.
unfold FtoRradix; apply Rle_trans with (1:=C).
unfold Rdiv; apply Rmult_le_compat_r; auto with real.
apply Rle_powerRZ; try apply Rle_IZR; auto with real zarith.
apply Zplus_le_compat_l.
apply FcanonicLeastExp with radix b t; auto with zarith.
rewrite FnormalizeCorrect; auto with zarith real.
right; auto.
exists v; split; auto with zarith.
split; auto with zarith.
intros T; absurd (FtoRradix x=(Fnormalize radix b t x))%R.
unfold FtoRradix; apply NormalAndSubNormalNotEq with b t; auto with zarith.
unfold FtoRradix; rewrite FnormalizeCorrect; auto with real zarith.
rewrite FnormalizeCorrect; auto with real.
rewrite FnormalizeCorrect; auto with real.
Qed.

Theorem VeltkampEven: forall x p q hx:float,
     (Fbounded b x)
  -> (EvenClosest b radix t (x*((powerRZ radix s)+1))%R p)
  -> (EvenClosest b radix t (x-p)%R q)
  -> (EvenClosest b radix t (q+p)%R hx)
  -> (exists hx':float, (FtoRradix hx'=hx) /\ (EvenClosest b' radix (t-s) x hx')).
intros.
cut (Fcanonic radix b (Fnormalize radix b t x));
  [intros C|apply FnormalizeCanonic; auto with zarith].
case C; clear C; intros.
generalize VeltkampEvenN; intros T.
elim T with radix b s t (Fnormalize radix b t x) p q hx; auto.
intros v T'; elim T'; intros ; clear T T'.
rewrite FnormalizeCorrect in H5; auto.
exists v; split; auto with zarith.
rewrite FnormalizeCorrect; auto with real.
rewrite FnormalizeCorrect; auto with real.
generalize VeltkampEvenS; intros T.
elim T with radix b s t (Fnormalize radix b t x) p q hx; auto; clear T.
intros v T'; elim T'; intros ; clear T'.
rewrite FnormalizeCorrect in H5; auto.
exists v; split; auto with zarith.
rewrite FnormalizeCorrect; auto with real.
rewrite FnormalizeCorrect; auto with real.
Qed.

End VeltUlt.
Section VeltTail.

Variable radix : Z.
Variable b : Fbound.
Variables s t:nat.

Let b' := Bound
    (P_of_succ_nat (pred (Z.abs_nat (Zpower_nat radix (minus t s)))))
    (dExp b).

Let bt := Bound
    (P_of_succ_nat (pred (Z.abs_nat (Zpower_nat radix s))))
    (dExp b).

Let bt2 := Bound
    (P_of_succ_nat (pred (Z.abs_nat (Zpower_nat radix (minus s 1)))))
    (dExp b).

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.
Hypothesis radixMoreThanOne : (1 < radix)%Z.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ radixMoreThanOne).

Hypothesis pGivesBound: Zpos (vNum b)=(Zpower_nat radix t).
Hypothesis SLe: (2 <= s)%nat.
Hypothesis SGe: (s <= t-2)%nat.

Theorem Veltkamp_tail_aux: forall x p q hx tx:float,
     (Fcanonic radix b x)
  -> (Closest b radix (x*((powerRZ radix s)+1))%R p)
  -> (Closest b radix (x-p)%R q)
  -> (Closest b radix (q+p)%R hx)
  -> (Closest b radix (x-hx)%R tx)
  -> (exists v:float, (FtoRradix v=hx) /\
     (Fexp (Fminus radix x v) = Fexp x) /\
      (Z.abs (Fnum (Fminus radix x v)) <= (powerRZ radix s)/2)%R).
intros.
cut (Zpos (vNum b') = Zpower_nat radix (t - s));[intros I|idtac].
2: unfold b'; apply p'GivesBound; auto with zarith.
generalize Veltkamp; intros W.
elim W with radix b s t x p q hx; auto.
2: apply FcanonicBound with radix; auto.
intros C TT; elim TT; intros v' W'; elim W';
  fold FtoRradix; fold b'; intros W1 T; elim T; intros W2 W3; clear W W' TT T.
cut (exists v:float, (Fcanonic radix b' v) /\ (FtoRradix v=v')).
2: exists (Fnormalize radix b' (t-s) v'); unfold b'; elim W2; intros; split.
2:apply FnormalizeCanonic; auto with zarith.
2: simpl.
2: rewrite <- p'GivesBound with radix b s t; simpl; auto with zarith.
2: rewrite nat_of_P_o_P_of_succ_nat_eq_succ; auto with zarith.
2: unfold FtoRradix; apply FnormalizeCorrect; auto.
intros W; elim W; intros v W'; elim W'; intros; clear W W'.
exists v; split.
rewrite H5; auto.
cut (Rabs (x-v) <= (powerRZ radix (s+Fexp x)) /2)%R;[intros T1|idtac].
cut (Fexp (Fminus radix x v) = Fexp x);[intros T2|idtac].
split; auto.
apply Rmult_le_reg_l with (powerRZ radix (Fexp x)); auto with real zarith.
apply powerRZ_lt, Rlt_IZR; omega.
apply Rle_trans with (Rabs (x-v))%R;[right|idtac].
unfold FtoRradix; rewrite <- Fminus_correct; auto;
   rewrite <- Fabs_correct; auto.
rewrite <- T2; unfold FtoR, Fabs; simpl; ring.
apply Rle_trans with (1:= T1); rewrite powerRZ_add.
unfold Rdiv; right; ring.
apply IZR_neq; omega.
unfold Fminus; simpl.
apply Zmin_le1.
case H; intros.
apply Z.le_trans with (Fexp (Float (nNormMin radix (t-s)) (Fexp x)));
   [simpl; auto with zarith|idtac].
apply Fcanonic_Rle_Zle with radix b' (t-s); auto with zarith.
apply FcanonicNnormMin; auto with zarith.
cut (Fbounded b x); [intros T; elim T; intros; unfold b'; simpl;
  auto with zarith| apply FcanonicBound with radix; auto].
rewrite Rabs_right.
apply RoundAbsMonotonel with b' (t-s) (Closest b' radix) x; auto with zarith.
apply ClosestRoundedModeP with (t-s); auto with zarith.
apply FcanonicBound with radix; auto.
apply FcanonicNnormMin; auto with zarith.
elim H6; intros T2 T3; elim T2; intros; unfold b'; simpl; auto.
apply ClosestCompatible with (1:=W2); auto.
apply FcanonicBound with radix; auto.
unfold FtoRradix; rewrite <- Fabs_correct; auto.
unfold FtoR; simpl.
apply Rmult_le_compat_r; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
elim H6; intros; apply Rle_IZR.
apply Zmult_le_reg_r with radix; auto with zarith.
apply Z.le_trans with (Z.abs (radix * Fnum x))%Z;
   [idtac|rewrite Zabs_Zmult; rewrite Z.abs_eq; auto with zarith].
apply Z.le_trans with (2:=H8).
unfold nNormMin; rewrite pGivesBound.
apply Z.le_trans with (Zpower_nat radix (t-s)); auto with zarith.
pattern (t-s) at 2; replace (t-s) with (pred (t-s)+1); auto with zarith.
rewrite Zpower_nat_is_exp; unfold Zpower_nat; simpl; auto with zarith.
ring_simplify (radix*1)%Z; auto with zarith.
apply Zpower_nat_monotone_le; auto with zarith.
apply Rle_ge; apply LeFnumZERO; simpl; unfold nNormMin; auto with real zarith.
apply Zpower_NR0; auto with zarith.
cut (Fbounded b' v);[intros T; elim T; unfold b'; simpl; intros|
  apply FcanonicBound with radix; auto].
elim H6; auto with zarith.
rewrite H5; rewrite W1;auto with real.
Qed.

Theorem Veltkamp_tail: forall x p q hx tx:float,
     (Fbounded b x)
  -> (Closest b radix (x*((powerRZ radix s)+1))%R p)
  -> (Closest b radix (x-p)%R q)
  -> (Closest b radix (q+p)%R hx)
  -> (Closest b radix (x-hx)%R tx)
  -> (exists tx':float, (FtoRradix tx'=tx) /\
         (hx+tx'=x)%R /\ (Fbounded bt tx') /\
         (Fexp (Fnormalize radix b t x) <= Fexp tx')%Z).
intros.
generalize Veltkamp_tail_aux; intros T.
elim T with (Fnormalize radix b t x) p q hx tx; auto; clear T.
intros v T; elim T; intros H4 T'; elim T'; intros H5 H6; clear T T'.
2: apply FnormalizeCanonic; auto with zarith.
2: unfold FtoRradix; rewrite FnormalizeCorrect; auto with real.
2: unfold FtoRradix; rewrite FnormalizeCorrect; auto with real.
2: unfold FtoRradix; rewrite FnormalizeCorrect; auto with real.
exists (Fminus radix (Fnormalize radix b t x) v).
split.
unfold FtoRradix; apply RoundedModeProjectorIdemEq with b t (Closest b radix);
  auto with zarith.
apply ClosestRoundedModeP with t; auto with zarith.
split.
apply Zlt_Rlt.
apply Rle_lt_trans with (1:=H6); rewrite pGivesBound;
  rewrite Zpower_nat_Z_powerRZ.
apply Rlt_le_trans with (powerRZ radix s*1)%R;
   [unfold Rdiv; apply Rmult_lt_compat_l; auto with real zarith|
    ring_simplify (powerRZ radix s*1)%R; apply Rle_powerRZ; auto with real zarith].
apply powerRZ_lt, Rlt_IZR; omega.
apply Rlt_le_trans with (/1)%R; auto with real.
apply Rle_IZR; omega.
rewrite H5; cut (Fbounded b (Fnormalize radix b t x));
   [intros T; elim T; auto|apply FnormalizeBounded; auto with zarith].
rewrite Fminus_correct; auto; rewrite FnormalizeCorrect; auto with real.
fold FtoRradix; rewrite H4; auto.
split.
unfold FtoRradix; rewrite Fminus_correct; auto.
rewrite FnormalizeCorrect; auto with real.
fold FtoRradix; rewrite H4; ring.
split.
split.
apply Z.lt_le_trans with (Zpower_nat radix s).
apply Zlt_Rlt.
apply Rle_lt_trans with (1:=H6).
rewrite Zpower_nat_Z_powerRZ; apply Rlt_le_trans with (powerRZ radix s*1)%R;
   auto with real.
unfold Rdiv; apply Rmult_lt_compat_l; auto with real zarith.
apply powerRZ_lt, Rlt_IZR; omega.
apply Rlt_le_trans with (/1)%R; auto with real.
apply Zeq_le; apply sym_eq.
unfold bt in |- *; unfold vNum in |- * .
apply
 trans_eq
  with
    (Z_of_nat
       (nat_of_P
          (P_of_succ_nat
             (pred (Z.abs_nat (Zpower_nat radix (s))))))).
unfold Z_of_nat in |- *; rewrite nat_of_P_o_P_of_succ_nat_eq_succ;
 auto with zarith.
rewrite nat_of_P_o_P_of_succ_nat_eq_succ; auto with arith zarith.
rewrite <- S_pred with (Z.abs_nat (Zpower_nat radix (s))) 0; auto with zarith.
rewrite <- Zabs_absolu; rewrite Z.abs_eq; auto with zarith.
apply Zpower_NR0; auto with zarith.
cut ( 0 < Z.abs_nat (Zpower_nat radix s))%Z; auto with zarith.
simpl; rewrite <- Zabs_absolu; rewrite Z.abs_eq; auto with zarith.
apply Zpower_nat_less; omega.
apply Zpower_NR0; auto with zarith.
rewrite H5; unfold bt; simpl.
cut (Fbounded b (Fnormalize radix b t x));
  [intros T; elim T; auto|apply FnormalizeBounded; auto with zarith].
rewrite H5; auto with zarith.
Qed.

Theorem Veltkamp_tail2: forall x p q hx tx:float,
     (radix=2)%Z
  -> (Fbounded b x)
  -> (Closest b radix (x*((powerRZ radix s)+1))%R p)
  -> (Closest b radix (x-p)%R q)
  -> (Closest b radix (q+p)%R hx)
  -> (Closest b radix (x-hx)%R tx)
  -> (exists tx':float, (FtoRradix tx'=tx) /\
         (hx+tx'=x)%R /\ (Fbounded bt2 tx') /\
         (Fexp (Fnormalize radix b t x) <= Fexp tx')%Z).
intros x p q hx tx I; intros.
generalize Veltkamp_tail_aux; intros T.
elim T with (Fnormalize radix b t x) p q hx tx; auto; clear T.
intros v T; elim T; intros H4 T'; elim T'; intros H5 H6; clear T T'.
2: apply FnormalizeCanonic; auto with zarith.
2: unfold FtoRradix; rewrite FnormalizeCorrect; auto with real.
2: unfold FtoRradix; rewrite FnormalizeCorrect; auto with real.
2: unfold FtoRradix; rewrite FnormalizeCorrect; auto with real.
generalize FboundedMbound2; intros T.
elim T with bt2 radix (s-1) (Fexp (Fminus radix (Fnormalize radix b t x) v))
     (Fnum (Fminus radix (Fnormalize radix b t x) v)); auto with zarith.
clear T; intros c T'; elim T'; intros H7 T''; elim T''; intros H8 H9; clear T' T''.
cut (FtoRradix c=x-hx)%R;[intros J|idtac].
exists c; split.
unfold FtoRradix; apply RoundedModeProjectorIdemEq with b t (Closest b radix);
  auto with zarith.
apply ClosestRoundedModeP with t; auto with zarith.
elim H7; intros.
split.
apply Z.lt_le_trans with (1:=H10); rewrite pGivesBound.
unfold bt2; simpl; auto with zarith.
apply Z.le_trans with (Zpower_nat radix (s-1)); auto with zarith.
apply Zeq_le.
apply
 trans_eq
  with
    (Z_of_nat
       (nat_of_P
          (P_of_succ_nat
             (pred (Z.abs_nat (Zpower_nat radix (s-1))))))).
unfold Z_of_nat in |- *; rewrite nat_of_P_o_P_of_succ_nat_eq_succ;
 auto with zarith.
rewrite nat_of_P_o_P_of_succ_nat_eq_succ; auto with arith zarith.
rewrite <- S_pred with (Z.abs_nat (Zpower_nat radix (s-1))) 0; auto with zarith.
rewrite <- Zabs_absolu; rewrite Z.abs_eq; auto with zarith.
apply Zpower_NR0; omega.
cut ( 0 < Z.abs_nat (Zpower_nat radix (s-1)))%Z; auto with zarith.
simpl; rewrite <- Zabs_absolu; rewrite Z.abs_eq; auto with zarith.
apply Zpower_nat_less; omega.
apply Zpower_NR0; omega.
apply Zpower_nat_monotone_le; omega.
generalize H11; unfold bt2; simpl; auto.
fold FtoRradix; rewrite J; auto with real.
split; [rewrite J; ring|split; auto].
rewrite <- H5; auto.
apply trans_eq with (FtoRradix (Fminus radix (Fnormalize radix b t x) v)).
unfold FtoRradix; rewrite H8; unfold FtoR; simpl; ring.
unfold FtoRradix; rewrite Fminus_correct; auto;
  rewrite FnormalizeCorrect; auto; fold FtoRradix; rewrite H4; ring.
unfold bt2; simpl.
apply
 trans_eq
  with
    (Z_of_nat
       (nat_of_P
          (P_of_succ_nat
             (pred (Z.abs_nat (Zpower_nat radix (s-1))))))).
unfold Z_of_nat in |- *; rewrite nat_of_P_o_P_of_succ_nat_eq_succ;
 auto with zarith.
rewrite nat_of_P_o_P_of_succ_nat_eq_succ; auto with arith zarith.
rewrite <- S_pred with (Z.abs_nat (Zpower_nat radix (s-1))) 0; auto with zarith.
rewrite <- Zabs_absolu; rewrite Z.abs_eq; auto with zarith.
apply Zpower_NR0; omega.
cut ( 0 < Z.abs_nat (Zpower_nat radix (s-1)))%Z; auto with zarith.
simpl; rewrite <- Zabs_absolu; rewrite Z.abs_eq; auto with zarith.
apply Zpower_nat_less; omega.
apply Zpower_NR0; omega.
apply Zle_Rle; clear T.
apply Rle_trans with (1:=H6); rewrite Zpower_nat_Z_powerRZ.
rewrite inj_minus1; auto with zarith.
unfold Zminus; rewrite powerRZ_add.
rewrite I; simpl; right; field.
apply IZR_neq; omega.
clear T; rewrite H5; unfold bt; simpl.
cut (Fbounded b (Fnormalize radix b t x));
  [intros T; elim T; auto|apply FnormalizeBounded; auto with zarith].
Qed.

End VeltTail.

Section VeltUtile.
Variable radix : Z.
Variable b : Fbound.
Variables s t:nat.

Let b' := Bound
    (P_of_succ_nat (pred (Z.abs_nat (Zpower_nat radix (minus t s)))))
    (dExp b).

Let bt := Bound
    (P_of_succ_nat (pred (Z.abs_nat (Zpower_nat radix s))))
    (dExp b).

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.
Hypothesis radixMoreThanOne : (1 < radix)%Z.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ radixMoreThanOne).

Hypothesis pGivesBound: Zpos (vNum b)=(Zpower_nat radix t).
Hypothesis SLe: (2 <= s)%nat.
Hypothesis SGe: (s <= t-2)%nat.

Theorem VeltkampU: forall x p q hx tx:float,
     (Fcanonic radix b x)
  -> (Closest b radix (x*((powerRZ radix s)+1))%R p)
  -> (Closest b radix (x-p)%R q)
  -> (Closest b radix (q+p)%R hx)
  -> (Closest b radix (x-hx)%R tx)
  -> (Rabs (x-hx) <= (powerRZ radix (s+Fexp x)) /2)%R /\
     (FtoRradix x=hx+tx)%R /\

     (exists hx':float, (FtoRradix hx'=hx)%R
                     /\ (Fbounded b' hx')
                     /\ ((Fnormal radix b x) -> (s+Fexp x <= Fexp hx')%Z)) /\

     (exists tx':float, (FtoRradix tx'=tx)%R
                     /\ (Fbounded bt tx')
                     /\ (Fexp x <= Fexp tx')%Z).
intros.
generalize Veltkamp; intros T.
elim T with radix b s t x p q hx; auto.
2: apply FcanonicBound with radix; auto.
clear T; intros H4 T; elim T; intros hx' T'; elim T'; intros H5 T''; clear T T'.
elim T''; intros H6 H7; clear T''.
generalize Veltkamp_tail; intros T.
elim T with radix b s t x p q hx tx; auto.
2: apply FcanonicBound with radix; auto.
clear T; intros tx' T'; elim T'; intros H8 T''; clear T'.
elim T''; intros H9 T; elim T; intros H10 H11; clear T T''.
split; auto.
split; auto with real.
unfold FtoRradix; rewrite <- H9; rewrite H8; auto with real.
split.
exists hx'; split; auto.
split; auto.
elim H6; auto with zarith.
exists tx'.
split; auto.
split; auto.
rewrite <- FcanonicFnormalizeEq with radix b t x; auto with zarith.
Qed.

End VeltUtile.

Section GenericDek.
Variable b : Fbound.
Variable radix : Z.
Variable p : nat.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.
Hypothesis radixMoreThanOne : (1 < radix)%Z.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ radixMoreThanOne).

Hypothesis pGivesBound : Zpos (vNum b) = Zpower_nat radix p.
Hypothesis precisionGreaterThanOne : 1 < p.

Theorem BoundedL: forall (r:R) (x:float) (e:Z),
   (e <=Fexp x)%Z -> (-dExp b <= e)%Z -> (FtoRradix x=r)%R ->
   (Rabs r < powerRZ radix (e+p))%R ->
       (exists x':float, (FtoRradix x'=r) /\ (Fbounded b x') /\ Fexp x'=e).
intros.
exists (Float (Fnum x*Zpower_nat radix (Z.abs_nat (Fexp x -e)))%Z e).
split.
rewrite <- H1; unfold FtoRradix, FtoR; simpl.
rewrite mult_IZR; rewrite Zpower_nat_Z_powerRZ.
rewrite Rmult_assoc; rewrite <- powerRZ_add.
replace (Z.abs_nat (Fexp x - e) + e)%Z with (Fexp x); auto with real.
rewrite <- Zabs_absolu; rewrite Z.abs_eq; auto with zarith.
apply IZR_neq; omega.
split;[idtac|simpl; auto].
split; simpl; auto.
apply Zlt_Rlt.
rewrite pGivesBound; rewrite <- Rabs_Zabs; rewrite mult_IZR.
repeat rewrite Zpower_nat_Z_powerRZ.
rewrite <- Zabs_absolu; rewrite Z.abs_eq; auto with zarith.
rewrite Rabs_mult; rewrite (Rabs_right ( powerRZ radix (Fexp x - e))).
2: apply Rle_ge, powerRZ_le, Rlt_IZR; omega.
apply Rmult_lt_reg_l with (powerRZ radix e); auto with real zarith.
apply powerRZ_lt, Rlt_IZR; omega.
rewrite <- powerRZ_add.
apply Rle_lt_trans with (2:=H2); rewrite <- H1.
unfold FtoRradix, FtoR; rewrite Rabs_mult.
rewrite (Rabs_right (powerRZ radix (Fexp x))).
2: apply Rle_ge; auto with real zarith.
right; apply trans_eq with (Rabs (Fnum x)*(powerRZ radix e*powerRZ radix (Fexp x-e)))%R;[ring|idtac].
rewrite <- powerRZ_add.
ring_simplify (e+(Fexp x-e))%Z; auto with real.
apply IZR_neq; omega.
apply powerRZ_le, Rlt_IZR; omega.
apply IZR_neq; omega.
Qed.

Theorem ClosestZero: forall (r:R) (x:float),
  (Closest b radix r x) -> (r=0)%R -> (FtoRradix x=0)%R.
intros.
cut (0 <= FtoRradix x)%R;[intros |idtac].
cut (FtoRradix x <= 0)%R;[intros; auto with real |idtac].
unfold FtoRradix; apply RleRoundedLessR0 with b p (Closest b radix) r; auto with real.
apply ClosestRoundedModeP with p; auto.
unfold FtoRradix; apply RleRoundedR0 with b p (Closest b radix) r; auto with real.
apply ClosestRoundedModeP with p; auto.
Qed.

Theorem Closestbbext: forall bext:Fbound, forall fext f:float,
    (vNum bext=vNum b) -> (dExp b < dExp bext)%Z ->
    (-dExp b <= Fexp fext)%Z ->
    (Closest b radix fext f) -> (Closest bext radix fext f).
intros bext fext f K1 K2; intros.
elim H0; intros.
split.
elim H1; intros; split; auto with zarith.
intros g Hg.
case (Zle_or_lt (-(dExp b)) (Fexp g)); intros.
apply H2.
elim Hg; split; auto with zarith.
case (Zle_lt_or_eq (-(dExp b)) (Fexp (Fnormalize radix b p f))).
cut (Fbounded b (Fnormalize radix b p f));[intros T; elim T; auto|idtac].
apply FnormalizeBounded; auto with zarith.
intros; apply Rle_trans with ((Fulp b radix p f)/2)%R.
apply Rmult_le_reg_l with (INR 2); auto with zarith real.
apply Rle_trans with (Fulp b radix p f);[idtac|simpl; right; field; auto with real].
rewrite <- Rabs_Ropp.
replace (- (FtoR radix f - fext))%R with (fext - FtoR radix f)%R;[idtac|ring].
apply ClosestUlp; auto with zarith.
rewrite <- Rabs_Ropp.
replace (- (FtoR radix g - fext))%R with (fext - FtoR radix g)%R;[idtac|ring].
apply Rle_trans with (Rabs fext -Rabs (FtoR radix g))%R;[idtac|apply Rabs_triang_inv].
apply Rle_trans with ((powerRZ radix (p-1+Fexp (Fnormalize radix b p f))
      - powerRZ radix (-1+ Fexp (Fnormalize radix b p f)))
      - powerRZ radix (p-1-dExp b))%R; [idtac|unfold Rminus; apply Rplus_le_compat].
apply Rplus_le_reg_l with (powerRZ radix (-1 + Fexp (Fnormalize radix b p f))).
ring_simplify ( powerRZ radix (-1 + Fexp (Fnormalize radix b p f)) +
    (powerRZ radix (p - 1 + Fexp (Fnormalize radix b p f)) -
     powerRZ radix (-1 + Fexp (Fnormalize radix b p f)) -
     powerRZ radix (p - 1 - dExp b)))%R.
apply Rle_trans with (powerRZ radix (Fexp (Fnormalize radix b p f))).
unfold Fulp, Rdiv; apply Rle_trans with
  ((/2+/radix)* powerRZ radix (Fexp (Fnormalize radix b p f)))%R.
rewrite powerRZ_add; try apply IZR_neq; auto with real zarith; simpl; right; field.
apply IZR_neq; auto with real zarith.
apply Rle_trans with (1 * powerRZ radix (Fexp (Fnormalize radix b p f)))%R;
  [apply Rmult_le_compat_r; auto with real zarith|right; ring].
apply powerRZ_le, Rlt_IZR; omega.
apply Rmult_le_reg_l with (2*radix)%R;
  [apply Rmult_lt_0_compat; auto with real zarith|idtac].
apply Rlt_IZR; omega.
apply Rle_trans with (2+radix)%R;
  [right; field; auto with real zarith| ring_simplify (2*radix*1)%R].
apply IZR_neq; omega.
apply Rle_trans with (radix+radix)%R;[idtac|right; ring].
apply Rplus_le_compat_r, Rle_IZR; omega.
apply Rle_trans with (powerRZ radix (p-2+Fexp (Fnormalize radix b p f)));
  [apply Rle_powerRZ; try apply Rle_IZR; auto with real zarith|idtac].
apply Rle_trans with (1*(powerRZ radix (p - 2 + Fexp (Fnormalize radix b p f))))%R;
   auto with real.
apply Rle_trans with ((radix -1)*(powerRZ radix (p - 2 + Fexp
  (Fnormalize radix b p f))))%R;[apply Rmult_le_compat_r; auto with real zarith|idtac].
apply powerRZ_le, Rlt_IZR; omega.
apply Rplus_le_reg_l with 1%R.
ring_simplify (1+(radix-1))%R; try rewrite <- plus_IZR; try apply IZR_le; auto with real zarith.
apply Rle_trans with ( - powerRZ radix (p - 2+ Fexp (Fnormalize radix b p f)) +
    powerRZ radix (p - 1 + Fexp (Fnormalize radix b p f)))%R.
right; unfold Zminus; repeat rewrite powerRZ_add; try apply IZR_neq; auto with real zarith.
simpl; field.
apply IZR_neq; omega.
unfold Rminus; rewrite Rplus_comm; apply Rplus_le_compat_l; apply Ropp_le_contravar;
   apply Rle_powerRZ; auto with real zarith.
apply Rle_IZR; omega.
cut (powerRZ radix (p - 1 + Fexp (Fnormalize radix b p f)) +
    - powerRZ radix (-1 + Fexp (Fnormalize radix b p f))=
   (Float (pPred (vNum b)) (-1+Fexp (Fnormalize radix b p f))))%R.
intros W; rewrite W.
2: unfold FtoRradix, FtoR, pPred.
2: apply trans_eq with (Z.pred (Zpos (vNum b))*powerRZ radix
  (-1+Fexp (Fnormalize radix b p f)))%R;[idtac|simpl; auto with real].
2: unfold Z.pred, Zminus; rewrite plus_IZR.
2: rewrite pGivesBound; rewrite Zpower_nat_Z_powerRZ.
2: repeat rewrite powerRZ_add; try apply IZR_neq; auto with real zarith; simpl; field.
2: apply IZR_neq; omega.
case (Rle_or_lt (Float (pPred (vNum b)) (-1 + Fexp (Fnormalize radix b p f)))
    (Rabs fext)); auto with real; intros V.
absurd ( Rabs f <= Float (pPred (vNum b)) (-1 + Fexp (Fnormalize radix b p f)))%R.
apply Rlt_not_le.
apply Rlt_le_trans with (powerRZ radix (p-1+Fexp (Fnormalize radix b p f))).
rewrite <- W; apply Rlt_le_trans with (powerRZ radix (p - 1 +
   Fexp (Fnormalize radix b p f))+-0)%R; auto with real zarith.
apply Rplus_lt_compat_l, Ropp_lt_contravar.
apply powerRZ_lt, Rlt_IZR; omega.
right; ring.
unfold FtoRradix; rewrite <- FnormalizeCorrect with radix b p f; auto with zarith.
rewrite <- Fabs_correct; auto.
rewrite powerRZ_add; unfold FtoRradix, FtoR, Fabs; simpl.
apply Rmult_le_compat_r; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
apply Rmult_le_reg_l with radix; auto with real zarith.
now apply Rlt_IZR.
apply Rle_trans with (powerRZ radix p).
unfold Zminus; rewrite powerRZ_add; try apply IZR_neq; auto with real zarith;
  simpl; right; field ; auto with real.
apply IZR_neq; omega.
cut (Fnormal radix b (Fnormalize radix b p f));[intros Nf|idtac].
rewrite <- Zpower_nat_Z_powerRZ; rewrite <- pGivesBound; rewrite <- mult_IZR;
  elim Nf; intros.
rewrite Zabs_Zmult in H6; rewrite Z.abs_eq in H6; auto with zarith real.
apply Rle_IZR; omega.
cut (Fcanonic radix b (Fnormalize radix b p f));[intros X|apply FnormalizeCanonic; auto with zarith].
case X; auto; intros X'.
elim X'; intros H5 H6; elim H6; intros.
absurd (-dExp b < dExp b)%Z; auto with zarith.
apply IZR_neq; omega.
unfold FtoRradix; apply RoundAbsMonotoner with b p (Closest b radix) fext;
  auto with real zarith.
apply ClosestRoundedModeP with p; auto with zarith.
split.
apply Z.le_lt_trans with (pPred (vNum b)); auto with zarith.
simpl; rewrite Z.abs_eq; auto with zarith.
apply Zlt_le_weak; apply pPredMoreThanOne with radix p; auto with zarith.
unfold pPred; auto with zarith.
apply Z.le_trans with (Z.pred (Fexp (Fnormalize radix b p f))); auto with zarith.
unfold Z.pred; apply Z.le_trans with (-1+Fexp (Fnormalize radix b p f))%Z;auto with zarith.
apply Ropp_le_contravar; rewrite <- Fabs_correct; auto.
unfold FtoR, Fabs; simpl.
apply Rle_trans with ((powerRZ radix p)*(powerRZ radix (-1-dExp b)))%R.
apply Rmult_le_compat; auto with real zarith.
apply Rle_IZR; auto with zarith.
apply powerRZ_le, Rlt_IZR; omega.
elim Hg; intros; rewrite <- Zpower_nat_Z_powerRZ;
   rewrite <- pGivesBound;rewrite <- K1; auto with real zarith.
apply Rle_IZR; omega.
apply Rle_powerRZ; auto with real zarith.
apply Rle_IZR; omega.
unfold Zminus; repeat rewrite powerRZ_add; try apply IZR_neq; auto with real zarith; right; ring.
intros H4.
apply Rle_trans with 0%R; try apply Rabs_pos; right.
rewrite <- FnormalizeCorrect with radix b p f; auto with zarith.
unfold FtoRradix; rewrite <- Fminus_correct; auto.
rewrite <- Fabs_correct; auto.
unfold FtoR.
replace (Fnum (Fabs (Fminus radix (Fnormalize radix b p f) fext))) with 0%Z;
   [simpl; ring|idtac].
apply sym_eq; apply trans_eq with (Z.abs (Fnum (Fminus radix
   (Fnormalize radix b p f) fext)));[simpl; auto with zarith|idtac].
cut ( 0 <= Z.abs (Fnum (Fminus radix (Fnormalize radix b p f) fext)))%Z;
  auto with real zarith.
cut (Z.abs (Fnum (Fminus radix (Fnormalize radix b p f) fext)) < 1)%Z;
  auto with real zarith.
apply Zlt_Rlt.
apply Rmult_lt_reg_l with (powerRZ radix (-(dExp b))); auto with real zarith.
apply powerRZ_lt, Rlt_IZR; omega.
apply Rle_lt_trans with (Rabs (f-fext))%R.
unfold FtoRradix; rewrite <- FnormalizeCorrect with radix b p f; auto with zarith.
rewrite <- Fminus_correct; auto; rewrite <- Fabs_correct; auto.
unfold FtoR; simpl.
replace (Z.min (Fexp (Fnormalize radix b p f)) (Fexp fext)) with (-(dExp b))%Z;
  [right; ring|idtac].
rewrite Zmin_le1; auto with zarith.
apply Rlt_le_trans with (Fulp b radix p f);
   [idtac|unfold Fulp; simpl; rewrite H4; auto with real zarith].
rewrite <- Rabs_Ropp.
replace (- (f - fext))%R with (fext -f)%R;[idtac|ring].
unfold FtoRradix; apply RoundedModeUlp with (Closest b radix); auto with zarith real.
apply ClosestRoundedModeP with p; auto with zarith.
Qed.

Variable b' : Fbound.

Definition Underf_Err (a a' : float) (ra n:R) :=
   (Closest b radix ra a) /\ (Fbounded b' a') /\
   (Rabs (a-a') <= n*powerRZ radix (-(dExp b)))%R /\
   ( ((-dExp b) <= Fexp a')%Z -> (FtoRradix a =a')%R).

Theorem Underf_Err1: forall (a' a:float),
    vNum b=vNum b' -> (dExp b <= dExp b')%Z ->
   (Fbounded b' a') -> (Closest b radix a' a) ->
   (Underf_Err a a' (FtoRradix a') (/2)%R).
intros.
unfold Underf_Err.
split; auto.
split; auto.
case (Zle_or_lt (- dExp b)%Z (Fexp a')); intros.
cut (FtoRradix a'=a);[intros H4|idtac].
rewrite H4; split; auto with real.
ring_simplify (a-a)%R; rewrite Rabs_R0; apply Rlt_le; apply Rmult_lt_0_compat; auto with real zarith.
apply powerRZ_lt, Rlt_IZR; omega.
unfold FtoRradix; apply RoundedModeProjectorIdemEq with b p (Closest b radix); auto.
apply ClosestRoundedModeP with p; auto with zarith.
elim H1; intros; split; auto.
rewrite H; auto.
split.
apply Rmult_le_reg_l with (INR 2); auto with real zarith.
apply Rle_trans with (powerRZ radix (- dExp b));[idtac|simpl; right; field; auto with real].
replace (a-a')%R with (-(a'-a))%R;[rewrite Rabs_Ropp|ring].
apply Rle_trans with (Fulp b radix p a).
unfold FtoRradix; apply ClosestUlp; auto with zarith.
unfold Fulp; apply Rle_powerRZ; try apply Rle_IZR; auto with real zarith.
apply Z.le_trans with (Fexp (firstNormalPos radix b p));[idtac|unfold firstNormalPos; simpl; auto with zarith].
apply Fcanonic_Rle_Zle with radix b p; auto with zarith.
apply FnormalizeCanonic; auto with zarith; elim H2; auto.
left; apply firstNormalPosNormal; auto with zarith.
rewrite (Rabs_right ((FtoR radix (firstNormalPos radix b p)))).
rewrite FnormalizeCorrect; auto with zarith.
apply RoundAbsMonotoner with b p (Closest b radix) (FtoRradix a'); auto.
apply ClosestRoundedModeP with p; auto with zarith.
assert (Fnormal radix b (firstNormalPos radix b p));
 [apply firstNormalPosNormal; auto with zarith| elim H4; auto].
unfold FtoRradix; rewrite <- Fabs_correct; auto.
unfold firstNormalPos, Fabs, FtoR; simpl.
apply Rle_trans with (powerRZ radix p * powerRZ radix (Fexp a'))%R.
apply Rmult_le_compat_r; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
elim H1; intros; apply Rle_trans with (IZR (Zpos (vNum b'))); auto with real zarith.
apply Rle_IZR; auto with zarith.
rewrite <- H; rewrite pGivesBound; rewrite Zpower_nat_Z_powerRZ; auto with real zarith.
unfold nNormMin; rewrite Zpower_nat_Z_powerRZ.
repeat rewrite <- powerRZ_add; try apply IZR_neq; auto with real zarith.
apply Rle_powerRZ; try apply Rle_IZR; auto with real zarith.
apply Rle_ge; apply LeFnumZERO; auto.
unfold firstNormalPos, nNormMin; simpl; auto with zarith.
apply Zpower_NR0; omega.
intros ; absurd (Fexp a' < - dExp b)%Z; auto with zarith.
Qed.

Theorem Underf_Err2_aux: forall (r:R) (x1:float),
    vNum b=vNum b' -> (dExp b <= dExp b')%Z ->
    (Fcanonic radix b x1) ->
    (Closest b radix r x1) ->
    (exists x2:float, (Underf_Err x1 x2 r (3/4)%R) /\ (Closest b' radix r x2)).
intros.
assert (ZH: (0 < 3/4)%R).
apply Rmult_lt_reg_l with 4%R; auto with real.
ring_simplify (4*0)%R; apply Rlt_le_trans with 3%R; auto with real.
right; field; auto with real.
case (Zle_lt_or_eq (-(dExp b))%Z (Fexp x1)).
elim H2; intros I1 I2; elim I1; auto.
intros I.
exists x1; split.
split; auto.
assert (Fbounded b x1);[elim H2; auto|idtac].
split.
split; auto with zarith.
elim H3; intros; rewrite <- H; auto.
split;[idtac|intros; auto with real].
ring_simplify (x1-x1)%R; rewrite Rabs_R0.
apply Rlt_le; apply Rmult_lt_0_compat; auto with real zarith.
apply powerRZ_lt, Rlt_IZR; omega.
split;[elim H2; intros T T'; elim T; intros; split; try rewrite <- H; auto with zarith|idtac].
intros f H3.
case (Zle_or_lt (-(dExp b)) (Fexp f)); intros.
elim H2; intros T1 T2; apply T2.
elim H3; intros; split; try rewrite H; auto with zarith.
fold FtoRradix; replace (f-r)%R with (-((x1-f)-(x1-r)))%R;[rewrite Rabs_Ropp|ring].
apply Rle_trans with (Rabs (x1 - f) - Rabs (x1 - r))%R;[idtac|apply Rabs_triang_inv].
apply Rplus_le_reg_l with (Rabs (x1-r)).
apply Rle_trans with ((INR 2)*(Rabs (x1-r)))%R;[right; simpl; ring|idtac].
apply Rle_trans with (Rabs (x1 - f));[idtac|right; ring].
apply Rle_trans with (Fulp b radix p x1).
replace (x1-r)%R with (-(r-x1))%R;[rewrite Rabs_Ropp|ring].
unfold FtoRradix; apply ClosestUlp; auto with zarith.
rewrite CanonicFulp; auto with zarith.
apply Rle_trans with (powerRZ radix (Fexp x1));[right; unfold FtoR; simpl; ring|idtac].
apply Rle_trans with ((Rabs x1)-Rabs f)%R;[idtac|apply Rabs_triang_inv].
apply Rplus_le_reg_l with (Rabs f).
apply Rle_trans with (Rabs x1);[idtac|right;ring].
apply Rle_trans with (powerRZ radix (p-2+Fexp x1)+powerRZ radix (p-2+Fexp x1))%R.
apply Rplus_le_compat.
apply Rle_trans with (FtoRradix (Float (Zpos (vNum b')) (Fexp f))).
apply Rlt_le; unfold FtoRradix; apply MaxFloat; auto.
unfold FtoRradix, FtoR; rewrite <- H; rewrite pGivesBound;simpl.
rewrite Zpower_nat_Z_powerRZ; rewrite <- powerRZ_add.
apply Rle_powerRZ; try apply Rle_IZR; auto with zarith real.
apply IZR_neq; omega.
apply Rle_powerRZ; try apply Rle_IZR; auto with zarith real.
apply Rle_trans with (2*powerRZ radix (p - 2 + Fexp x1))%R;[right; ring|idtac].
apply Rle_trans with (radix*powerRZ radix (p - 2 + Fexp x1))%R;
  [apply Rmult_le_compat_r; auto with real zarith|idtac].
apply powerRZ_le, Rlt_IZR; omega.
apply Rle_IZR; omega.
unfold FtoRradix; rewrite <- Fabs_correct; auto; unfold FtoR, Fabs; simpl.
rewrite powerRZ_add.
rewrite <- Rmult_assoc; apply Rmult_le_compat_r; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
case H1; intros T.
elim T; intros H5 H6.
apply Rmult_le_reg_l with (IZR radix); auto with real zarith.
apply Rlt_IZR; omega.
apply Rle_trans with (IZR (Zpos (vNum b))).
right; rewrite pGivesBound; rewrite Zpower_nat_Z_powerRZ.
unfold Zminus; rewrite powerRZ_add; simpl.
ring_simplify (radix*1)%R; field; auto with real zarith.
apply IZR_neq; omega.
apply IZR_neq; omega.
apply Rle_trans with (IZR(Z.abs (radix * Fnum x1))); auto with real zarith.
apply Rle_IZR; auto with zarith real.
rewrite Zabs_Zmult; rewrite Z.abs_eq; auto with zarith real.
right; rewrite mult_IZR; ring.
elim T; intros T1 T2; elim T2; intros T3 T4.
absurd (- dExp b < Fexp x1)%Z; auto with zarith.
apply IZR_neq; omega.
intros I.
generalize ClosestTotal; unfold TotalP.
intros T; elim T with b' radix p r; auto.
2: rewrite <- H; auto.
intros x2 H3'; clear T.
case (Zle_or_lt (-(dExp b)) (Fexp x2)); intros.
exists x1; split.
split; auto.
assert (Fbounded b x1);[elim H2; auto|idtac].
split.
elim H4; intros; split; auto with zarith.
split;[idtac|intros; auto with real].
ring_simplify (x1-x1)%R; rewrite Rabs_R0; apply Rlt_le.
apply Rmult_lt_0_compat; auto with real zarith.
apply powerRZ_lt, Rlt_IZR; omega.
split.
elim H2; intros T1 T2; elim T1; intros; split; try rewrite <- H; auto with zarith.
intros f H4.
apply Rle_trans with (Rabs (FtoR radix x2 - r)).
elim H2; intros T1 T2; apply T2.
elim H3'; intros T1' T2'; elim T1'; intros; split; try rewrite H; auto with zarith.
elim H3'; intros T1 T2; apply T2; auto.
exists x2; split; auto.
split; auto.
split;[elim H3'; auto|idtac].
split.
replace (x1-x2)%R with ((-(r-x1))+(r-x2))%R;[idtac|ring].
apply Rle_trans with (1:=Rabs_triang (-(r-x1))%R (r-x2)%R).
rewrite Rabs_Ropp; apply Rmult_le_reg_l with (INR 2); auto with real zarith.
apply Rle_trans with (S 1 * (Rabs (r - x1)) + S 1 *Rabs (r - x2))%R;[right; ring|idtac].
apply Rle_trans with ( powerRZ radix (- dExp b)+ (/2)*powerRZ radix (- dExp b))%R;[idtac|simpl; right; field].
apply Rplus_le_compat.
apply Rle_trans with (Fulp b radix p x1).
unfold FtoRradix; apply ClosestUlp; auto.
rewrite CanonicFulp; auto.
rewrite <- I; unfold FtoR; simpl; right; ring.
apply Rle_trans with (Fulp b' radix p x2).
unfold FtoRradix; apply ClosestUlp; auto.
rewrite <- H; auto.
apply Rle_trans with (powerRZ radix (Fexp x2)).
unfold Fulp; apply Rle_powerRZ; auto with zarith real.
apply Rle_IZR; omega.
apply FcanonicLeastExp with radix b' p; auto with zarith.
rewrite FnormalizeCorrect; auto with zarith real.
elim H3'; auto.
apply FnormalizeCanonic; auto with zarith.
elim H3'; auto.
apply Rmult_le_reg_l with 2%R; auto with real.
apply Rle_trans with (powerRZ radix (- dExp b));[idtac|right; field; auto with real].
apply Rle_trans with (radix * powerRZ radix (Fexp x2))%R;[apply Rmult_le_compat_r; auto with real zarith|idtac].
apply powerRZ_le, Rlt_IZR; omega.
apply Rle_IZR; omega.
apply Rle_trans with (powerRZ radix (Fexp x2+1)).
rewrite powerRZ_add; try apply IZR_neq; auto with real zarith; simpl; right; ring.
apply Rle_powerRZ; auto with real zarith.
apply Rle_IZR; omega.
intros H5; absurd ((Fexp x2 < - dExp b)%Z); auto with zarith.
Qed.

Theorem Underf_Err2: forall (r:R) (x1:float),
    vNum b=vNum b' -> (dExp b <= dExp b')%Z ->
    (Closest b radix r x1) ->
    (exists x2:float, (Underf_Err x1 x2 r (3/4)%R) /\ (Closest b' radix r x2)).
intros.
elim Underf_Err2_aux with r (Fnormalize radix b p x1); auto with zarith.
unfold Underf_Err; intros x2 tmp; elim tmp; intros T Z; elim T; intros V1 T'; elim T';
   intros V2 T''; elim T''; intros V3 V4; clear T T' T'' tmp.
exists x2; split; auto.
split; auto.
split; auto.
split; auto.
replace (x1-x2)%R with (Fnormalize radix b p x1 - x2)%R; auto with real.
unfold FtoRradix; rewrite FnormalizeCorrect; auto.
intros; apply trans_eq with (FtoRradix (Fnormalize radix b p x1)).
unfold FtoRradix; rewrite FnormalizeCorrect; auto.
apply V4; auto.
apply FnormalizeCanonic; auto with zarith; elim H1; auto.
apply ClosestCompatible with (1:=H1); auto.
rewrite <- FnormalizeCorrect with radix b p x1; auto.
apply FnormalizeBounded; auto with zarith; elim H1; auto.
Qed.

Theorem Underf_Err3: forall (x x' y y' z' z:float) (rx ry epsx epsy:R),
    vNum b=vNum b' -> (dExp b <= dExp b')%Z ->
   (Underf_Err x x' rx epsx) -> (Underf_Err y y' ry epsy) ->
   (epsx+epsy <= (powerRZ radix (p-1) -1))%R ->
   (Fbounded b' z') -> (FtoRradix z'=x'-y')%R ->
   (Fexp z' <= Fexp x')%Z -> (Fexp z' <= Fexp y')%Z ->
   (Closest b radix (x-y) z) ->
   (Underf_Err z z' (x-y) (epsx+epsy)%R).
intros.
unfold Underf_Err.
split; auto.
split; auto.
unfold Underf_Err in H1; unfold Underf_Err in H2.
case (Zle_or_lt (- dExp b)%Z (Fexp z')); intros.
elim H1; intros V1 T; elim T; intros V2 T'; elim T'; intros V3 V4; clear T T' H1.
elim H2; intros W1 T; elim T; intros W2 T'; elim T'; intros W3 W4; clear T T' H2.
cut (FtoRradix z=z')%R;[intros H9'; rewrite H9'; split; auto|idtac].
ring_simplify (z'-z')%R; rewrite Rabs_R0.
apply Rle_trans with (0* powerRZ radix (- dExp b))%R;[right; ring|apply Rmult_le_compat_r; auto with real zarith].
apply powerRZ_le, Rlt_IZR; omega.
apply Rle_trans with (0+0)%R; [right; ring|apply Rplus_le_compat; auto with real].
apply Rmult_le_reg_l with (powerRZ radix (- dExp b))%R; auto with real zarith;
    ring_simplify (powerRZ radix (- dExp b) * 0)%R; try rewrite Rmult_comm.
apply powerRZ_lt, Rlt_IZR; omega.
apply Rle_trans with (2:=V3); auto with real.
apply Rabs_pos.
apply Rmult_le_reg_l with (powerRZ radix (- dExp b))%R; auto with real zarith;
    ring_simplify (powerRZ radix (- dExp b) * 0)%R; try rewrite Rmult_comm.
apply powerRZ_lt, Rlt_IZR; omega.
apply Rle_trans with (2:=W3); auto with real.
apply Rabs_pos.
unfold FtoRradix; apply sym_eq.
apply RoundedModeProjectorIdemEq with b p (Closest b radix); auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
elim H4; intros; split; auto with zarith.
fold FtoRradix; rewrite H5.
rewrite <- V4; auto with zarith.
rewrite <- W4; auto with zarith real.
elim H1; intros V1 T; elim T; intros V2 T'; elim T'; intros V3 V4; clear T T' H1.
elim H2; intros W1 T; elim T; intros W2 T'; elim T'; intros W3 W4; clear T T' H2.
split;[idtac|intros; absurd (- dExp b <= Fexp z')%Z; auto with zarith].
replace (z-z')%R with (-((x-y)-z)+((x-x')+-(y-y')))%R;[idtac|rewrite H5; ring].
apply Rle_trans with (1:=Rabs_triang (- (x - y - z))%R ((x - x') + - (y - y'))%R).
apply Rle_trans with ((Rabs (- (x - y - z))) + (Rabs (x - x') +(Rabs (- (y - y')))))%R;
   [apply Rplus_le_compat_l; apply Rabs_triang|idtac].
rewrite Rabs_Ropp; rewrite Rabs_Ropp.
apply Rle_trans with (0 + ( epsx * powerRZ radix (- dExp b)
       + epsy * powerRZ radix (- dExp b)))%R;[idtac|right; ring].
apply Rplus_le_compat;[idtac|apply Rplus_le_compat; auto with real].
cut (FtoRradix (Fnormalize radix b p z)=x-y)%R.
unfold FtoRradix; rewrite FnormalizeCorrect; auto.
fold FtoRradix; intros T; rewrite T; ring_simplify (x - y - (x - y))%R; rewrite Rabs_R0; auto with real.
unfold FtoRradix, Rminus; rewrite <- Fopp_correct; auto.
apply plusExact1 with b p; auto.
elim V1; auto.
apply oppBounded; elim W1; auto.
rewrite Fopp_correct; auto with real.
apply ClosestCompatible with (1:=H8); auto.
rewrite FnormalizeCorrect; auto with real.
apply FnormalizeBounded; auto with zarith; elim H8; auto.
apply Z.le_trans with (-(dExp b))%Z.
2: apply Zmin_Zle.
2: elim V1; intros T1 T2; elim T1; auto.
2: elim W1; intros T1 T2; elim T1; auto.
apply Z.le_trans with (Fexp (Float (pPred (vNum b)) (-(dExp b))%Z));
   [idtac| simpl; auto with zarith].
apply Fcanonic_Rle_Zle with radix b p; auto with zarith.
apply FnormalizeCanonic; auto with zarith; elim H8; auto.
apply FcanonicPpred with p; auto with zarith.
rewrite (Rabs_right ((FtoR radix (Float (pPred (vNum b)) (- dExp b))))).
rewrite FnormalizeCorrect; auto with zarith.
apply RoundAbsMonotoner with b p (Closest b radix) (x-y)%R; auto.
apply ClosestRoundedModeP with p; auto with zarith.
assert (Fcanonic radix b (Float (pPred (vNum b)) (- dExp b)));
 [apply FcanonicPpred with p; auto with zarith | apply FcanonicBound with radix; auto].
replace (x-y)%R with ((x-x')+-(y-y')+z')%R;[idtac|rewrite H5; ring].
apply Rle_trans with (1:=Rabs_triang (x - x' + - (y - y'))%R z').
apply Rle_trans with ((powerRZ radix (p - 1) - 1)*powerRZ radix (-(dExp b)) +
       (powerRZ radix p - 1)*powerRZ radix (-(dExp b)-1))%R;[apply Rplus_le_compat|idtac].
apply Rle_trans with (1:=Rabs_triang (x-x')%R (-(y-y'))%R); rewrite Rabs_Ropp.
apply Rle_trans with (epsx * powerRZ radix (- dExp b)+ epsy * powerRZ radix (- dExp b))%R; auto with real.
apply Rle_trans with ((epsx+epsy) * powerRZ radix (- dExp b))%R;[right; ring|idtac].
apply Rmult_le_compat_r; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
unfold FtoRradix; rewrite <- Fabs_correct; auto.
unfold Fabs, FtoR; simpl.
apply Rmult_le_compat; auto with real zarith.
apply Rle_IZR; auto with zarith.
apply powerRZ_le, Rlt_IZR; omega.
elim H4; intros.
apply Rle_trans with (IZR (Z.pred (Zpos (vNum b')))); auto with real zarith.
apply Rle_IZR; omega.
unfold Z.pred, Zminus; rewrite plus_IZR.
rewrite <- H; rewrite pGivesBound; rewrite Zpower_nat_Z_powerRZ; auto with real zarith.
apply Rle_powerRZ; try apply Rle_IZR; auto with real zarith.
apply Rle_trans with ((powerRZ radix p - 1) * powerRZ radix (- dExp b))%R.
apply Rle_trans with (- powerRZ radix (- dExp b)+powerRZ radix p
   * powerRZ radix (- dExp b))%R;[idtac|right; ring].
apply Rle_trans with (- powerRZ radix (- dExp b)+ ((powerRZ radix (p - 1)* powerRZ radix (- dExp b)+
  (powerRZ radix p*powerRZ radix (- dExp b - 1) - powerRZ radix (- dExp b - 1)))))%R;
   [right; ring|apply Rplus_le_compat_l].
repeat rewrite <- powerRZ_add; try apply IZR_neq; auto with real zarith.
replace ((p + (- dExp b - 1)))%Z with (p - 1 + - dExp b)%Z;[idtac|ring].
apply Rle_trans with (powerRZ radix (p - 1 + - dExp b) +
    (powerRZ radix (p - 1 + - dExp b) - 0))%R; auto with real zarith.
apply Rplus_le_compat_l; unfold Rminus; apply Rplus_le_compat_l; auto with real zarith.
apply Ropp_le_contravar.
apply powerRZ_le, Rlt_IZR; omega.
apply Rle_trans with (2*(powerRZ radix (p - 1 + - dExp b)))%R;[right; ring|idtac].
apply Rle_trans with (radix*(powerRZ radix (p - 1 + - dExp b)))%R;
 [apply Rmult_le_compat_r; auto with real zarith|idtac].
apply powerRZ_le, Rlt_IZR; omega.
apply Rle_IZR; omega.
unfold Zminus; repeat rewrite powerRZ_add; try apply IZR_neq; auto with real zarith; simpl; right; field; auto with real zarith.
apply IZR_neq; omega.
unfold FtoR; simpl.
unfold pPred, Z.pred, Zminus; rewrite plus_IZR.
rewrite pGivesBound; rewrite Zpower_nat_Z_powerRZ; simpl; auto with real zarith.
apply Rle_ge; apply LeFnumZERO; auto.
simpl; apply Zlt_le_weak.
apply pPredMoreThanOne with radix p; auto with zarith.
Qed.

Theorem Underf_Err3_bis: forall (x x' y y' z' z:float) (rx ry epsx epsy:R),
   (4 <= p) ->
    vNum b=vNum b' -> (dExp b <= dExp b')%Z ->
   (Underf_Err x x' rx epsx) -> (Underf_Err y y' ry epsy) ->
   (epsx+epsy <= 7)%R ->
   (Fbounded b' z') -> (FtoRradix z'=x'-y')%R ->
   (Fexp z' <= Fexp x')%Z -> (Fexp z' <= Fexp y')%Z ->
   (Closest b radix (x-y) z) ->
   (Underf_Err z z' (x-y) (epsx+epsy)%R).
intros.
apply Underf_Err3 with x' y' rx ry; auto.
apply Rle_trans with (1:=H4).
apply Rle_trans with (8-1)%R;[right; ring|unfold Rminus; apply Rplus_le_compat_r].
apply Rle_trans with (powerRZ radix 3)%R; auto with real zarith.
apply Rle_trans with (powerRZ 2 3)%R; auto with real zarith.
simpl; right; ring.
simpl; auto with real zarith.
assert (2 <= radix)%R;[apply Rle_IZR; auto with real zarith|idtac].
ring_simplify (2*1)%R; ring_simplify (radix*1)%R.
apply Rmult_le_compat; auto with real zarith.
apply Rlt_le; apply Rmult_lt_0_compat; auto with real.
apply Rle_powerRZ; try apply Rle_IZR; auto with zarith real.
Qed.

End GenericDek.

Section Sec1.

Variable radix : Z.
Variable b : Fbound.
Variables s t:nat.

Let b' := Bound
    (P_of_succ_nat (pred (Z.abs_nat (Zpower_nat radix (minus t s)))))
    (dExp b).

Let bt := Bound
    (P_of_succ_nat (pred (Z.abs_nat (Zpower_nat radix s))))
    (dExp b).

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.
Hypothesis radixMoreThanOne : (1 < radix)%Z.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ radixMoreThanOne).

Hypothesis pGivesBound: Zpos (vNum b)=(Zpower_nat radix t).

Hypothesis SLe: (2 <= s)%nat.
Hypothesis SGe: (s <= t-2)%nat.

Hypothesis Hst1: (t-1 <= s+s)%Z.
Hypothesis Hst2: (s+s <= t+1)%Z.

Variables x x1 x2 y y1 y2 r e: float.

Hypotheses Nx: Fnormal radix b x.
Hypotheses Ny: Fnormal radix b y.

Hypothesis K: (-dExp b <= Fexp x +Fexp y)%Z.

Hypotheses rDef: Closest b radix (x*y) r.

Hypotheses eeq: (x*y=r+e)%R.
Hypotheses Xeq: (FtoRradix x=x1+x2)%R.
Hypotheses Yeq: (FtoRradix y=y1+y2)%R.

Hypotheses x2Le: (Rabs x2 <= (powerRZ radix (s+Fexp x)) /2)%R.
Hypotheses y2Le: (Rabs y2 <= (powerRZ radix (s+Fexp y)) /2)%R.
Hypotheses x1Exp: (s+Fexp x <= Fexp x1)%Z.
Hypotheses y1Exp: (s+Fexp y <= Fexp y1)%Z.
Hypotheses x2Exp: (Fexp x <= Fexp x2)%Z.
Hypotheses y2Exp: (Fexp y <= Fexp y2)%Z.

Lemma x2y2Le: (Rabs (x2*y2) <= (powerRZ radix (2*s+Fexp x+Fexp y)) /4)%R.
rewrite Rabs_mult.
apply Rle_trans with ((powerRZ radix (s + Fexp x) / 2)*(powerRZ radix (s + Fexp y) / 2))%R.
apply Rmult_le_compat; try apply Rabs_pos; auto with real.
replace (2*s)%Z with (s+s)%Z; auto with zarith.
repeat rewrite powerRZ_add; try apply IZR_neq; auto with real zarith.
right; field.
Qed.

Lemma x2y1Le: (Rabs (x2*y1) < (powerRZ radix (t+s+Fexp x+Fexp y)) /2
          + (powerRZ radix (2*s+Fexp x+Fexp y)) /4)%R.
replace (x2*y1)%R with (x2*y+(-(x2*y2)))%R;[idtac|rewrite Yeq; ring].
apply Rle_lt_trans with (1:=Rabs_triang (x2*y)%R (-(x2*y2))%R).
rewrite Rabs_Ropp.
cut ((Rabs (x2 * y) < powerRZ radix (t + s + Fexp x + Fexp y) / 2))%R;[intros I1|idtac].
generalize x2y2Le; auto with real.
rewrite Rabs_mult.
apply Rlt_le_trans with ((powerRZ radix (s + Fexp x) / 2)*powerRZ radix (t+Fexp y))%R.
cut (Rabs y <powerRZ radix (t + Fexp y))%R; auto with real zarith.
intros I; apply Rle_lt_trans with (powerRZ radix (s + Fexp x) / 2 *Rabs y)%R; auto with real.
apply Rmult_le_compat; try apply Rabs_pos; auto with real.
apply Rmult_lt_compat_l; auto with real zarith.
unfold Rdiv; apply Rmult_lt_0_compat; auto with real zarith.
apply powerRZ_lt, Rlt_IZR; omega.
unfold FtoRradix; rewrite <- Fabs_correct; auto with zarith.
unfold FtoR, Fabs; simpl; rewrite powerRZ_add.
apply Rmult_lt_compat_r; auto with real zarith.
apply powerRZ_lt, Rlt_IZR; omega.
elim Ny; intros I1 I2; elim I1; intros.
apply Rlt_le_trans with (IZR (Zpos (vNum b))); try apply Rlt_IZR; auto with real zarith.
right; rewrite pGivesBound;rewrite Zpower_nat_Z_powerRZ; auto with real.
apply IZR_neq; omega.
repeat rewrite powerRZ_add; try apply IZR_neq; auto with real zarith.
unfold Rdiv; right; ring.
Qed.

Lemma x1y2Le: (Rabs (x1*y2) < (powerRZ radix (t+s+Fexp x+Fexp y)) /2
          + (powerRZ radix (2*s+Fexp x+Fexp y)) /4)%R.
replace (x1*y2)%R with (x*y2+(-(x2*y2)))%R;[idtac|rewrite Xeq; ring].
apply Rle_lt_trans with (1:=Rabs_triang (x*y2)%R (-(x2*y2))%R).
rewrite Rabs_Ropp.
cut ((Rabs (x * y2) < powerRZ radix (t + s + Fexp x + Fexp y) / 2))%R;[intros I1|idtac].
generalize x2y2Le; auto with real.
rewrite Rabs_mult.
apply Rlt_le_trans with (powerRZ radix (t+Fexp x)*(powerRZ radix (s + Fexp y) / 2))%R.
cut (Rabs x <powerRZ radix (t + Fexp x))%R; auto with real zarith.
intros I; apply Rle_lt_trans with (Rabs x*(powerRZ radix (s + Fexp y) / 2))%R; auto with real.
apply Rmult_le_compat; try apply Rabs_pos; auto with real.
apply Rmult_lt_compat_r; auto with real zarith.
unfold Rdiv; apply Rmult_lt_0_compat; auto with real zarith.
apply powerRZ_lt, Rlt_IZR; omega.
unfold FtoRradix; rewrite <- Fabs_correct; auto with zarith.
unfold FtoR, Fabs; simpl; rewrite powerRZ_add.
apply Rmult_lt_compat_r; auto with real zarith.
apply powerRZ_lt, Rlt_IZR; omega.
elim Nx; intros I1 I2; elim I1; intros.
apply Rlt_le_trans with (IZR (Zpos (vNum b))); try apply IZR_lt; auto with real zarith.
right; rewrite pGivesBound;rewrite Zpower_nat_Z_powerRZ; auto with real.
apply IZR_neq; omega.
repeat rewrite powerRZ_add; try apply IZR_neq; auto with real zarith.
unfold Rdiv; right; ring.
Qed.

Lemma eLe: (Rabs e <= (powerRZ radix (t+Fexp x+Fexp y)) /2)%R.
apply Rmult_le_reg_l with (INR 2); auto with real zarith.
replace (FtoRradix e) with (x*y-r)%R;[idtac|rewrite eeq; ring].
apply Rle_trans with (Fulp b radix t r).
unfold FtoRradix; apply ClosestUlp; auto with zarith.
apply Rle_trans with (powerRZ radix (t + Fexp x + Fexp y));
  [idtac|simpl; right; field; auto with real].
unfold Fulp; apply Rle_powerRZ; try apply Rle_IZR; auto with real zarith.
apply Z.le_trans with (Fexp (Float (pPred (vNum b)) (t+Fexp x+Fexp y)));
  [idtac|simpl; auto with zarith].
apply Fcanonic_Rle_Zle with radix b t; auto with zarith.
apply FnormalizeCanonic; auto with zarith.
elim rDef; auto.
replace (Float (pPred (vNum b)) (t + Fexp x + Fexp y)) with
   (FPred b radix t (Float (nNormMin radix t) (t+1+Fexp x+Fexp y))).
apply FPredCanonic; auto with zarith.
apply FcanonicNnormMin; auto with zarith.
rewrite FPredSimpl2; auto with zarith.
simpl; unfold Z.pred; auto with zarith.
replace (t+1+Fexp x +Fexp y+-1)%Z with (t+Fexp x+Fexp y)%Z; auto with zarith.
simpl; auto with zarith.
rewrite FnormalizeCorrect; auto with zarith.
rewrite (Rabs_right (FtoR radix (Float (pPred (vNum b)) (t + Fexp x + Fexp y)))%R).
2: apply Rle_ge; apply LeFnumZERO; simpl; auto with zarith.
2: apply Zlt_le_weak; apply pPredMoreThanOne with radix t; auto with zarith.
apply RoundAbsMonotoner with b t (Closest b radix) (x*y)%R; auto with zarith.
apply ClosestRoundedModeP with t; auto with zarith.
replace (Float (pPred (vNum b)) (t + Fexp x + Fexp y)) with
   (FPred b radix t (Float (nNormMin radix t) (t+1+Fexp x+Fexp y))).
apply FBoundedPred; auto with zarith.
elim FnormalNnormMin with radix b t (t + 1 + Fexp x + Fexp y)%Z; auto with zarith.
rewrite FPredSimpl2; auto with zarith.
simpl; unfold Z.pred; auto with zarith.
replace (t+1+Fexp x +Fexp y+-1)%Z with (t+Fexp x+Fexp y)%Z; auto with zarith.
simpl; auto with zarith.
rewrite Rabs_mult.
apply Rle_trans with ((FtoRradix (Float (pPred (vNum b)) (Fexp x)))*(powerRZ radix (t+Fexp y)))%R.
apply Rmult_le_compat; try apply Rabs_pos; auto with real zarith.
unfold FtoRradix; rewrite <- Fabs_correct; auto with zarith.
unfold FtoR, Fabs; simpl.
apply Rmult_le_compat_r; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
elim Nx; intros I1 I2; elim I1; intros; unfold pPred; auto with real zarith.
apply Rle_IZR; omega.
unfold FtoRradix; rewrite <- Fabs_correct; auto with zarith.
unfold FtoR, Fabs; simpl.
rewrite powerRZ_add.
apply Rmult_le_compat_r; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
elim Ny; intros I1 I2; elim I1; intros; apply Rle_trans with (IZR (Zpos (vNum b))); auto with real zarith.
apply Rle_IZR; omega.
rewrite pGivesBound; rewrite Zpower_nat_Z_powerRZ; auto with real zarith.
apply IZR_neq; omega.
unfold FtoRradix, FtoR; simpl; repeat rewrite powerRZ_add; try apply IZR_neq; auto with real zarith; right; ring.
Qed.

Lemma rExp: (t - 1 + Fexp x + Fexp y <= Fexp r)%Z.
apply Z.le_trans with (Fexp (Float (nNormMin radix t) (t-1+Fexp x+Fexp y)));
  [simpl; auto with zarith|idtac].
apply Z.le_trans with (Fexp (Fnormalize radix b t r)).
apply Fcanonic_Rle_Zle with radix b t; auto with zarith.
apply FcanonicNnormMin; auto with zarith.
apply FnormalizeCanonic; auto with zarith.
elim rDef; auto.
rewrite FnormalizeCorrect; auto with zarith.
rewrite (Rabs_right (FtoR radix (Float (nNormMin radix t) (t-1 + Fexp x + Fexp y)))%R).
2: apply Rle_ge; apply LeFnumZERO; simpl; auto with zarith.
2: unfold nNormMin; auto with zarith.
apply RoundAbsMonotonel with b t (Closest b radix) (x*y)%R; auto with zarith.
apply ClosestRoundedModeP with t; auto with zarith.
apply FcanonicBound with radix; auto with zarith.
apply FcanonicNnormMin; auto with zarith.
fold FtoRradix; replace (FtoRradix (Float (nNormMin radix t) (t-1 + Fexp x + Fexp y))) with
  ((Float (nNormMin radix t) (Fexp x))*(Float (nNormMin radix t) (Fexp y)))%R.
rewrite Rabs_mult.
apply Rmult_le_compat.
unfold FtoRradix; apply LeFnumZERO; simpl; unfold nNormMin; try apply Zpower_NR0; auto with zarith.
unfold FtoRradix; apply LeFnumZERO; simpl; unfold nNormMin; try apply Zpower_NR0; auto with zarith.
unfold FtoRradix; rewrite <- Fabs_correct; auto with zarith; unfold FtoR; simpl.
apply Rmult_le_compat_r; auto with real zarith; apply Rmult_le_reg_l with (IZR radix); try apply Rlt_IZR; auto with real zarith.
apply Rmult_le_compat_l.
apply Rle_IZR; omega.
apply powerRZ_le, Rlt_IZR; omega.
rewrite <- mult_IZR; rewrite <- (PosNormMin radix b t); auto with zarith.
elim Nx; intros I1 I2; rewrite Zabs_Zmult in I2; rewrite Z.abs_eq in I2; auto with real zarith.
rewrite <- mult_IZR; apply Rle_IZR; auto with real zarith.
unfold FtoRradix; rewrite <- Fabs_correct; auto with zarith; unfold FtoR; simpl.
apply Rmult_le_compat_r; auto with real zarith; apply Rmult_le_reg_l with (IZR radix); try apply Rlt_IZR; auto with real zarith.
apply Rmult_le_compat_l.
apply Rle_IZR; omega.
apply powerRZ_le, Rlt_IZR; omega.
rewrite <- mult_IZR; rewrite <- (PosNormMin radix b t); auto with zarith.
elim Ny; intros I1 I2; rewrite Zabs_Zmult in I2; rewrite Z.abs_eq in I2; auto with real zarith.
rewrite <- mult_IZR; apply Rle_IZR; auto with real zarith.
unfold FtoRradix, FtoR; simpl.
rewrite powerRZ_add; try apply IZR_neq; auto with real zarith.
rewrite powerRZ_add; try apply IZR_neq; auto with real zarith.
replace (IZR (nNormMin radix t)) with (powerRZ radix (t-1));[ring|idtac].
unfold nNormMin; rewrite Zpower_nat_Z_powerRZ.
rewrite inj_pred; auto with zarith; unfold Z.pred; auto with real zarith.
apply Zpower_NR0; omega.
apply FcanonicLeastExp with radix b t; auto with real zarith.
rewrite FnormalizeCorrect; auto with zarith real.
elim rDef; auto.
apply FnormalizeCanonic; auto with zarith; elim rDef; auto.
Qed.

Lemma powerRZSumRle: forall (e1 e2:Z),
  (e2<= e1)%Z ->
  (powerRZ radix e1 + powerRZ radix e2 <= powerRZ radix (e1+1))%R.
Proof.
clear -radixMoreThanOne; intros.
apply Rle_trans with (powerRZ radix e1 + powerRZ radix e1)%R;
  [apply Rplus_le_compat_l; apply Rle_powerRZ; auto with real zarith|idtac].
apply Rle_IZR; omega.
apply Rle_trans with (powerRZ radix e1*2)%R;[right; ring|rewrite powerRZ_add].
apply Rmult_le_compat_l; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
simpl; ring_simplify (radix*1)%R; apply Rle_IZR; auto with real zarith.
apply IZR_neq; omega.
Qed.

Lemma Boundedt1: (exists x':float, (FtoRradix x'=r-x1*y1)%R /\ (Fbounded b x')
                   /\ (Fexp x'=t-1+Fexp x+Fexp y)%Z).
unfold FtoRradix; apply BoundedL with t (Fminus radix r (Fmult x1 y1)); auto with zarith.
unfold Fminus, Fopp, Fplus, Fmult; simpl.
apply Zmin_Zle; auto with zarith.
apply rExp.
rewrite Fminus_correct; auto with zarith; rewrite Fmult_correct; auto with real zarith.
fold FtoRradix.
replace (r-x1*y1)%R with ((-e)+x1*y2+x2*y1+x2*y2)%R.
2: apply trans_eq with (-e+x*y-(x+0)*y+x1*y2+x2*y1+x2*y2)%R;[ring|idtac].
2: rewrite eeq; rewrite Xeq; rewrite Yeq; ring.
apply Rle_lt_trans with ((Rabs e)+(Rabs (x1 * y2)) + (Rabs (x2 * y1)) + (Rabs (x2 * y2)))%R.
apply Rle_trans with (1:= Rabs_triang (-e+ x1 * y2 + x2 * y1) (x2 * y2)%R).
apply Rplus_le_compat_r.
apply Rle_trans with (1:= Rabs_triang (-e+ x1 * y2) (x2 * y1)%R).
apply Rplus_le_compat_r.
apply Rle_trans with (1:= Rabs_triang (-e) (x1 * y2)%R).
rewrite Rabs_Ropp; right; ring.
generalize eLe; generalize x1y2Le; generalize x2y1Le; generalize x2y2Le; intros.
apply Rlt_le_trans with ( (powerRZ radix (t + Fexp x + Fexp y) / 2) +
    (powerRZ radix (t + s + Fexp x + Fexp y) / 2 +powerRZ radix (2 * s + Fexp x + Fexp y) / 4) +
    (powerRZ radix (t + s + Fexp x + Fexp y) / 2 +powerRZ radix (2 * s + Fexp x + Fexp y) / 4) +
    (powerRZ radix (2 * s + Fexp x + Fexp y) / 4))%R;
  auto with real.
apply Rlt_le_trans with (powerRZ radix (t + Fexp x + Fexp y) / 2 +
    (powerRZ radix (t + s + Fexp x + Fexp y) / 2 +
     powerRZ radix (2 * s + Fexp x + Fexp y) / 4) +
    (powerRZ radix (t + s + Fexp x + Fexp y) / 2 +
     powerRZ radix (2 * s + Fexp x + Fexp y) / 4) + Rabs (x2 * y2))%R; auto with real.
apply Rplus_lt_compat_r.
apply Rle_lt_trans with (powerRZ radix (t + Fexp x + Fexp y) / 2 +
    (powerRZ radix (t + s + Fexp x + Fexp y) / 2 +
     powerRZ radix (2 * s + Fexp x + Fexp y) / 4) + Rabs (x2 * y1))%R; auto with real.
apply Rle_trans with (powerRZ radix (t + s + Fexp x + Fexp y)+
  powerRZ radix (t + Fexp x + Fexp y) / 2 +3*powerRZ radix (2 * s + Fexp x + Fexp y) / 4)%R;
  [right; field; auto with real|idtac].
assert (0 < 8)%R; auto with real.
apply Rle_trans with (powerRZ radix (t + s + Fexp x + Fexp y) +
    powerRZ radix (t + Fexp x + Fexp y) +
    powerRZ radix (2 * s + Fexp x + Fexp y))%R;
  [apply Rplus_le_compat; try apply Rplus_le_compat_l |idtac].
unfold Rdiv; apply Rle_trans with (powerRZ radix (t + Fexp x + Fexp y)*1)%R;[idtac|right; ring].
apply Rmult_le_compat_l; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
apply Rle_trans with (/1)%R; auto with real.
unfold Rdiv; apply Rle_trans with (powerRZ radix (2*s + Fexp x + Fexp y)*1)%R;[idtac|right; ring].
 apply Rle_trans with (powerRZ radix (2*s + Fexp x + Fexp y)*(3*/4))%R;[right; ring|idtac].
apply Rmult_le_compat_l; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
assert (0<4)%R.
auto with real.
apply Rmult_le_reg_l with (4%R); auto with real.
apply Rle_trans with 3%R;[right; field|idtac]; auto with real.
apply Rle_trans with 4%R; auto with real.
apply Rle_trans with (powerRZ radix (t + s + Fexp x + Fexp y) +
    powerRZ radix (t +1+ Fexp x + Fexp y) +
    powerRZ radix (2 * s + Fexp x + Fexp y))%R.
apply Rplus_le_compat_r; apply Rplus_le_compat_l.
apply Rle_powerRZ; auto with real zarith.
apply Rle_IZR; omega.
apply Rle_trans with (powerRZ radix (t + s + Fexp x + Fexp y) +
    powerRZ radix (t+1 + Fexp x + Fexp y+1))%R.
rewrite Rplus_assoc; apply Rplus_le_compat_l.
apply powerRZSumRle; auto with zarith.
apply Rle_trans with (powerRZ radix (t + s + Fexp x + Fexp y+1)).
apply powerRZSumRle; auto with zarith.
apply Rle_powerRZ; auto with real zarith.
apply Rle_IZR; omega.
Qed.

Lemma Boundedt2: (exists x':float, (FtoRradix x'=r-x1*y1-x1*y2)%R /\ (Fbounded b x')
                   /\ (Fexp x'=s+Fexp x+Fexp y)%Z).
elim Boundedt1; intros t1 T; elim T; intros H1 T'; elim T'; intros H2 H3; clear T T'.
unfold FtoRradix; apply BoundedL with t (Fminus radix t1 (Fmult x1 y2)); auto with zarith.
unfold Fminus, Fopp, Fplus, Fmult; simpl.
apply Zmin_Zle; auto with zarith.
rewrite Fminus_correct; auto with zarith; rewrite Fmult_correct; auto with real zarith.
fold FtoRradix; rewrite H1; ring.
fold FtoRradix; replace (r-x1*y1-x1*y2)%R with ((-e)+x2*y1+x2*y2)%R.
2: apply trans_eq with (-e+x*y-(x+0)*y+x2*y1+x2*y2)%R;[ring|idtac].
2: rewrite eeq; rewrite Xeq; rewrite Yeq; ring.
apply Rle_lt_trans with ((Rabs e) + (Rabs (x2 * y1)) + (Rabs (x2 * y2)))%R.
apply Rle_trans with (1:= Rabs_triang (-e + x2 * y1) (x2 * y2)%R).
apply Rplus_le_compat_r.
apply Rle_trans with (1:= Rabs_triang (-e) (x2 * y1)%R).
rewrite Rabs_Ropp; right; ring.
generalize eLe; generalize x2y1Le; generalize x2y2Le; intros.
apply Rle_lt_trans with
  ( (Rabs e + Rabs (x2 * y1)+ (powerRZ radix (2 * s + Fexp x + Fexp y) / 4)))%R; auto with real.
apply Rlt_le_trans with (Rabs e + (powerRZ radix (t + s + Fexp x + Fexp y) / 2 +
        powerRZ radix (2 * s + Fexp x + Fexp y) / 4)+powerRZ radix (2 * s + Fexp x + Fexp y) / 4)%R; auto with real.
apply Rle_trans with (powerRZ radix (t + Fexp x + Fexp y) / 2+ (powerRZ radix (t + s + Fexp x + Fexp y) / 2 +
     powerRZ radix (2 * s + Fexp x + Fexp y) / 4) +
    powerRZ radix (2 * s + Fexp x + Fexp y) / 4)%R; auto with real.
replace (s + Fexp x + Fexp y + t)%Z with (t+s+Fexp x+Fexp y)%Z;[idtac|ring].
apply Rplus_le_reg_l with (-((powerRZ radix (t + s + Fexp x + Fexp y) / 2)))%R.
apply Rle_trans with (/2* (powerRZ radix (t + Fexp x + Fexp y)+ powerRZ radix (2 * s + Fexp x + Fexp y)))%R;
   [right; field; auto with real|idtac].
apply Rle_trans with (/2* powerRZ radix (t + s + Fexp x + Fexp y))%R;[idtac|right; field; auto with real].
apply Rmult_le_compat_l; auto with real.
apply Rle_trans with (powerRZ radix (t+1 + Fexp x + Fexp y) +
    powerRZ radix (2 * s + Fexp x + Fexp y))%R; auto with real zarith.
apply Rplus_le_compat_r.
apply Rle_powerRZ; try apply Rle_IZR; auto with zarith.
apply Rle_trans with (powerRZ radix (t+1 + Fexp x + Fexp y+1)).
apply powerRZSumRle; auto with real zarith.
apply Rle_powerRZ; try apply Rle_IZR; auto with real zarith.
Qed.

Lemma Boundedt3: (exists x':float, (FtoRradix x'=r-x1*y1-x1*y2-x2*y1)%R /\ (Fbounded b x')
                   /\ (Fexp x'=s+Fexp x+Fexp y)%Z).
elim Boundedt2; intros t2 T; elim T; intros H1 T'; elim T'; intros H2 H3; clear T T'.
unfold FtoRradix; apply BoundedL with t (Fminus radix t2 (Fmult x2 y1)); auto with zarith.
unfold Fminus, Fopp, Fplus, Fmult; simpl.
apply Zmin_Zle; auto with zarith.
rewrite Fminus_correct; auto with zarith; rewrite Fmult_correct; auto with real zarith.
fold FtoRradix; rewrite H1; ring.
fold FtoRradix; replace (r-x1*y1-x1*y2-x2*y1)%R with ((-e)+x2*y2)%R.
2: apply trans_eq with (-e+x*y-(x+0)*y+x2*y2)%R;[ring|idtac].
2: rewrite eeq; rewrite Xeq; rewrite Yeq; ring.
apply Rle_lt_trans with ((Rabs e) + (Rabs (x2 * y2)))%R.
apply Rle_trans with (1:= Rabs_triang (-e) (x2 * y2)%R).
rewrite Rabs_Ropp; right; ring.
generalize eLe; generalize x2y2Le; intros.
apply Rle_lt_trans with
  (powerRZ radix (t + Fexp x + Fexp y) / 2+powerRZ radix (2 * s + Fexp x + Fexp y) / 4)%R; auto with real.
apply Rlt_le_trans with (powerRZ radix (t + Fexp x + Fexp y) +
    powerRZ radix (2 * s + Fexp x + Fexp y) / 4)%R.
apply Rplus_lt_compat_r.
apply Rlt_le_trans with (powerRZ radix (t + Fexp x + Fexp y)*1)%R;[idtac|right; ring].
unfold Rdiv; apply Rmult_lt_compat_l; auto with real zarith.
apply powerRZ_lt, Rlt_IZR; omega.
apply Rlt_le_trans with (/1)%R; auto with real.
apply Rle_trans with (powerRZ radix (t+1 + Fexp x + Fexp y) +
    powerRZ radix (2 * s + Fexp x + Fexp y))%R;[apply Rplus_le_compat|idtac].
apply Rle_powerRZ; auto with real zarith.
apply Rle_IZR; omega.
apply Rle_trans with (powerRZ radix (2*s + Fexp x + Fexp y)*1)%R;[idtac|right; ring].
unfold Rdiv; apply Rmult_le_compat_l; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
assert (0 < 4)%R;[apply Rlt_le_trans with 2%R; auto with real|idtac].
apply Rmult_le_reg_l with 4%R; auto with real.
apply Rle_trans with 1%R;[right; field|ring_simplify (4*1)%R]; auto with real.
apply Rle_trans with (powerRZ radix (t + 1 + Fexp x + Fexp y+1)).
apply powerRZSumRle; auto with zarith.
apply Rle_powerRZ; try apply Rle_IZR; auto with real zarith.
Qed.

Lemma Boundedt4: (exists x':float, (FtoRradix x'=r-x1*y1-x1*y2-x2*y1-x2*y2)%R /\ (Fbounded b x')).
elim errorBoundedMult with b radix t (Closest b radix) x y r; auto with zarith.
2: apply ClosestRoundedModeP with t; auto with zarith.
2: elim Nx; auto.
2: elim Ny; auto.
intros g T; elim T; intros H1 T'; elim T'; intros; clear T T'.
exists (Fopp g); split.
unfold FtoRradix;rewrite Fopp_correct; rewrite H1; fold FtoRradix.
rewrite Xeq; rewrite Yeq; ring.
apply oppBounded; auto.
Qed.

Lemma Boundedt4_aux: (exists x':float, (FtoRradix x'=r-x1*y1-x1*y2-x2*y1-x2*y2)%R /\ (Fbounded b x')
  /\ (Fexp x'=Fexp x+Fexp y)%Z).
elim errorBoundedMult with b radix t (Closest b radix) x y r; auto with zarith.
2: apply ClosestRoundedModeP with t; auto with zarith.
2: elim Nx; auto.
2: elim Ny; auto.
intros g T; elim T; intros H1 T'; elim T'; intros; clear T T'.
exists (Fopp g); split.
unfold FtoRradix;rewrite Fopp_correct; rewrite H1; fold FtoRradix.
rewrite Xeq; rewrite Yeq; ring.
split;[apply oppBounded; auto|simpl; auto].
Qed.

Hypotheses Fx1: Fbounded b' x1.
Hypotheses Fx2: Fbounded bt x2.
Hypotheses Fy1: Fbounded b' y1.
Hypotheses Fy2: Fbounded bt y2.
Hypothesis Hst3: (t <= s+s)%Z.

Lemma p''GivesBound: Zpos (vNum bt)=(Zpower_nat radix s).
clear -radixMoreThanOne.
unfold bt in |- *; unfold vNum in |- *.
apply
 trans_eq
  with
    (Z_of_nat
       (nat_of_P
          (P_of_succ_nat
             (pred (Z.abs_nat (Zpower_nat radix s)))))).
unfold Z_of_nat in |- *; rewrite nat_of_P_o_P_of_succ_nat_eq_succ;
 auto with zarith.
rewrite nat_of_P_o_P_of_succ_nat_eq_succ; auto with arith zarith.
cut (Z.abs (Zpower_nat radix s) = Zpower_nat radix s).
intros H; pattern (Zpower_nat radix s) at 2 in |- *; rewrite <- H.
rewrite Zabs_absolu.
rewrite <- (S_pred (Z.abs_nat (Zpower_nat radix s)) 0);
 auto with arith zarith.
apply lt_Zlt_inv; simpl in |- *; auto with zarith arith.
rewrite <- Zabs_absolu; rewrite H; auto with arith zarith.
apply Zpower_nat_less; omega.
apply Z.abs_eq; auto with arith zarith.
apply Zpower_NR0; omega.
Qed.

Lemma Boundedx1y1_aux: (exists x':float, (FtoRradix x'=x1*y1)%R /\ (Fbounded b x')
  /\ (Fexp x'=Fexp x1+Fexp y1)%Z ).
exists (Fmult x1 y1).
split;[unfold FtoRradix; rewrite Fmult_correct; auto with real zarith|idtac].
split.
unfold Fmult; split; simpl; auto with zarith.
rewrite Zabs_Zmult.
elim Fx1; elim Fy1; intros.
apply Z.lt_le_trans with (Zpos (vNum b')*Zpos (vNum b'))%Z; auto with zarith.
case (Zle_lt_or_eq 0%Z (Z.abs (Fnum x1))); auto with zarith.
intros I; apply Z.lt_le_trans with (Z.abs (Fnum x1) * Zpos (vNum b'))%Z; auto with zarith.
apply Zmult_lt_compat_l; auto with zarith.
intros I; rewrite <- I, Zmult_0_l; auto with zarith.
apply Z.mul_pos_pos; auto with zarith.
unfold b'; rewrite p'GivesBound; auto with zarith.
rewrite <- Zpower_nat_is_exp.
rewrite pGivesBound; auto with zarith.
apply Zpower_nat_monotone_le; omega.
simpl; auto.
Qed.

Lemma Boundedx1y1: (exists x':float, (FtoRradix x'=x1*y1)%R /\ (Fbounded b x')).
elim Boundedx1y1_aux; intros f T; elim T ; intros T1 T2; elim T2; intros.
exists f; split; auto.
Qed.

Lemma Boundedx1y2_aux: (exists x':float, (FtoRradix x'=x1*y2)%R /\ (Fbounded b x')
   /\ (Fexp x'=Fexp x1+Fexp y2)%Z ).
exists (Fmult x1 y2).
split;[unfold FtoRradix; rewrite Fmult_correct; auto with real zarith|idtac].
split;[idtac|simpl; auto].
unfold Fmult; split; simpl; auto with zarith.
rewrite Zabs_Zmult.
elim Fx1; elim Fy2; intros.
apply Z.lt_le_trans with (Zpos (vNum b')*Zpos (vNum bt))%Z; auto with zarith.
case (Zle_lt_or_eq 0%Z (Z.abs (Fnum x1))); auto with zarith.
intros I; apply Z.lt_le_trans with (Z.abs (Fnum x1) * Zpos (vNum bt))%Z; auto with zarith.
apply Zmult_lt_compat_l; auto with zarith.
intros I; rewrite <- I; auto with zarith.
rewrite Zmult_0_l.
apply Z.mul_pos_pos; auto with zarith.
unfold b'; rewrite p'GivesBound; auto with zarith.
rewrite p''GivesBound; auto with zarith.
rewrite <- Zpower_nat_is_exp.
rewrite pGivesBound; auto with zarith.
apply Zpower_nat_monotone_le; omega.
Qed.

Lemma Boundedx1y2: (exists x':float, (FtoRradix x'=x1*y2)%R /\ (Fbounded b x')).
elim Boundedx1y2_aux; intros f T; elim T ; intros T1 T2; elim T2; intros.
exists f; split; auto.
Qed.

Lemma Boundedx2y1_aux: (exists x':float, (FtoRradix x'=x2*y1)%R /\ (Fbounded b x')
   /\ (Fexp x'=Fexp x2+Fexp y1)%Z ).
exists (Fmult x2 y1).
split;[unfold FtoRradix; rewrite Fmult_correct; auto with real zarith|idtac].
split;[idtac|simpl; auto].
unfold Fmult; split; simpl; auto with zarith.
rewrite Zabs_Zmult.
elim Fx2; elim Fy1; intros.
apply Z.lt_le_trans with (Zpos (vNum bt)*Zpos (vNum b'))%Z; auto with zarith.
case (Zle_lt_or_eq 0%Z (Z.abs (Fnum x2))); auto with zarith.
intros I; apply Z.lt_le_trans with (Z.abs (Fnum x2) * Zpos (vNum b'))%Z; auto with zarith.
apply Zmult_lt_compat_l; auto with zarith.
intros I; rewrite <- I; auto with zarith.
rewrite Zmult_0_l.
apply Z.mul_pos_pos; auto with zarith.
unfold b'; rewrite p'GivesBound; auto with zarith.
rewrite p''GivesBound; auto with zarith.
rewrite <- Zpower_nat_is_exp.
rewrite pGivesBound; auto with zarith.
apply Zpower_nat_monotone_le; omega.
Qed.

Lemma Boundedx2y1: (exists x':float, (FtoRradix x'=x2*y1)%R /\ (Fbounded b x')).
elim Boundedx2y1_aux; intros f T; elim T ; intros T1 T2; elim T2; intros.
exists f; split; auto.
Qed.

End Sec1.

Section Algo.

Variable radix : Z.
Variable b : Fbound.
Variables t:nat.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.
Hypothesis radixMoreThanOne : (1 < radix)%Z.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ radixMoreThanOne).

Hypothesis pGivesBound: Zpos (vNum b)=(Zpower_nat radix t).
Hypotheses pGe: (4 <= t).

Variables x y p q hx tx p' q' hy ty x1y1 x1y2 x2y1 x2y2 r t1 t2 t3 t4:float.

Hypothesis Cx: (Fnormal radix b x).
Hypothesis Cy: (Fnormal radix b y).

Hypothesis Expoxy: (-dExp b <= Fexp x+Fexp y)%Z.

Let s:= t- div2 t.

Hypothesis A1: (Closest b radix (x*((powerRZ radix s)+1))%R p).
Hypothesis A2: (Closest b radix (x-p)%R q).
Hypothesis A3: (Closest b radix (q+p)%R hx).
Hypothesis A4: (Closest b radix (x-hx)%R tx).

Hypothesis B1: (Closest b radix (y*((powerRZ radix s)+1))%R p').
Hypothesis B2: (Closest b radix (y-p')%R q').
Hypothesis B3: (Closest b radix (q'+p')%R hy).
Hypothesis B4: (Closest b radix (y-hy)%R ty).

Hypothesis C1: (Closest b radix (hx*hy)%R x1y1).
Hypothesis C2: (Closest b radix (hx*ty)%R x1y2).
Hypothesis C3: (Closest b radix (tx*hy)%R x2y1).
Hypothesis C4: (Closest b radix (tx*ty)%R x2y2).

Hypothesis D1: (Closest b radix (x*y)%R r).
Hypothesis D2: (Closest b radix (r-x1y1)%R t1).
Hypothesis D3: (Closest b radix (t1-x1y2)%R t2).
Hypothesis D4: (Closest b radix (t2-x2y1)%R t3).
Hypothesis D5: (Closest b radix (t3-x2y2)%R t4).

Lemma SLe: (2 <= s)%nat.
unfold s; auto with zarith.
assert (2<= t-div2 t)%Z; auto with zarith.
apply Zmult_le_reg_r with 2%Z; auto with zarith.
replace ((t-div2 t)*2)%Z with (2*t-2*div2 t)%Z; auto with zarith.
replace (2*div2 t)%Z with (Z_of_nat (Div2.double (div2 t))).
case (even_or_odd t); intros I.
rewrite <- even_double; auto with zarith.
apply Z.le_trans with (2*t+1-(S ( Div2.double (div2 t))))%Z; auto with zarith.
rewrite <- odd_double; auto with zarith.
replace (Z_of_nat (S ( Div2.double (div2 t)))) with (1+ Div2.double (div2 t))%Z; auto with zarith.
unfold Div2.double; rewrite inj_plus; ring.
Qed.

Lemma SGe: (s <= t-2)%nat.
unfold s; auto with zarith.
assert (2<= div2 t)%Z; auto with zarith.
apply Zmult_le_reg_r with 2%Z; auto with zarith.
replace (div2 t*2)%Z with (Z_of_nat (Div2.double (div2 t))).
case (even_or_odd t); intros I.
rewrite <- even_double; auto with zarith.
apply Z.le_trans with (-1+(S ( Div2.double (div2 t))))%Z; auto with zarith.
rewrite <- odd_double; auto with zarith.
case (Zle_lt_or_eq 4 t); auto with zarith.
intros I2; absurd (odd t); auto.
intros I3; apply not_even_and_odd with t; auto.
replace t with (4%nat); auto with zarith.
apply even_S; apply odd_S; apply even_S; apply odd_S; apply even_O.
unfold Div2.double; rewrite inj_plus; ring.
Qed.

Lemma s2Ge: (t <= s + s)%Z.
unfold s.
assert (2*(div2 t) <= t)%Z; auto with zarith.
case (even_or_odd t); intros I.
apply Z.le_trans with (Div2.double (div2 t)).
unfold Div2.double; rewrite inj_plus; auto with zarith.
rewrite <- even_double; auto with zarith.
apply Z.le_trans with (-1+(S ( Div2.double (div2 t))))%Z; auto with zarith.
rewrite inj_S; unfold Z.succ; auto with zarith.
unfold Div2.double; rewrite inj_plus; auto with zarith.
rewrite <- odd_double; auto with zarith.
Qed.

Lemma s2Le: (s + s <= t + 1)%Z.
unfold s.
rewrite inj_minus1; auto with zarith.
2: generalize (lt_div2 t); auto with zarith.
assert (t<= 2*(div2 t)+1)%Z; auto with zarith.
case (even_or_odd t); intros I.
apply Z.le_trans with ((Div2.double (div2 t)+1))%Z.
2:unfold Div2.double; rewrite inj_plus; auto with zarith.
rewrite <- even_double; auto with zarith.
apply Z.le_trans with ((S ( Div2.double (div2 t))))%Z; auto with zarith.
2: rewrite inj_S; unfold Z.succ; auto with zarith.
2: unfold Div2.double; rewrite inj_plus; auto with zarith.
rewrite <- odd_double; auto with zarith.
Qed.

Theorem Dekker_aux: (exists x':float, (FtoRradix x'=tx*ty)%R /\ (Fbounded b x'))
    -> (x*y=r-t4)%R.
intros L1.
generalize SLe; intros Sle; generalize SGe; intros Sge.
generalize s2Le; intros s2le; generalize s2Ge; intros s2ge.
generalize VeltkampU; intros V.
elim V with radix b s t x p q hx tx; auto.
2: left; auto.
intros MX1 T; elim T; intros MX2 T'; clear T; elim T'; intros T1 T2; clear T'.
elim T1; intros hx' T1'; elim T1'; intros MX3 T1''; elim T1''; intros MX4 MX5; clear T1 T1' T1''.
lapply MX5; auto; clear MX5; intros MX5.
elim T2; intros tx' T1'; elim T1'; intros MX6 T1''; elim T1''; intros MX7 MX8; clear T2 T1' T1''.
elim V with radix b s t y p' q' hy ty; auto.
2: left; auto.
intros MY1 T; elim T; intros MY2 T'; clear T; elim T'; intros T1 T2; clear T'.
elim T1; intros hy' T1'; elim T1'; intros MY3 T1''; elim T1''; intros MY4 MY5; clear T1 T1' T1''.
lapply MY5; auto; clear MY5; intros MY5.
elim T2; intros ty' T1'; elim T1'; intros MY6 T1''; elim T1''; intros MY7 MY8; clear T2 T1' T1'' V.
generalize Boundedt1; intros V.
elim V with radix b s t x hx' tx' y hy' ty' r (Fminus radix (Fmult x y) r); auto with zarith real; clear V.
2:rewrite Fminus_correct; auto with zarith; rewrite Fmult_correct; auto with zarith; ring.
2:rewrite MX6; rewrite MX3; exact MX2.
2:rewrite MY6; rewrite MY3; exact MY2.
2:rewrite MX6; replace (FtoR radix tx) with (FtoR radix x-FtoR radix hx)%R; auto with real.
2:rewrite MX2; ring.
2:rewrite MY6; replace (FtoR radix ty) with (FtoR radix y-FtoR radix hy)%R; auto with real.
2:rewrite MY2; ring.
intros t1' T; elim T; intros M11 T'; elim T'; intros M12 M13; clear T T'.
generalize Boundedt2; intros V.
elim V with radix b s t x hx' tx' y hy' ty' r (Fminus radix (Fmult x y) r); auto with zarith real; clear V.
2:rewrite Fminus_correct; auto with zarith; rewrite Fmult_correct; auto with zarith; ring.
2:rewrite MX6; rewrite MX3; exact MX2.
2:rewrite MY6; rewrite MY3; exact MY2.
2:rewrite MX6; replace (FtoR radix tx) with (FtoR radix x-FtoR radix hx)%R; auto with real.
2:rewrite MX2; ring.
2:rewrite MY6; replace (FtoR radix ty) with (FtoR radix y-FtoR radix hy)%R; auto with real.
2:rewrite MY2; ring.
intros t2' T; elim T; intros M21 T'; elim T'; intros M22 M23; clear T T'.
generalize Boundedt3; intros V.
elim V with radix b s t x hx' tx' y hy' ty' r (Fminus radix (Fmult x y) r); auto with zarith real; clear V.
2:rewrite Fminus_correct; auto with zarith; rewrite Fmult_correct; auto with zarith; ring.
2:rewrite MX6; rewrite MX3; exact MX2.
2:rewrite MY6; rewrite MY3; exact MY2.
2:rewrite MX6; replace (FtoR radix tx) with (FtoR radix x-FtoR radix hx)%R; auto with real.
2:rewrite MX2; ring.
2:rewrite MY6; replace (FtoR radix ty) with (FtoR radix y-FtoR radix hy)%R; auto with real.
2:rewrite MY2; ring.
intros t3' T; elim T; intros M31 T'; elim T'; intros M32 M33; clear T T'.
generalize Boundedt4; intros V.
elim V with radix b s t x hx' tx' y hy' ty' r ; auto with zarith real; clear V.
2:rewrite MX6; rewrite MX3; exact MX2.
2:rewrite MY6; rewrite MY3; exact MY2.
intros t4' T; elim T; intros M41 M42; clear T.
cut (FtoRradix t4=r-x*y)%R; auto with real.
intros V; rewrite V; ring.
apply sym_eq.
apply trans_eq with (FtoRradix t4').
unfold FtoRradix; rewrite M41; rewrite MX2; rewrite MY2.
rewrite MX3; rewrite MX6; rewrite MY3; rewrite MY6; ring.
unfold FtoRradix; apply RoundedModeProjectorIdemEq with b t (Closest b radix); auto with zarith.
apply ClosestRoundedModeP with t; auto with zarith.
replace (FtoR radix t4') with (t3 - x2y2)%R; auto.
replace (FtoRradix t3) with (FtoRradix t3').
replace (FtoRradix x2y2) with (tx*ty)%R.
unfold FtoRradix; rewrite M31; rewrite M41.
rewrite <- MY6; rewrite <- MX6; ring.
elim L1; intros v T; elim T; intros L2 L3.
rewrite <- L2; unfold FtoRradix.
apply RoundedModeProjectorIdemEq with b t (Closest b radix); auto with zarith.
apply ClosestRoundedModeP with t; auto with zarith.
replace (FtoR radix v) with (tx*ty)%R; auto with real.
unfold FtoRradix; apply RoundedModeProjectorIdemEq with b t (Closest b radix); auto with zarith.
apply ClosestRoundedModeP with t; auto with zarith.
replace (FtoR radix t3') with (t2-x2y1)%R; auto with real.
replace (FtoRradix t2) with (FtoRradix t2').
replace (FtoRradix x2y1) with (tx*hy)%R.
unfold FtoRradix; rewrite M21; rewrite M31.
rewrite <- MX6; rewrite <- MY3; ring.
elim Boundedx2y1 with radix b s t x tx' y hy'; auto with zarith.
intros v T; elim T; intros L2 L3.
apply trans_eq with (FtoR radix v).
unfold FtoRradix; rewrite L2; rewrite MX6; rewrite MY3; ring.
unfold FtoRradix; apply RoundedModeProjectorIdemEq with b t (Closest b radix); auto with zarith.
apply ClosestRoundedModeP with t; auto with zarith.
replace (FtoR radix v) with (tx*hy)%R; auto with real.
unfold FtoRradix; rewrite L2; rewrite MX6; rewrite MY3; ring.
unfold FtoRradix; apply RoundedModeProjectorIdemEq with b t (Closest b radix); auto with zarith.
apply ClosestRoundedModeP with t; auto with zarith.
replace (FtoR radix t2') with (t1-x1y2)%R; auto with real.
replace (FtoRradix t1) with (FtoRradix t1').
replace (FtoRradix x1y2) with (hx*ty)%R.
unfold FtoRradix; rewrite M21; rewrite M11.
rewrite <- MX3; rewrite <- MY6; ring.
elim Boundedx1y2 with radix b s t x hx' y ty'; auto with zarith.
intros v T; elim T; intros L2 L3; clear T.
apply trans_eq with (FtoR radix v).
unfold FtoRradix; rewrite L2; rewrite MY6; rewrite MX3; ring.
unfold FtoRradix; apply RoundedModeProjectorIdemEq with b t (Closest b radix); auto with zarith.
apply ClosestRoundedModeP with t; auto with zarith.
replace (FtoR radix v) with (hx*ty)%R; auto with real.
unfold FtoRradix; rewrite L2; rewrite MY6; rewrite MX3; ring.
unfold FtoRradix; apply RoundedModeProjectorIdemEq with b t (Closest b radix); auto with zarith.
apply ClosestRoundedModeP with t; auto with zarith.
replace (FtoR radix t1') with (r-x1y1)%R; auto with real.
replace (FtoRradix x1y1) with (hx*hy)%R.
unfold FtoRradix; rewrite M11; rewrite MY3; rewrite MX3; ring.
elim Boundedx1y1 with radix b s t x hx' y hy'; auto with zarith.
intros v T; elim T; intros L2 L3; clear T.
apply trans_eq with (FtoR radix v).
unfold FtoRradix; rewrite L2; rewrite MY3; rewrite MX3; ring.
unfold FtoRradix; apply RoundedModeProjectorIdemEq with b t (Closest b radix); auto with zarith.
apply ClosestRoundedModeP with t; auto with zarith.
replace (FtoR radix v) with (hx*hy)%R; auto with real.
unfold FtoRradix; rewrite L2; rewrite MY3; rewrite MX3; ring.
Qed.

Theorem Boundedx2y2: (radix=2)%Z \/ (even t) ->
    (exists x':float, (FtoRradix x'=tx*ty)%R /\ (Fbounded b x') /\ (Fexp x+Fexp y <= Fexp x')%Z).
intros H; case H; clear H; intros H.
generalize SLe; intros Sle; generalize SGe; intros Sge.
elim Veltkamp_tail2 with radix b s t x p q hx tx; auto.
2: elim Cx; auto.
intros x2 T; elim T; intros G1 T'; elim T'; intros G2 T''; elim T''; intros G3 G4; clear T T' T''.
elim Veltkamp_tail2 with radix b s t y p' q' hy ty; auto.
2: elim Cy; auto.
intros y2 T; elim T; intros J1 T'; elim T'; intros J2 T''; elim T''; intros J3 J4; clear T T' T''.
exists (Fmult x2 y2).
split;[unfold FtoRradix; rewrite Fmult_correct; auto with real zarith|idtac].
rewrite G1; rewrite J1; ring.
split.
unfold Fmult; split; simpl; auto with zarith.
rewrite Zabs_Zmult.
elim J3; elim G3; replace (Zpos
          (vNum
             (Bound
                (P_of_succ_nat (pred (Z.abs_nat (Zpower_nat radix (s - 1)))))
                (dExp b))))%Z with (Zpower_nat radix (s - 1)); intros.
apply Z.lt_le_trans with (Zpower_nat radix (s - 1)*Zpower_nat radix (s - 1))%Z; auto with zarith.
case (Zle_lt_or_eq 0%Z (Z.abs (Fnum x2))); auto with zarith.
intros I; apply Z.lt_le_trans with (Z.abs (Fnum x2) * Zpower_nat radix (s-1))%Z; auto with zarith.
apply Zmult_lt_compat_l; auto with zarith.
intros I; rewrite <- I; auto with zarith.
rewrite Zmult_0_l; apply Z.mul_pos_pos; apply Zpower_nat_less; omega.
rewrite pGivesBound; rewrite <- Zpower_nat_is_exp; auto with zarith.
apply Zpower_nat_monotone_le; auto with zarith.
generalize s2Le; auto with zarith.
apply sym_eq; unfold vNum in |- *.
apply
 trans_eq
  with
    (Z_of_nat
       (nat_of_P
          (P_of_succ_nat
             (pred (Z.abs_nat (Zpower_nat radix (s-1))))))).
unfold Z_of_nat in |- *; rewrite nat_of_P_o_P_of_succ_nat_eq_succ;
 auto with zarith.
rewrite nat_of_P_o_P_of_succ_nat_eq_succ; auto with arith zarith.
cut (Z.abs (Zpower_nat radix (s-1)) = Zpower_nat radix (s-1)).
intros HA; pattern (Zpower_nat radix (s-1)) at 2 in |- *; rewrite <- HA.
rewrite Zabs_absolu.
rewrite <- (S_pred (Z.abs_nat (Zpower_nat radix (s-1))) 0);
 auto with arith zarith.
apply lt_Zlt_inv; simpl in |- *; auto with zarith arith.
rewrite <- Zabs_absolu; rewrite HA; auto with arith zarith.
apply Zpower_nat_less; omega.
apply Z.abs_eq; auto with arith zarith.
apply Zpower_NR0; omega.
apply Z.le_trans with (Fexp (Fnormalize radix b t x)+Fexp (Fnormalize radix b t y))%Z; auto with zarith.
rewrite FcanonicFnormalizeEq; auto with zarith.
rewrite FcanonicFnormalizeEq; auto with zarith.
left; auto.
left; auto.
apply Z.le_trans with (Fexp (Fnormalize radix b t x)+Fexp (Fnormalize radix b t y))%Z; simpl; auto with zarith.
rewrite FcanonicFnormalizeEq; auto with zarith.
rewrite FcanonicFnormalizeEq; auto with zarith.
left; auto.
left; auto.
generalize SLe; intros Sle; generalize SGe; intros Sge.
elim Veltkamp_tail with radix b s t x p q hx tx; auto.
2: elim Cx; auto.
intros x2 T; elim T; intros G1 T'; elim T'; intros G2 T''; elim T''; intros G3 G4; clear T T' T''.
elim Veltkamp_tail with radix b s t y p' q' hy ty; auto.
2: elim Cy; auto.
intros y2 T; elim T; intros J1 T'; elim T'; intros J2 T''; elim T''; intros J3 J4; clear T T' T''.
exists (Fmult x2 y2).
split;[unfold FtoRradix; rewrite Fmult_correct; auto with real zarith|idtac].
rewrite G1; rewrite J1; auto with real.
split.
unfold Fmult; split; simpl; auto with zarith.
rewrite Zabs_Zmult.
elim J3; elim G3; replace (Zpos
          (vNum
             (Bound
                (P_of_succ_nat (pred (Z.abs_nat (Zpower_nat radix (s)))))
                (dExp b))))%Z with (Zpower_nat radix (s)); intros.
apply Z.lt_le_trans with (Zpower_nat radix s*Zpower_nat radix s)%Z; auto with zarith.
case (Zle_lt_or_eq 0%Z (Z.abs (Fnum x2))); auto with zarith.
intros I; apply Z.lt_le_trans with (Z.abs (Fnum x2) * Zpower_nat radix s)%Z; auto with zarith.
apply Zmult_lt_compat_l; auto with zarith.
intros I; rewrite <- I; auto with zarith.
rewrite Zmult_0_l; apply Z.mul_pos_pos; apply Zpower_nat_less; omega.
rewrite <- Zpower_nat_is_exp; rewrite pGivesBound; auto with zarith.
apply Zpower_nat_monotone_le; auto with zarith.
assert (2*s <= t)%Z; auto with zarith.
unfold s.
rewrite inj_minus1 by (generalize (lt_div2 t); auto with zarith).
assert (t <= 2*(div2 t))%Z; auto with zarith.
apply Z.le_trans with (Div2.double (div2 t)).
2: unfold Div2.double; rewrite inj_plus; auto with zarith.
rewrite <- even_double; auto with zarith.
apply sym_eq; apply p''GivesBound; auto.
apply Z.le_trans with (Fexp (Fnormalize radix b t x)+Fexp (Fnormalize radix b t y))%Z; auto with zarith.
rewrite FcanonicFnormalizeEq; auto with zarith.
rewrite FcanonicFnormalizeEq; auto with zarith.
left; auto.
left; auto.
apply Z.le_trans with (Fexp (Fnormalize radix b t x)+Fexp (Fnormalize radix b t y))%Z; auto with zarith.
rewrite FcanonicFnormalizeEq; auto with zarith.
rewrite FcanonicFnormalizeEq; auto with zarith.
left; auto.
left; auto.
simpl; auto with zarith.
Qed.

Theorem DekkerN: (radix=2)%Z \/ (even t) -> (x*y=r-t4)%R.
intros H; apply Dekker_aux.
elim Boundedx2y2; auto.
intros f T; exists f; intuition.
Qed.

End Algo.

Section AlgoS1.

Variable radix : Z.
Variable b : Fbound.
Variables t:nat.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.
Hypothesis radixMoreThanOne : (1 < radix)%Z.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ radixMoreThanOne).

Hypothesis pGivesBound: Zpos (vNum b)=(Zpower_nat radix t).
Hypotheses pGe: (4 <= t).

Variables x y p q hx tx p' q' hy ty x1y1 x1y2 x2y1 x2y2 r t1 t2 t3 t4:float.

Hypothesis Cx: (Fnormal radix b x).
Hypothesis Cy: (Fsubnormal radix b y).

Hypothesis Expoxy: (-dExp b <= Fexp x+Fexp y)%Z.

Let s:= t- div2 t.

Hypothesis A1: (Closest b radix (x*((powerRZ radix s)+1))%R p).
Hypothesis A2: (Closest b radix (x-p)%R q).
Hypothesis A3: (Closest b radix (q+p)%R hx).
Hypothesis A4: (Closest b radix (x-hx)%R tx).

Hypothesis B1: (Closest b radix (y*((powerRZ radix s)+1))%R p').
Hypothesis B2: (Closest b radix (y-p')%R q').
Hypothesis B3: (Closest b radix (q'+p')%R hy).
Hypothesis B4: (Closest b radix (y-hy)%R ty).

Hypothesis C1: (Closest b radix (hx*hy)%R x1y1).
Hypothesis C2: (Closest b radix (hx*ty)%R x1y2).
Hypothesis C3: (Closest b radix (tx*hy)%R x2y1).
Hypothesis C4: (Closest b radix (tx*ty)%R x2y2).

Hypothesis D1: (Closest b radix (x*y)%R r).
Hypothesis D2: (Closest b radix (r-x1y1)%R t1).
Hypothesis D3: (Closest b radix (t1-x1y2)%R t2).
Hypothesis D4: (Closest b radix (t2-x2y1)%R t3).
Hypothesis D5: (Closest b radix (t3-x2y2)%R t4).

Theorem DekkerS1: (radix=2)%Z \/ (even t) -> (x*y=r-t4)%R.
intros H; unfold FtoRradix.
case (Req_dec 0%R y); intros Ny.
cut (FtoRradix r=0)%R;[intros Z1|idtac].
cut (FtoRradix t4=0)%R;[intros Z2|idtac].
fold FtoRradix; rewrite Z1; rewrite Z2; rewrite <- Ny; ring.
cut (FtoRradix hy=0)%R;[intros Z3|idtac].
cut (FtoRradix ty=0)%R;[intros Z4|idtac].
unfold FtoRradix; apply ClosestZero with b t (t3-x2y2)%R; auto with zarith.
cut (FtoRradix t3=0)%R;[intros Z5|idtac].
cut (FtoRradix x2y2=0)%R;[intros Z6|idtac].
rewrite Z5; rewrite Z6; ring.
unfold FtoRradix; apply ClosestZero with b t (tx*ty)%R; auto with zarith.
rewrite Z4; ring.
unfold FtoRradix; apply ClosestZero with b t (t2-x2y1)%R; auto with zarith.
cut (FtoRradix t2=0)%R;[intros Z5|idtac].
cut (FtoRradix x2y1=0)%R;[intros Z6|idtac].
rewrite Z5; rewrite Z6; ring.
unfold FtoRradix; apply ClosestZero with b t (tx*hy)%R; auto with zarith.
rewrite Z3; ring.
unfold FtoRradix; apply ClosestZero with b t (t1-x1y2)%R; auto with zarith.
cut (FtoRradix t1=0)%R;[intros Z5|idtac].
cut (FtoRradix x1y2=0)%R;[intros Z6|idtac].
rewrite Z5; rewrite Z6; ring.
unfold FtoRradix; apply ClosestZero with b t (hx*ty)%R; auto with zarith.
rewrite Z4; ring.
unfold FtoRradix; apply ClosestZero with b t (r-x1y1)%R; auto with zarith.
cut (FtoRradix x1y1=0)%R;[intros Z6|idtac].
rewrite Z1; rewrite Z6; ring.
unfold FtoRradix; apply ClosestZero with b t (hx*hy)%R; auto with zarith.
rewrite Z3; ring.
elim VeltkampU with radix b s t y p' q' hy ty; auto.
intros T1 T; elim T; intros H' T'; clear T1 T T'.
fold FtoRradix in H'; rewrite Z3 in H'; rewrite <- Ny in H'; auto with real.
apply trans_eq with (0+ty)%R; auto with real.
unfold s; apply SLe; auto.
unfold s; apply SGe; auto.
right; auto.
elim Veltkamp with radix b s t y p' q' hy; auto with zarith arith.
2: apply SLe; auto.
2: apply SGe; auto.
2: elim Cy; auto.
intros T1 T; elim T; intros hy' T'; elim T'; intros G1 T''; elim T''; intros ; clear T1 T T' T''.
unfold FtoRradix; rewrite <- G1.
 apply ClosestZero with (Bound (P_of_succ_nat (pred (Z.abs_nat (Zpower_nat radix (t - s)))))
            (dExp b)) (t-s) (FtoR radix y)%R; auto with zarith.
apply p'GivesBound; auto with zarith.
apply SGe; auto.
cut (s <= t-2)%nat ;[idtac | apply SGe; auto].
unfold s; auto with zarith.
unfold FtoRradix; apply ClosestZero with b t (x*y)%R; auto with zarith.
rewrite <- Ny; ring.
elim bimplybplusNorm with radix b s t y; auto.
2: unfold s; apply SLe; auto.
2: unfold s; apply SGe; auto.
2: elim Cy; auto.
intros yy T; elim T; intros X1 X2; clear T.
rewrite <- X1.
assert (Fnormal radix (plusExp t b) x).
elim Cx; intros F1 F2; elim F1; intros.
split;[split|idtac]; unfold plusExp; simpl; auto with zarith.
assert (- dExp (plusExp t b) <= Fexp x + Fexp yy)%Z.
elim X2; intros F1 F2; elim F1; intros.
assert (0 <= Fexp x)%Z; auto with zarith.
apply Zplus_le_reg_l with (Fexp y).
rewrite (Zplus_comm (Fexp y) (Fexp x)); apply Z.le_trans with (2:=Expoxy).
elim Cy; intros F1' F2'; elim F2'; auto with zarith.
assert (Closest (plusExp t b) radix
   (FtoR radix x * (powerRZ radix (t - div2 t)%nat + 1)) p).
cut (FtoR radix x * (powerRZ radix (t - div2 t)%nat + 1) =
    (FtoRradix (Fmult x (Float (Zpower_nat radix (t - div2 t)%nat + 1) 0))))%R.
intros K'; rewrite K'.
unfold FtoRradix; apply Closestbbplus with 2 t; auto with zarith.
unfold Fmult; simpl.
assert (K:Fbounded b x);[elim Cx; auto|elim K; auto with zarith].
fold FtoRradix; rewrite <- K'; auto with real.
unfold FtoRradix; rewrite Fmult_correct; auto.
unfold FtoR; simpl; rewrite plus_IZR; rewrite Zpower_nat_Z_powerRZ; simpl;ring.
assert (Closest (plusExp t b) radix (FtoR radix x - FtoR radix p) q).
rewrite <- Fminus_correct; auto.
apply Closestbbplus with 2 t; auto with zarith.
unfold Fminus; simpl; apply Zmin_Zle.
assert (K:Fbounded b x);[elim Cx; auto|elim K; auto with zarith].
assert (K:Fbounded b p);[elim A1; auto|elim K; auto with zarith].
rewrite Fminus_correct; auto.
assert (Closest (plusExp t b) radix (FtoR radix q + FtoR radix p) hx).
rewrite <- Fplus_correct; auto.
apply Closestbbplus with 2 t; auto with zarith.
unfold Fplus; simpl; apply Zmin_Zle.
assert (K:Fbounded b q);[elim A2; auto|elim K; auto with zarith].
assert (K:Fbounded b p);[elim A1; auto|elim K; auto with zarith].
rewrite Fplus_correct; auto.
assert (Closest (plusExp t b) radix (FtoR radix x - FtoR radix hx) tx).
rewrite <- Fminus_correct; auto.
apply Closestbbplus with 2 t; auto with zarith.
unfold Fminus; simpl; apply Zmin_Zle.
assert (K:Fbounded b x);[elim Cx; auto|elim K; auto with zarith].
assert (K:Fbounded b hx);[elim A3; auto|elim K; auto with zarith].
rewrite Fminus_correct; auto.
assert (Closest (plusExp t b) radix (FtoR radix yy * (powerRZ radix (t - div2 t)%nat + 1)) p').
rewrite X1; cut (FtoR radix y * (powerRZ radix (t - div2 t)%nat + 1) =
    (FtoRradix (Fmult y (Float (Zpower_nat radix (t - div2 t)%nat + 1) 0))))%R.
intros K'; rewrite K'.
unfold FtoRradix; apply Closestbbplus with 2 t; auto with zarith.
unfold Fmult; simpl.
assert (K:Fbounded b y);[elim Cy; auto|elim K; auto with zarith].
fold FtoRradix; rewrite <- K'; auto with real.
unfold FtoRradix; rewrite Fmult_correct; auto.
unfold FtoR; simpl; rewrite plus_IZR; rewrite Zpower_nat_Z_powerRZ; simpl; ring.
assert (Closest (plusExp t b) radix (FtoR radix yy - FtoR radix p') q').
rewrite X1; rewrite <- Fminus_correct; auto.
apply Closestbbplus with 2 t; auto with zarith.
unfold Fminus; simpl; apply Zmin_Zle.
assert (K:Fbounded b y);[elim Cy; auto|elim K; auto with zarith].
assert (K:Fbounded b p');[elim B1; auto|elim K; auto with zarith].
rewrite Fminus_correct; auto.
assert (Closest (plusExp t b) radix (FtoR radix q' + FtoR radix p') hy).
rewrite <- Fplus_correct; auto.
apply Closestbbplus with 2 t; auto with zarith.
unfold Fplus; simpl; apply Zmin_Zle.
assert (K:Fbounded b q');[elim B2; auto|elim K; auto with zarith].
assert (K:Fbounded b p');[elim B1; auto|elim K; auto with zarith].
rewrite Fplus_correct; auto.
assert (Closest (plusExp t b) radix (FtoR radix yy - FtoR radix hy) ty).
rewrite X1; rewrite <- Fminus_correct; auto.
apply Closestbbplus with 2 t; auto with zarith.
unfold Fminus; simpl; apply Zmin_Zle.
assert (K:Fbounded b y);[elim Cy; auto|elim K; auto with zarith].
assert (K:Fbounded b hy);[elim B3; auto|elim K; auto with zarith].
rewrite Fminus_correct; auto.
generalize VeltkampU; intros V.
elim V with radix b s t x p q hx tx; auto.
2: unfold s; apply SLe; auto.
2: unfold s; apply SGe; auto.
2: left; auto.
intros M1 T; elim T; intros M2 T'; elim T'; intros T1 T2; clear T T'.
elim T1; intros hx' T1'; elim T1'; intros M3 T; elim T; intros M4 T'; clear T1 T1' T.
lapply T'; auto; intros M5; clear T'.
elim T2; intros tx' T1'; elim T1'; intros M6 T; elim T; intros M7 M8; clear T2 T1' T V.
elim Veltkamp_tail with radix b s t y p' q' hy ty; auto.
2: unfold s; apply SLe; auto.
2: unfold s; apply SGe; auto.
2: elim Cy; auto.
intros ty' T1'; elim T1'; intros N5 T; elim T; intros N7 T'; elim T'; intros N8 N9; clear T1' T T'.
rewrite FcanonicFnormalizeEq in N9; auto with zarith;[idtac|right; auto].
assert (Fexp y <= Fexp hy)%Z.
elim Cy; intros T1 T2; elim T2; intros T3 T4; rewrite T3.
elim B3; intros G1 G2; elim G1; auto.
apply DekkerN with (plusExp t b) t p q hx tx p' q' hy ty x1y1 x1y2 x2y1 x2y2 t1 t2 t3; auto.
rewrite <- M3.
rewrite <- Fmult_correct; auto.
apply Closestbbplus with 2 t; auto with zarith.
unfold Fmult; simpl; auto with zarith.
rewrite Fmult_correct; auto with real; rewrite M3; auto.
rewrite <- M3; rewrite <- N5.
rewrite <- Fmult_correct; auto.
apply Closestbbplus with 2 t; auto with zarith.
unfold Fmult; simpl; auto with zarith.
rewrite Fmult_correct; auto with real; rewrite M3; rewrite N5; auto.
rewrite <- M6.
rewrite <- Fmult_correct; auto.
apply Closestbbplus with 2 t; auto with zarith.
unfold Fmult; simpl; auto with zarith.
rewrite Fmult_correct; auto with real; rewrite M6; auto.
rewrite <- M6; rewrite <- N5.
rewrite <- Fmult_correct; auto.
apply Closestbbplus with 2 t; auto with zarith.
unfold Fmult; simpl; auto with zarith.
rewrite Fmult_correct; auto with real; rewrite M6; rewrite N5; auto.
rewrite X1.
rewrite <- Fmult_correct; auto.
apply Closestbbplus with 2 t; auto with zarith.
rewrite Fmult_correct; auto with real.
rewrite <- Fminus_correct; auto.
apply Closestbbplus with 2 t; auto with zarith.
unfold Fminus; simpl; apply Zmin_Zle.
assert (K:Fbounded b r);[elim D1; auto|elim K; auto with zarith].
assert (K:Fbounded b x1y1);[elim C1; auto|elim K; auto with zarith].
rewrite Fminus_correct; auto.
rewrite <- Fminus_correct; auto.
apply Closestbbplus with 2 t; auto with zarith.
unfold Fminus; simpl; apply Zmin_Zle.
assert (K:Fbounded b t1);[elim D2; auto|elim K; auto with zarith].
assert (K:Fbounded b x1y2);[elim C2; auto|elim K; auto with zarith].
rewrite Fminus_correct; auto.
rewrite <- Fminus_correct; auto.
apply Closestbbplus with 2 t; auto with zarith.
unfold Fminus; simpl; apply Zmin_Zle.
assert (K:Fbounded b t2);[elim D3; auto|elim K; auto with zarith].
assert (K:Fbounded b x2y1);[elim C3; auto|elim K; auto with zarith].
rewrite Fminus_correct; auto.
rewrite <- Fminus_correct; auto.
apply Closestbbplus with 2 t; auto with zarith.
unfold Fminus; simpl; apply Zmin_Zle.
assert (K:Fbounded b t3);[elim D4; auto|elim K; auto with zarith].
assert (K:Fbounded b x2y2);[elim C4; auto|elim K; auto with zarith].
rewrite Fminus_correct; auto.
Qed.

End AlgoS1.

Section AlgoS2.

Variable radix : Z.
Variable b : Fbound.
Variables t:nat.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.
Hypothesis radixMoreThanOne : (1 < radix)%Z.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ radixMoreThanOne).

Hypothesis pGivesBound: Zpos (vNum b)=(Zpower_nat radix t).
Hypotheses pGe: (4 <= t).

Variables x y p q hx tx p' q' hy ty x1y1 x1y2 x2y1 x2y2 r t1 t2 t3 t4:float.

Hypothesis Cx: (Fsubnormal radix b x).
Hypothesis Cy: (Fnormal radix b y).

Hypothesis Expoxy: (-dExp b <= Fexp x+Fexp y)%Z.

Let s:= t- div2 t.

Hypothesis A1: (Closest b radix (x*((powerRZ radix s)+1))%R p).
Hypothesis A2: (Closest b radix (x-p)%R q).
Hypothesis A3: (Closest b radix (q+p)%R hx).
Hypothesis A4: (Closest b radix (x-hx)%R tx).

Hypothesis B1: (Closest b radix (y*((powerRZ radix s)+1))%R p').
Hypothesis B2: (Closest b radix (y-p')%R q').
Hypothesis B3: (Closest b radix (q'+p')%R hy).
Hypothesis B4: (Closest b radix (y-hy)%R ty).

Hypothesis C1: (Closest b radix (hx*hy)%R x1y1).
Hypothesis C2: (Closest b radix (hx*ty)%R x1y2).
Hypothesis C3: (Closest b radix (tx*hy)%R x2y1).
Hypothesis C4: (Closest b radix (tx*ty)%R x2y2).

Hypothesis D1: (Closest b radix (x*y)%R r).
Hypothesis D2: (Closest b radix (r-x1y1)%R t1).
Hypothesis D3: (Closest b radix (t1-x1y2)%R t2).
Hypothesis D4: (Closest b radix (t2-x2y1)%R t3).
Hypothesis D5: (Closest b radix (t3-x2y2)%R t4).

Theorem DekkerS2: (radix=2)%Z \/ (even t) -> (x*y=r-t4)%R.
intros H; unfold FtoRradix.
case (Req_dec 0%R x); intros Ny.
cut (FtoRradix r=0)%R;[intros Z1|idtac].
cut (FtoRradix t4=0)%R;[intros Z2|idtac].
fold FtoRradix; rewrite Z1; rewrite Z2; rewrite <- Ny; ring.
cut (FtoRradix hx=0)%R;[intros Z3|idtac].
cut (FtoRradix tx=0)%R;[intros Z4|idtac].
unfold FtoRradix; apply ClosestZero with b t (t3-x2y2)%R; auto with zarith.
cut (FtoRradix t3=0)%R;[intros Z5|idtac].
cut (FtoRradix x2y2=0)%R;[intros Z6|idtac].
rewrite Z5; rewrite Z6; ring.
unfold FtoRradix; apply ClosestZero with b t (tx*ty)%R; auto with zarith.
rewrite Z4; ring.
unfold FtoRradix; apply ClosestZero with b t (t2-x2y1)%R; auto with zarith.
cut (FtoRradix t2=0)%R;[intros Z5|idtac].
cut (FtoRradix x2y1=0)%R;[intros Z6|idtac].
rewrite Z5; rewrite Z6; ring.
unfold FtoRradix; apply ClosestZero with b t (tx*hy)%R; auto with zarith.
rewrite Z4; ring.
unfold FtoRradix; apply ClosestZero with b t (t1-x1y2)%R; auto with zarith.
cut (FtoRradix t1=0)%R;[intros Z5|idtac].
cut (FtoRradix x1y2=0)%R;[intros Z6|idtac].
rewrite Z5; rewrite Z6; ring.
unfold FtoRradix; apply ClosestZero with b t (hx*ty)%R; auto with zarith.
rewrite Z3; ring.
unfold FtoRradix; apply ClosestZero with b t (r-x1y1)%R; auto with zarith.
cut (FtoRradix x1y1=0)%R;[intros Z6|idtac].
rewrite Z1; rewrite Z6; ring.
unfold FtoRradix; apply ClosestZero with b t (hx*hy)%R; auto with zarith.
rewrite Z3; ring.
elim VeltkampU with radix b s t x p q hx tx; auto.
intros T1 T; elim T; intros H' T'; clear T1 T T'.
fold FtoRradix in H'; rewrite Z3 in H'; rewrite <- Ny in H'; auto with real.
apply trans_eq with (0+tx)%R; auto with real.
unfold s; apply SLe; auto.
unfold s; apply SGe; auto.
right; auto.
elim Veltkamp with radix b s t x p q hx; auto.
2: unfold s; apply SLe; auto.
2: unfold s; apply SGe; auto.
2: elim Cx; auto.
intros T1 T; elim T; intros hy' T'; elim T'; intros G1 T''; elim T''; intros ; clear T1 T T' T''.
unfold FtoRradix; rewrite <- G1.
 apply ClosestZero with (Bound (P_of_succ_nat (pred (Z.abs_nat (Zpower_nat radix (t - s)))))
            (dExp b)) (t-s) (FtoR radix x)%R; auto with zarith.
apply p'GivesBound; auto with zarith.
apply SGe; auto.
cut (s <= t - 2)%nat; [unfold s; auto with zarith| apply SGe; auto].
unfold FtoRradix; apply ClosestZero with b t (x*y)%R; auto with zarith.
rewrite <- Ny; ring.
elim bimplybplusNorm with radix b s t x; auto.
2: unfold s; apply SLe; auto.
2: unfold s; apply SGe; auto.
2: elim Cx; auto.
intros xx T; elim T; intros X1 X2; clear T.
rewrite <- X1.
assert (Fnormal radix (plusExp t b) y).
elim Cy; intros F1 F2; elim F1; intros.
split;[split|idtac]; unfold plusExp; simpl; auto with zarith.
assert (- dExp (plusExp t b) <= Fexp xx + Fexp y)%Z.
elim X2; intros F1 F2; elim F1; intros.
assert (0 <= Fexp y)%Z; auto with zarith.
apply Zplus_le_reg_l with (Fexp x).
apply Z.le_trans with (2:=Expoxy).
elim Cx; intros F1' F2'; elim F2'; auto with zarith.
assert (Closest (plusExp t b) radix
   (FtoR radix y * (powerRZ radix (t - div2 t)%nat + 1)) p').
cut (FtoR radix y * (powerRZ radix (t - div2 t)%nat + 1) =
    (FtoRradix (Fmult y (Float (Zpower_nat radix (t - div2 t)%nat + 1) 0))))%R.
intros K'; rewrite K'.
unfold FtoRradix; apply Closestbbplus with 2 t; auto with zarith.
unfold Fmult; simpl.
assert (K:Fbounded b y);[elim Cy; auto|elim K; auto with zarith].
fold FtoRradix; rewrite <- K'; auto with real.
unfold FtoRradix; rewrite Fmult_correct; auto.
unfold FtoR; simpl; rewrite plus_IZR; rewrite Zpower_nat_Z_powerRZ; simpl; ring.
assert (Closest (plusExp t b) radix (FtoR radix y - FtoR radix p') q').
rewrite <- Fminus_correct; auto.
apply Closestbbplus with 2 t; auto with zarith.
unfold Fminus; simpl; apply Zmin_Zle.
assert (K:Fbounded b y);[elim Cy; auto|elim K; auto with zarith].
assert (K:Fbounded b p');[elim B1; auto|elim K; auto with zarith].
rewrite Fminus_correct; auto.
assert (Closest (plusExp t b) radix (FtoR radix q' + FtoR radix p') hy).
rewrite <- Fplus_correct; auto.
apply Closestbbplus with 2 t; auto with zarith.
unfold Fplus; simpl; apply Zmin_Zle.
assert (K:Fbounded b q');[elim B2; auto|elim K; auto with zarith].
assert (K:Fbounded b p');[elim B1; auto|elim K; auto with zarith].
rewrite Fplus_correct; auto.
assert (Closest (plusExp t b) radix (FtoR radix y - FtoR radix hy) ty).
rewrite <- Fminus_correct; auto.
apply Closestbbplus with 2 t; auto with zarith.
unfold Fminus; simpl; apply Zmin_Zle.
assert (K:Fbounded b y);[elim Cy; auto|elim K; auto with zarith].
assert (K:Fbounded b hy);[elim B3; auto|elim K; auto with zarith].
rewrite Fminus_correct; auto.
assert (Closest (plusExp t b) radix (FtoR radix xx * (powerRZ radix (t - div2 t)%nat + 1)) p).
rewrite X1; cut (FtoR radix x * (powerRZ radix (t - div2 t)%nat + 1) =
    (FtoRradix (Fmult x (Float (Zpower_nat radix (t - div2 t)%nat + 1) 0))))%R.
intros K'; rewrite K'.
unfold FtoRradix; apply Closestbbplus with 2 t; auto with zarith.
unfold Fmult; simpl.
assert (K:Fbounded b x);[elim Cx; auto|elim K; auto with zarith].
fold FtoRradix; rewrite <- K'; auto with real.
unfold FtoRradix; rewrite Fmult_correct; auto.
unfold FtoR; simpl; rewrite plus_IZR; rewrite Zpower_nat_Z_powerRZ; simpl; ring.
assert (Closest (plusExp t b) radix (FtoR radix xx - FtoR radix p) q).
rewrite X1; rewrite <- Fminus_correct; auto.
apply Closestbbplus with 2 t; auto with zarith.
unfold Fminus; simpl; apply Zmin_Zle.
assert (K:Fbounded b x);[elim Cx; auto|elim K; auto with zarith].
assert (K:Fbounded b p);[elim A1; auto|elim K; auto with zarith].
rewrite Fminus_correct; auto.
assert (Closest (plusExp t b) radix (FtoR radix q + FtoR radix p) hx).
rewrite <- Fplus_correct; auto.
apply Closestbbplus with 2 t; auto with zarith.
unfold Fplus; simpl; apply Zmin_Zle.
assert (K:Fbounded b q);[elim A2; auto|elim K; auto with zarith].
assert (K:Fbounded b p);[elim A1; auto|elim K; auto with zarith].
rewrite Fplus_correct; auto.
assert (Closest (plusExp t b) radix (FtoR radix xx - FtoR radix hx) tx).
rewrite X1; rewrite <- Fminus_correct; auto.
apply Closestbbplus with 2 t; auto with zarith.
unfold Fminus; simpl; apply Zmin_Zle.
assert (K:Fbounded b x);[elim Cx; auto|elim K; auto with zarith].
assert (K:Fbounded b hx);[elim A3; auto|elim K; auto with zarith].
rewrite Fminus_correct; auto.
generalize VeltkampU; intros V.
elim V with radix b s t y p' q' hy ty; auto.
2: unfold s; apply SLe; auto.
2: unfold s; apply SGe; auto.
2: left; auto.
intros M1 T; elim T; intros M2 T'; elim T'; intros T1 T2; clear T T'.
elim T1; intros hy' T1'; elim T1'; intros M3 T; elim T; intros M4 T'; clear T1 T1' T.
lapply T'; auto; intros M5; clear T'.
elim T2; intros ty' T1'; elim T1'; intros M6 T; elim T; intros M7 M8; clear T2 T1' T V.
elim Veltkamp_tail with radix b s t x p q hx tx; auto.
2: unfold s; apply SLe; auto.
2: unfold s; apply SGe; auto.
2: elim Cx; auto.
intros tx' T1'; elim T1'; intros N5 T; elim T; intros N7 T'; elim T'; intros N8 N9; clear T1' T T'.
rewrite FcanonicFnormalizeEq in N9; auto with zarith;[idtac|right; auto].
assert (Fexp x <= Fexp hx)%Z.
elim Cx; intros T1 T2; elim T2; intros T3 T4; rewrite T3.
elim A3; intros G1 G2; elim G1; auto.
apply DekkerN with (plusExp t b) t p q hx tx p' q' hy ty x1y1 x1y2 x2y1 x2y2 t1 t2 t3; auto.
rewrite <- M3.
rewrite <- Fmult_correct; auto.
apply Closestbbplus with 2 t; auto with zarith.
unfold Fmult; simpl; auto with zarith.
rewrite Fmult_correct; auto with real; rewrite M3; auto.
rewrite <- M6.
rewrite <- Fmult_correct; auto.
apply Closestbbplus with 2 t; auto with zarith.
unfold Fmult; simpl; auto with zarith.
rewrite Fmult_correct; auto with real; rewrite M6; auto.
rewrite <- M3; rewrite <- N5.
rewrite <- Fmult_correct; auto.
apply Closestbbplus with 2 t; auto with zarith.
unfold Fmult; simpl; auto with zarith.
rewrite Fmult_correct; auto with real; rewrite M3; rewrite N5; auto.
rewrite <- M6; rewrite <- N5.
rewrite <- Fmult_correct; auto.
apply Closestbbplus with 2 t; auto with zarith.
unfold Fmult; simpl; auto with zarith.
rewrite Fmult_correct; auto with real; rewrite M6; rewrite N5; auto.
rewrite X1.
rewrite <- Fmult_correct; auto.
apply Closestbbplus with 2 t; auto with zarith.
rewrite Fmult_correct; auto with real.
rewrite <- Fminus_correct; auto.
apply Closestbbplus with 2 t; auto with zarith.
unfold Fminus; simpl; apply Zmin_Zle.
assert (K:Fbounded b r);[elim D1; auto|elim K; auto with zarith].
assert (K:Fbounded b x1y1);[elim C1; auto|elim K; auto with zarith].
rewrite Fminus_correct; auto.
rewrite <- Fminus_correct; auto.
apply Closestbbplus with 2 t; auto with zarith.
unfold Fminus; simpl; apply Zmin_Zle.
assert (K:Fbounded b t1);[elim D2; auto|elim K; auto with zarith].
assert (K:Fbounded b x1y2);[elim C2; auto|elim K; auto with zarith].
rewrite Fminus_correct; auto.
rewrite <- Fminus_correct; auto.
apply Closestbbplus with 2 t; auto with zarith.
unfold Fminus; simpl; apply Zmin_Zle.
assert (K:Fbounded b t2);[elim D3; auto|elim K; auto with zarith].
assert (K:Fbounded b x2y1);[elim C3; auto|elim K; auto with zarith].
rewrite Fminus_correct; auto.
rewrite <- Fminus_correct; auto.
apply Closestbbplus with 2 t; auto with zarith.
unfold Fminus; simpl; apply Zmin_Zle.
assert (K:Fbounded b t3);[elim D4; auto|elim K; auto with zarith].
assert (K:Fbounded b x2y2);[elim C4; auto|elim K; auto with zarith].
rewrite Fminus_correct; auto.
Qed.

End AlgoS2.

Section Algo1.

Variable radix : Z.
Variable b : Fbound.
Variables t:nat.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.
Hypothesis radixMoreThanOne : (1 < radix)%Z.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ radixMoreThanOne).

Hypothesis pGivesBound: Zpos (vNum b)=(Zpower_nat radix t).
Hypotheses pGe: (4 <= t).

Variables x y p q hx tx p' q' hy ty x1y1 x1y2 x2y1 x2y2 r t1 t2 t3 t4:float.

Hypothesis Cx: (Fcanonic radix b x).
Hypothesis Cy: (Fcanonic radix b y).

Hypothesis Expoxy: (-dExp b <= Fexp x+Fexp y)%Z.

Let s:= t- div2 t.

Hypothesis A1: (Closest b radix (x*((powerRZ radix s)+1))%R p).
Hypothesis A2: (Closest b radix (x-p)%R q).
Hypothesis A3: (Closest b radix (q+p)%R hx).
Hypothesis A4: (Closest b radix (x-hx)%R tx).

Hypothesis B1: (Closest b radix (y*((powerRZ radix s)+1))%R p').
Hypothesis B2: (Closest b radix (y-p')%R q').
Hypothesis B3: (Closest b radix (q'+p')%R hy).
Hypothesis B4: (Closest b radix (y-hy)%R ty).

Hypothesis C1: (Closest b radix (hx*hy)%R x1y1).
Hypothesis C2: (Closest b radix (hx*ty)%R x1y2).
Hypothesis C3: (Closest b radix (tx*hy)%R x2y1).
Hypothesis C4: (Closest b radix (tx*ty)%R x2y2).

Hypothesis D1: (Closest b radix (x*y)%R r).
Hypothesis D2: (Closest b radix (r-x1y1)%R t1).
Hypothesis D3: (Closest b radix (t1-x1y2)%R t2).
Hypothesis D4: (Closest b radix (t2-x2y1)%R t3).
Hypothesis D5: (Closest b radix (t3-x2y2)%R t4).

Hypothesis dExpPos: ~(Z_of_N(dExp b)=0)%Z.

Theorem Dekker1: (radix=2)%Z \/ (even t) -> (x*y=r-t4)%R.
case Cy; case Cx; intros.
unfold FtoRradix; apply DekkerN with b t p q hx tx p' q' hy ty x1y1 x1y2 x2y1 x2y2 t1 t2 t3; auto.
unfold FtoRradix; apply DekkerS2 with b t p q hx tx p' q' hy ty x1y1 x1y2 x2y1 x2y2 t1 t2 t3; auto.
unfold FtoRradix; apply DekkerS1 with b t p q hx tx p' q' hy ty x1y1 x1y2 x2y1 x2y2 t1 t2 t3; auto.
absurd (- dExp b <= Fexp x + Fexp y)%Z; auto with zarith.
apply Zlt_not_le.
elim H; intros T1 T2; elim T2; intros G1 T; clear T1 T2 T.
elim H0; intros T1 T2; elim T2; intros G2 T; clear T1 T2 T.
rewrite G1; rewrite G2; auto with zarith.
Qed.

End Algo1.
Section Algo2.

Variable radix : Z.
Variable b : Fbound.
Variables t:nat.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.
Hypothesis radixMoreThanOne : (1 < radix)%Z.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ radixMoreThanOne).

Hypothesis pGivesBound: Zpos (vNum b)=(Zpower_nat radix t).
Hypotheses pGe: (4 <= t).
Let s:= t- div2 t.

Variables x y:float.

Let b' := Bound (vNum b) (Nplus (N.double (dExp b)) (N.double (Npos (P_of_succ_nat t)))).

Theorem Veltkampb': forall (f pf qf hf tf:float),
  (dExp b < dExp b')%Z ->
  (Fbounded b f) ->
  Closest b radix (f * (powerRZ radix s + 1)) pf -> Closest b radix (f - pf) qf ->
  Closest b radix (qf + pf) hf -> Closest b radix (f - hf) tf ->
  Closest b' radix (f * (powerRZ radix s + 1)) pf /\
  Closest b' radix (f - pf) qf /\ Closest b' radix (qf + pf) hf /\
  Closest b' radix (f - hf) tf.
intros.
split.
assert (f*(powerRZ radix s + 1)= (FtoRradix (Fplus radix (Fmult f (Float 1 s)) f)))%R.
unfold FtoRradix; rewrite Fplus_correct; auto; rewrite Fmult_correct; auto.
unfold FtoR; simpl; ring.
rewrite H5; unfold FtoRradix; apply Closestbbext with b t; auto with zarith.
simpl; rewrite Zmin_le2; auto with zarith; apply H0.
fold FtoRradix; rewrite <- H5; auto.
split.
unfold FtoRradix; rewrite <- Fminus_correct; auto.
apply Closestbbext with b t; auto with zarith.
simpl; apply Zmin_Zle; auto with zarith; try apply H0.
apply H1.
rewrite Fminus_correct; auto.
split.
unfold FtoRradix; rewrite <- Fplus_correct; auto.
apply Closestbbext with b t; auto with zarith.
simpl; apply Zmin_Zle; try apply H1; apply H2.
rewrite Fplus_correct; auto.
unfold FtoRradix; rewrite <- Fminus_correct; auto.
apply Closestbbext with b t; auto with zarith.
simpl; apply Zmin_Zle; try apply H3; apply H0.
rewrite Fminus_correct; auto.
Qed.

Variables p q hx tx p' q' hy ty x1y1 x1y2 x2y1 x2y2 r t1 t2 t3 t4:float.

Hypothesis Cx: (Fcanonic radix b x).
Hypothesis Cy: (Fcanonic radix b y).

Hypothesis Expoxy: (Fexp x+Fexp y < -dExp b)%Z.

Hypothesis A1: (Closest b radix (x*((powerRZ radix s)+1))%R p).
Hypothesis A2: (Closest b radix (x-p)%R q).
Hypothesis A3: (Closest b radix (q+p)%R hx).
Hypothesis A4: (Closest b radix (x-hx)%R tx).

Hypothesis B1: (Closest b radix (y*((powerRZ radix s)+1))%R p').
Hypothesis B2: (Closest b radix (y-p')%R q').
Hypothesis B3: (Closest b radix (q'+p')%R hy).
Hypothesis B4: (Closest b radix (y-hy)%R ty).

Hypothesis C1: (Closest b radix (hx*hy)%R x1y1).
Hypothesis C2: (Closest b radix (hx*ty)%R x1y2).
Hypothesis C3: (Closest b radix (tx*hy)%R x2y1).
Hypothesis C4: (Closest b radix (tx*ty)%R x2y2).

Hypothesis D1: (Closest b radix (x*y)%R r).
Hypothesis D2: (Closest b radix (r-x1y1)%R t1).
Hypothesis D3: (Closest b radix (t1-x1y2)%R t2).
Hypothesis D4: (Closest b radix (t2-x2y1)%R t3).
Hypothesis D5: (Closest b radix (t3-x2y2)%R t4).

Theorem dExpPrim: (dExp b < dExp b')%Z.
unfold b'; simpl; auto with zarith.
cut (forall (x:N) (y:positive), (x+(Zpos y)=(x +Npos y)%N)%Z).
intros T; simpl; rewrite <- T; auto with zarith.
apply Z.le_lt_trans with (N.double (dExp b)); auto with zarith.
unfold N.double; case (dExp b); auto with zarith.
intros; unfold Z_of_N; auto with zarith.
intros;unfold Nplus.
case x0; auto with zarith.
Qed.

Theorem dExpPrimEq: (Z_of_N (N.double (dExp b) + Npos (xO (P_of_succ_nat t)))
   =2*(dExp b)+2*t+2)%Z.
cut (forall (x:N) (y:positive), (x+(Zpos y)=(x +Npos y)%N)%Z).
intros T; rewrite <- T; auto with zarith.
2:intros;unfold Nplus.
2:case x0; auto with zarith.
replace (Zpos (xO (P_of_succ_nat t))) with (2*t+2)%Z.
unfold N.double; case (dExp b); auto with zarith.
apply trans_eq with (2*(Zpos (P_of_succ_nat t)))%Z; auto with zarith.
Qed.

Theorem NormalbPrim: forall (f:float), Fcanonic radix b f -> (FtoRradix f <>0) ->
   (exists f':float, (Fnormal radix b' f') /\ FtoRradix f'=f /\ (-t-dExp b <= Fexp f')%Z).
intros.
exists (Fnormalize radix b' t f).
assert (powerRZ radix (-(dExp b)) <= (Fabs (Fnormalize radix b' t f)))%R.
unfold FtoRradix; rewrite Fabs_correct; auto.
rewrite FnormalizeCorrect; auto with zarith; rewrite <- Fabs_correct; auto.
unfold FtoRradix, FtoR, Fabs; simpl.
apply Rle_trans with ((IZR 1)*powerRZ radix (- dExp b))%R;[right; simpl; ring|idtac].
apply Rmult_le_compat; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
case (Zle_lt_or_eq 0 (Z.abs (Fnum f))); auto with zarith real.
intros H1; apply Rle_IZR; omega.
intros; absurd (Rabs f =0)%R.
apply Rabs_no_R0; auto.
unfold FtoRradix; rewrite <- Fabs_correct; auto; unfold FtoR, Fabs; simpl; rewrite <- H1; simpl; ring.
assert (Fbounded b f);[apply FcanonicBound with radix; auto with zarith|idtac].
elim H1; intros; apply Rle_powerRZ; auto with zarith real.
apply Rle_IZR; omega.
assert (Fcanonic radix b' (Fnormalize radix b' t f)).
apply FnormalizeCanonic; auto with zarith.
assert (Fbounded b f);[apply FcanonicBound with radix; auto with zarith|idtac].
elim H2; generalize dExpPrim; intros; split; auto with zarith.
split.
case H2; auto.
intros; absurd (Fabs f < (firstNormalPos radix b' t))%R.
apply Rle_not_lt.
apply Rle_trans with (powerRZ radix (-(dExp b))).
unfold firstNormalPos, FtoRradix, FtoR; simpl.
unfold nNormMin; rewrite Zpower_nat_Z_powerRZ; rewrite <- powerRZ_add.
apply Rle_powerRZ; auto with real zarith.
apply Rle_IZR; omega.
rewrite dExpPrimEq.
rewrite inj_pred; auto with zarith; unfold Z.pred.
ring_simplify (t + -1 + - (2 * dExp b + 2 * t + 2))%Z; auto with zarith.
apply IZR_neq; omega.
apply Rle_trans with (1:=H1); unfold FtoRradix; repeat rewrite Fabs_correct; auto.
rewrite FnormalizeCorrect; auto with zarith real.
apply Rle_lt_trans with (Fabs (Fnormalize radix b' t f)).
unfold FtoRradix; repeat rewrite Fabs_correct; auto.
rewrite FnormalizeCorrect; auto with zarith real.
unfold FtoRradix; apply FsubnormalLtFirstNormalPos; auto with zarith.
apply FsubnormFabs; auto.
rewrite Fabs_correct; auto with real zarith; apply Rabs_pos.
split;[unfold FtoRradix; rewrite FnormalizeCorrect; auto with zarith|idtac].
apply Z.le_trans with (Fexp (Float (nNormMin radix t) (-t-dExp b))); auto with zarith.
apply Fcanonic_Rle_Zle with radix b' t; auto with zarith.
apply FcanonicNnormMin; auto with zarith.
unfold b'; simpl; rewrite dExpPrimEq; auto with zarith.
rewrite Rabs_right.
rewrite <- Fabs_correct; auto; fold FtoRradix; apply Rle_trans with (2:=H1).
unfold FtoRradix, FtoR, nNormMin; simpl; rewrite Zpower_nat_Z_powerRZ.
rewrite <- powerRZ_add; try apply IZR_neq; auto with real zarith; apply Rle_powerRZ; auto with zarith real.
apply Rle_IZR; omega.
apply Rle_ge; apply LeFnumZERO; auto with zarith.
unfold nNormMin; simpl; auto with zarith.
apply Zpower_NR0; omega.
Qed.

Theorem Dekker2_aux:
  (FtoRradix x <>0) -> (FtoRradix y <>0) ->
  (radix=2)%Z \/ (even t) -> (Rabs (x*y-(r-t4)) <= (7/2)*powerRZ radix (-(dExp b)))%R.
intros P1 P2.
intros; generalize dExpPrim; intros.
elim (NormalbPrim x); auto.
intros x' T; elim T; intros Nx' T'; elim T'; intros Hx' Ex'; clear T T'.
elim (NormalbPrim y); auto.
intros y' T; elim T; intros Ny' T'; elim T'; intros Hy' Ey'; clear T T'.
assert (MM:(-(dExp b') <= Fexp x'+Fexp y')%Z).
unfold b'; simpl; rewrite dExpPrimEq; auto with zarith.
generalize Underf_Err2; intros T.
elim T with b radix t b' (x*y)%R r; auto with zarith; clear T.
intros r' T; elim T; intros H1 H2; clear T.
elim Veltkampb' with x p q hx tx; auto.
2: apply FcanonicBound with radix; auto.
intros H4 T; elim T; intros H5 T'; elim T'; intros H6 H7; clear T T'.
elim Veltkampb' with y p' q' hy ty; auto.
2: apply FcanonicBound with radix; auto.
intros H8 T; elim T; intros H9 T'; elim T'; intros H10 H11; clear T T'.
assert (TotalP (Closest b' radix)).
apply ClosestTotal with t; auto with zarith.
unfold TotalP in H3.
elim (H3 (hx * hy)%R); intros x1y1' H12.
elim (H3 (hx * ty)%R); intros x1y2' H13.
elim (H3 (tx * hy)%R); intros x2y1' H14.
elim (H3 (tx * ty)%R); intros x2y2' H15.
elim (H3 (r' - x1y1')%R); intros t1' H16.
elim (H3 (t1' - x1y2')%R); intros t2' H17.
elim (H3 (t2' - x2y1')%R); intros t3' H18.
elim (H3 (t3' - x2y2')%R); intros t4' H19.
rewrite <- Hx'; rewrite <- Hy'; unfold FtoRradix.
rewrite DekkerN with radix b' t x' y' p q hx tx p' q' hy ty x1y1' x1y2' x2y1' x2y2' r' t1' t2' t3' t4';
  auto with zarith.
2: fold FtoRradix; rewrite Hx'; auto.
2: fold FtoRradix; rewrite Hx'; auto.
2: fold FtoRradix; rewrite Hx'; auto.
2: fold FtoRradix; rewrite Hy'; auto.
2: fold FtoRradix; rewrite Hy'; auto.
2: fold FtoRradix; rewrite Hy'; auto.
2: fold FtoRradix; rewrite Hx'; rewrite Hy'; auto.
fold FtoRradix.
replace (r' - t4' - (r - t4))%R with (-(r-r')+((t4-t4')))%R;[idtac|ring].
apply Rle_trans with (1:=Rabs_triang (-(r-r'))%R ((t4-t4'))%R).
apply Rle_trans with ((3/4)*powerRZ radix (- dExp b) +(11/4)*powerRZ radix (- dExp b))%R;
  [idtac|right; field; apply prod_neq_R0; auto with real; apply prod_neq_R0; auto with real].
apply Rplus_le_compat.
rewrite Rabs_Ropp; auto with real.
elim H1; intros G1 G2; elim G2; intros G3 G4; elim G4; intros G5 G6.
unfold FtoRradix; apply Rle_trans with (1:=G5); right; ring.
cut (2 <= s);[intros Sle|unfold s; apply SLe; auto].
cut (s <= t-2);[intros Sge|unfold s; apply SGe; auto].
cut (s+s <= t+1)%Z;[intros s2le|unfold s; apply s2Le; auto].
cut (t <=s+s)%Z;[intros s2ge|unfold s; apply s2Ge; auto].
generalize VeltkampU; intros V.
elim V with radix b' s t x' p q hx tx; auto.
2: left; auto.
2: fold FtoRradix; rewrite Hx'; auto.
2: fold FtoRradix; rewrite Hx'; auto.
2: fold FtoRradix; rewrite Hx'; auto.
intros MX1 T; elim T; intros MX2 T'; clear T; elim T'; intros T1 T2; clear T'.
elim T1; intros Chx' T1'; elim T1'; intros MX3 T1''; elim T1''; intros MX4 MX5; clear T1 T1' T1''.
lapply MX5; auto; clear MX5; intros MX5.
elim T2; intros Ctx' T1'; elim T1'; intros MX6 T1''; elim T1''; intros MX7 MX8; clear T2 T1' T1''.
elim V with radix b' s t y' p' q' hy ty; auto.
2: left; auto.
2: fold FtoRradix; rewrite Hy'; auto.
2: fold FtoRradix; rewrite Hy'; auto.
2: fold FtoRradix; rewrite Hy'; auto.
intros MY1 T; elim T; intros MY2 T'; clear T; elim T'; intros T1 T2; clear T'.
elim T1; intros Chy' T1'; elim T1'; intros MY3 T1''; elim T1''; intros MY4 MY5; clear T1 T1' T1''.
lapply MY5; auto; clear MY5; intros MY5.
elim T2; intros Cty' T1'; elim T1'; intros MY6 T1''; elim T1''; intros MY7 MY8; clear T2 T1' T1'' V.
generalize Boundedt1; intros V.
elim V with radix b' s t x' Chx' Ctx' y' Chy' Cty' r' (Fminus radix (Fmult x' y') r');
   auto with zarith real; clear V.
2: fold FtoRradix; rewrite Hy';rewrite Hx'; auto.
2:rewrite Fminus_correct; auto with zarith; rewrite Fmult_correct; auto with zarith; ring.
2:rewrite MX6; rewrite MX3; exact MX2.
2:rewrite MY6; rewrite MY3; exact MY2.
2:rewrite MX6; replace (FtoR radix tx) with (FtoR radix x'-FtoR radix hx)%R; auto with real.
2:rewrite MX2; ring.
2:rewrite MY6; replace (FtoR radix ty) with (FtoR radix y'-FtoR radix hy)%R; auto with real.
2:rewrite MY2; ring.
intros Ct1' T; elim T; intros M11 T'; elim T'; intros M12 M13; clear T T'.
generalize Boundedt2; intros V.
elim V with radix b' s t x' Chx' Ctx' y' Chy' Cty' r' (Fminus radix (Fmult x' y') r');
   auto with zarith real; clear V.
2: fold FtoRradix; rewrite Hy';rewrite Hx'; auto.
2:rewrite Fminus_correct; auto with zarith; rewrite Fmult_correct; auto with zarith; ring.
2:rewrite MX6; rewrite MX3; exact MX2.
2:rewrite MY6; rewrite MY3; exact MY2.
2:rewrite MX6; replace (FtoR radix tx) with (FtoR radix x'-FtoR radix hx)%R; auto with real.
2:rewrite MX2; ring.
2:rewrite MY6; replace (FtoR radix ty) with (FtoR radix y'-FtoR radix hy)%R; auto with real.
2:rewrite MY2; ring.
intros Ct2' T; elim T; intros M21 T'; elim T'; intros M22 M23; clear T T'.
generalize Boundedt3; intros V.
elim V with radix b' s t x' Chx' Ctx' y' Chy' Cty' r' (Fminus radix (Fmult x' y') r');
    auto with zarith real; clear V.
2: fold FtoRradix; rewrite Hy';rewrite Hx'; auto.
2:rewrite Fminus_correct; auto with zarith; rewrite Fmult_correct; auto with zarith; ring.
2:rewrite MX6; rewrite MX3; exact MX2.
2:rewrite MY6; rewrite MY3; exact MY2.
2:rewrite MX6; replace (FtoR radix tx) with (FtoR radix x'-FtoR radix hx)%R; auto with real.
2:rewrite MX2; ring.
2:rewrite MY6; replace (FtoR radix ty) with (FtoR radix y'-FtoR radix hy)%R; auto with real.
2:rewrite MY2; ring.
intros Ct3' T; elim T; intros M31 T'; elim T'; intros M32 M33; clear T T'.
generalize Boundedt4_aux; intros V.
elim V with radix b' s t x' Chx' Ctx' y' Chy' Cty' r' ; auto with zarith real; clear V.
2: fold FtoRradix; rewrite Hy';rewrite Hx'; auto.
2:rewrite MX6; rewrite MX3; exact MX2.
2:rewrite MY6; rewrite MY3; exact MY2.
intros Ct4' T; elim T; intros M41 T'; elim T'; intros M42 M43; clear T T'.
elim Boundedx1y1_aux with radix b' s t x' Chx' y' Chy'; auto with zarith.
intros Cx1y1' T; elim T; intros O1 T'; elim T'; intros O2 O3 ; clear T T'.
elim Boundedx1y2_aux with radix b' s t x' Chx' y' Cty'; auto with zarith.
intros Cx1y2' T; elim T; intros O4 T'; elim T'; intros O5 O6; clear T T'.
elim Boundedx2y1_aux with radix b' s t x' Ctx' y' Chy'; auto with zarith.
intros Cx2y1' T; elim T; intros O7 T'; elim T'; intros O8 O9; clear T T'.
assert (tmp:forall (f:float) (i:nat), (i <= t) ->
       (Fbounded (Bound (P_of_succ_nat (pred (Z.abs_nat (Zpower_nat radix (i)))))
             (dExp b')) f) -> (Fbounded b' f)).
intros f i J1 J2; elim J2; intros J3 J4; split; auto with zarith.
apply Z.lt_le_trans with (1:=J3).
apply Z.le_trans with (Zpower_nat radix i);[idtac|unfold b'; simpl; rewrite pGivesBound; auto with zarith].
simpl.
apply
 Z.le_trans
  with
    (Z_of_nat
       (nat_of_P
          (P_of_succ_nat
             (pred (Z.abs_nat (Zpower_nat radix (i))))))).
unfold Z_of_nat in |- *; rewrite nat_of_P_o_P_of_succ_nat_eq_succ;
 auto with zarith.
rewrite nat_of_P_o_P_of_succ_nat_eq_succ; auto with arith zarith.
rewrite <- S_pred with (Z.abs_nat (Zpower_nat radix (i))) 0; auto with zarith.
rewrite <- Zabs_absolu; rewrite Z.abs_eq; auto with zarith.
apply Zpower_NR0; auto with zarith.
cut ( 0 < Z.abs_nat (Zpower_nat radix (i)))%Z; auto with zarith.
simpl; rewrite <- Zabs_absolu; rewrite Z.abs_eq; auto with zarith.
apply Zpower_nat_less; omega.
apply Zpower_NR0; omega.
apply Zpower_nat_monotone_le; omega.
elim Boundedx2y2 with radix b' t x' y' p q Chx' Ctx' p' q' Chy' Cty'; auto with zarith.
2: fold FtoRradix; rewrite Hx'; auto.
2: fold FtoRradix; rewrite Hx'; auto.
2:apply ClosestCompatible with (1:=H6); auto.
2:apply tmp with (t-s); auto with zarith.
2:apply ClosestCompatible with (1:=H7); auto with real.
2: fold FtoRradix; rewrite Hx'; auto with real.
2:apply tmp with s; auto with zarith.
2: fold FtoRradix; rewrite Hy'; auto.
2: fold FtoRradix; rewrite Hy'; auto.
2:apply ClosestCompatible with (1:=H10); auto.
2:apply tmp with (t-s); auto with zarith.
2:fold FtoRradix; rewrite Hy'; apply ClosestCompatible with (1:=H11); auto with real.
2:apply tmp with s; auto with zarith.
intros Cx2y2' T; elim T; intros O10 T'; elim T'; intros O11 O12; clear T T' tmp.
assert (ZZ:RoundedModeP b' radix (Closest b' radix)).
apply ClosestRoundedModeP with t; auto with zarith.
assert (K1':FtoRradix x1y1'=Cx1y1').
unfold FtoRradix; apply sym_eq; apply RoundedModeProjectorIdemEq with b' t (Closest b' radix); auto with zarith.
rewrite O1; rewrite MY3; rewrite MX3; auto.
assert (K1:FtoRradix Ct1'=t1').
unfold FtoRradix; apply RoundedModeProjectorIdemEq with b' t (Closest b' radix); auto with zarith.
rewrite M11; replace (FtoR radix Chx' * FtoR radix Chy')%R with (FtoRradix x1y1'); auto.
rewrite <- O1; auto.
assert (K2':FtoRradix x1y2'=Cx1y2').
unfold FtoRradix; apply sym_eq; apply RoundedModeProjectorIdemEq with b' t (Closest b' radix); auto with zarith.
rewrite O4; rewrite MY6; rewrite MX3; auto.
assert (K2:FtoRradix Ct2'=t2').
unfold FtoRradix; apply RoundedModeProjectorIdemEq with b' t (Closest b' radix); auto with zarith.
rewrite M21; rewrite <- M11; fold FtoRradix; rewrite K1.
replace (Chx' * Cty')%R with (FtoRradix x1y2'); auto.
unfold FtoRradix; rewrite <- O4; auto.
assert (K3':FtoRradix x2y1'=Cx2y1').
unfold FtoRradix; apply sym_eq; apply RoundedModeProjectorIdemEq with b' t (Closest b' radix); auto with zarith.
rewrite O7; rewrite MY3; rewrite MX6; auto.
assert (K3:FtoRradix Ct3'=t3').
unfold FtoRradix; apply RoundedModeProjectorIdemEq with b' t (Closest b' radix); auto with zarith.
rewrite M31; rewrite <- M21; fold FtoRradix; rewrite K2.
replace (Ctx' * Chy')%R with (FtoRradix x2y1'); auto.
unfold FtoRradix; rewrite <- O7; auto.
assert (K4':FtoRradix x2y2'=Cx2y2').
unfold FtoRradix; apply sym_eq; apply RoundedModeProjectorIdemEq with b' t (Closest b' radix); auto with zarith.
rewrite O10; rewrite MX6; rewrite MY6; auto.
assert (K4:FtoRradix Ct4'=t4').
unfold FtoRradix; apply RoundedModeProjectorIdemEq with b' t (Closest b' radix); auto with zarith.
rewrite M41; rewrite <- M31; fold FtoRradix; rewrite K3.
replace (Ctx' * Cty')%R with (FtoRradix x2y2'); auto.
unfold FtoRradix; rewrite <- O10; auto.
rewrite <- K4.
cut (Underf_Err b radix b' t4 Ct4' (t3-x2y2)%R (11/4)).
unfold FtoRradix; intros G; elim G; intros G1 G2; elim G2; intros G3 G4; elim G4; auto with real.
replace (11/4)%R with (9/4+/2)%R;
   [idtac|field; apply prod_neq_R0; auto with real; apply prod_neq_R0; auto with real].
unfold FtoRradix; apply Underf_Err3_bis with t Ct3' Cx2y2' (t2-x2y1)%R (tx*ty)%R; auto with zarith.
replace (9/4)%R with (7/4+/2)%R;[idtac|
   field; apply prod_neq_R0; auto with real; apply prod_neq_R0; auto with real].
unfold FtoRradix; apply Underf_Err3_bis with t Ct2' Cx2y1' (t1-x1y2)%R (tx*hy)%R; auto with zarith.
replace (7/4)%R with (5/4+/2)%R;[idtac|
   field; apply prod_neq_R0; auto with real; apply prod_neq_R0; auto with real].
unfold FtoRradix; apply Underf_Err3_bis with t Ct1' Cx1y2' (r-x1y1)%R (hx*ty)%R; auto with zarith.
replace (5/4)%R with ((3/4)+/2)%R;[idtac|
   field; apply prod_neq_R0; auto with real; apply prod_neq_R0; auto with real].
unfold FtoRradix; apply Underf_Err3_bis with t r' Cx1y1' (x*y)%R (hx*hy)%R; auto with zarith.
cut (hx*hy=FtoRradix Cx1y1')%R.
intros P; rewrite P; unfold FtoRradix; apply Underf_Err1 with t; auto with zarith.
fold FtoRradix; rewrite <- P; auto.
rewrite <- K1'; unfold FtoRradix; rewrite <- MX3; rewrite <- MY3; rewrite <- O1; auto.
apply Rmult_le_reg_l with (IZR 4); auto with real zarith; simpl.
apply Rle_trans with (IZR 5);[simpl; right; field; auto with real|idtac].
repeat apply prod_neq_R0; auto with real.
apply Rle_trans with (IZR 28); [auto with real zarith|simpl; right; ring].
rewrite M11; rewrite <- O1; auto with real.
rewrite M13; apply rExp with radix b' s; auto.
fold FtoRradix; rewrite Hx'; rewrite Hy'; auto.
cut (hx*ty=FtoRradix Cx1y2')%R.
intros P; rewrite P; unfold FtoRradix; apply Underf_Err1 with t; auto with zarith.
fold FtoRradix; rewrite <- P; auto.
unfold FtoRradix; rewrite O4; rewrite MX3; rewrite MY6; auto with real.
apply Rmult_le_reg_l with (IZR 4); auto with real zarith; simpl.
apply Rle_trans with (IZR 7);[simpl; right; field; auto with real|idtac].
repeat apply prod_neq_R0; auto with real.
apply Rle_trans with (IZR 28); [auto with real zarith|simpl; right; ring].
rewrite M21; rewrite M11; rewrite O4; ring.
cut (tx*hy=FtoRradix Cx2y1')%R.
intros P; rewrite P; unfold FtoRradix; apply Underf_Err1 with t; auto with zarith.
fold FtoRradix; rewrite <- P; auto.
unfold FtoRradix; rewrite O7; rewrite MX6; rewrite MY3; auto with real.
apply Rmult_le_reg_l with (IZR 4); auto with real zarith; simpl.
apply Rle_trans with (IZR 9);[simpl; right; field; auto with real|idtac].
repeat apply prod_neq_R0; auto with real.
apply Rle_trans with (IZR 28); [auto with real zarith|simpl; right; ring].
rewrite M21; rewrite M31; rewrite O7; ring.
cut (tx*ty=FtoRradix Cx2y2')%R.
intros P; rewrite P; unfold FtoRradix; apply Underf_Err1 with t; auto with zarith.
fold FtoRradix; rewrite <- P; auto.
unfold FtoRradix; rewrite O10; rewrite MX6; rewrite MY6; auto with real.
apply Rmult_le_reg_l with (IZR 4); auto with real zarith; simpl.
apply Rle_trans with (IZR 11);[simpl; right; field; auto with real|idtac].
repeat apply prod_neq_R0; auto with real.
apply Rle_trans with (IZR 28); [auto with real zarith|simpl; right; ring].
rewrite M41; rewrite M31; rewrite O10; ring.
Qed.

Theorem Dekker2:
  (radix=2)%Z \/ (even t) -> (Rabs (x*y-(r-t4)) <= (7/2)*powerRZ radix (-(dExp b)))%R.
intros.
case (Req_dec 0%R x); intros Ny.
cut (FtoRradix r=0)%R;[intros Z1|idtac].
cut (FtoRradix t4=0)%R;[intros Z2|idtac].
replace ((x * y - (r - t4)))%R with 0%R.
rewrite Rabs_R0; apply Rlt_le; apply Rmult_lt_0_compat; auto with real zarith.
unfold Rdiv; apply Rmult_lt_0_compat; auto with real zarith.
apply powerRZ_lt, Rlt_IZR; omega.
fold FtoRradix; rewrite Z1; rewrite Z2; rewrite <- Ny; ring.
cut (FtoRradix hx=0)%R;[intros Z3|idtac].
cut (FtoRradix tx=0)%R;[intros Z4|idtac].
unfold FtoRradix; apply ClosestZero with b t (t3-x2y2)%R; auto with zarith.
cut (FtoRradix t3=0)%R;[intros Z5|idtac].
cut (FtoRradix x2y2=0)%R;[intros Z6|idtac].
rewrite Z5; rewrite Z6; ring.
unfold FtoRradix; apply ClosestZero with b t (tx*ty)%R; auto with zarith.
rewrite Z4; ring.
unfold FtoRradix; apply ClosestZero with b t (t2-x2y1)%R; auto with zarith.
cut (FtoRradix t2=0)%R;[intros Z5|idtac].
cut (FtoRradix x2y1=0)%R;[intros Z6|idtac].
rewrite Z5; rewrite Z6; ring.
unfold FtoRradix; apply ClosestZero with b t (tx*hy)%R; auto with zarith.
rewrite Z4; ring.
unfold FtoRradix; apply ClosestZero with b t (t1-x1y2)%R; auto with zarith.
cut (FtoRradix t1=0)%R;[intros Z5|idtac].
cut (FtoRradix x1y2=0)%R;[intros Z6|idtac].
rewrite Z5; rewrite Z6; ring.
unfold FtoRradix; apply ClosestZero with b t (hx*ty)%R; auto with zarith.
rewrite Z3; ring.
unfold FtoRradix; apply ClosestZero with b t (r-x1y1)%R; auto with zarith.
cut (FtoRradix x1y1=0)%R;[intros Z6|idtac].
rewrite Z1; rewrite Z6; ring.
unfold FtoRradix; apply ClosestZero with b t (hx*hy)%R; auto with zarith.
rewrite Z3; ring.
elim VeltkampU with radix b s t x p q hx tx; auto.
intros T1 T; elim T; intros H' T'; clear T1 T T'.
fold FtoRradix in H'; rewrite Z3 in H'; rewrite <- Ny in H'; auto with real.
apply trans_eq with (0+tx)%R; auto with real.
unfold s; apply SLe; auto.
unfold s; apply SGe; auto.
elim Veltkamp with radix b s t x p q hx; auto.
2: apply SLe; auto.
2: apply SGe; auto.
2: apply FcanonicBound with radix; auto.
intros T1 T; elim T; intros hy' T'; elim T'; intros G1 T''; elim T''; intros ; clear T1 T T' T''.
unfold FtoRradix; rewrite <- G1.
apply ClosestZero with (Bound (P_of_succ_nat (pred (Z.abs_nat (Zpower_nat radix (t - s)))))
            (dExp b)) (t-s) (FtoR radix x)%R; auto with zarith.
apply p'GivesBound; auto with zarith.
apply SGe; auto.
cut (s <= t-2)%nat; [unfold s; auto with zarith | apply SGe; auto].
unfold FtoRradix; apply ClosestZero with b t (x*y)%R; auto with zarith.
rewrite <- Ny; ring.
case (Req_dec 0%R y); intros Nx.
cut (FtoRradix r=0)%R;[intros Z1|idtac].
cut (FtoRradix t4=0)%R;[intros Z2|idtac].
replace ((x * y - (r - t4)))%R with 0%R.
rewrite Rabs_R0; apply Rlt_le; apply Rmult_lt_0_compat; auto with real zarith.
unfold Rdiv; apply Rmult_lt_0_compat; auto with real zarith.
apply powerRZ_lt, Rlt_IZR; omega.
fold FtoRradix; rewrite Z1; rewrite Z2; rewrite <- Nx; ring.
cut (FtoRradix hy=0)%R;[intros Z3|idtac].
cut (FtoRradix ty=0)%R;[intros Z4|idtac].
unfold FtoRradix; apply ClosestZero with b t (t3-x2y2)%R; auto with zarith.
cut (FtoRradix t3=0)%R;[intros Z5|idtac].
cut (FtoRradix x2y2=0)%R;[intros Z6|idtac].
rewrite Z5; rewrite Z6; ring.
unfold FtoRradix; apply ClosestZero with b t (tx*ty)%R; auto with zarith.
rewrite Z4; ring.
unfold FtoRradix; apply ClosestZero with b t (t2-x2y1)%R; auto with zarith.
cut (FtoRradix t2=0)%R;[intros Z5|idtac].
cut (FtoRradix x2y1=0)%R;[intros Z6|idtac].
rewrite Z5; rewrite Z6; ring.
unfold FtoRradix; apply ClosestZero with b t (tx*hy)%R; auto with zarith.
rewrite Z3; ring.
unfold FtoRradix; apply ClosestZero with b t (t1-x1y2)%R; auto with zarith.
cut (FtoRradix t1=0)%R;[intros Z5|idtac].
cut (FtoRradix x1y2=0)%R;[intros Z6|idtac].
rewrite Z5; rewrite Z6; ring.
unfold FtoRradix; apply ClosestZero with b t (hx*ty)%R; auto with zarith.
rewrite Z4; ring.
unfold FtoRradix; apply ClosestZero with b t (r-x1y1)%R; auto with zarith.
cut (FtoRradix x1y1=0)%R;[intros Z6|idtac].
rewrite Z1; rewrite Z6; ring.
unfold FtoRradix; apply ClosestZero with b t (hx*hy)%R; auto with zarith.
rewrite Z3; ring.
elim VeltkampU with radix b s t y p' q' hy ty; auto.
intros T1 T; elim T; intros H' T'; clear T1 T T'.
fold FtoRradix in H'; rewrite Z3 in H'; rewrite <- Nx in H'; auto with real.
apply trans_eq with (0+ty)%R; auto with real.
unfold s; apply SLe; auto.
unfold s; apply SGe; auto.
elim Veltkamp with radix b s t y p' q' hy; auto.
2: apply SLe; auto.
2: apply SGe; auto.
2: apply FcanonicBound with radix; auto.
intros T1 T; elim T; intros hy' T'; elim T'; intros G1 T''; elim T''; intros ; clear T1 T T' T''.
unfold FtoRradix; rewrite <- G1.
apply ClosestZero with (Bound (P_of_succ_nat (pred (Z.abs_nat (Zpower_nat radix (t - s)))))
            (dExp b)) (t-s) (FtoR radix y)%R; auto with zarith.
apply p'GivesBound; auto with zarith.
apply SGe; auto.
cut (s <= t-2)%nat; [unfold s; auto with zarith | apply SGe; auto].
unfold FtoRradix; apply ClosestZero with b t (x*y)%R; auto with zarith.
rewrite <- Nx; ring.
apply Dekker2_aux; auto.
Qed.

End Algo2.

Section AlgoT.

Variable radix : Z.
Variable b : Fbound.
Variables t:nat.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.
Hypothesis radixMoreThanOne : (1 < radix)%Z.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ radixMoreThanOne).

Hypothesis pGivesBound: Zpos (vNum b)=(Zpower_nat radix t).
Hypotheses pGe: (4 <= t).

Variables x y p q hx tx p' q' hy ty x1y1 x1y2 x2y1 x2y2 r t1 t2 t3 t4:float.

Hypothesis Cx: (Fcanonic radix b x).
Hypothesis Cy: (Fcanonic radix b y).

Let s:= t- div2 t.

Hypothesis A1: (Closest b radix (x*((powerRZ radix s)+1))%R p).
Hypothesis A2: (Closest b radix (x-p)%R q).
Hypothesis A3: (Closest b radix (q+p)%R hx).
Hypothesis A4: (Closest b radix (x-hx)%R tx).

Hypothesis B1: (Closest b radix (y*((powerRZ radix s)+1))%R p').
Hypothesis B2: (Closest b radix (y-p')%R q').
Hypothesis B3: (Closest b radix (q'+p')%R hy).
Hypothesis B4: (Closest b radix (y-hy)%R ty).

Hypothesis C1: (Closest b radix (hx*hy)%R x1y1).
Hypothesis C2: (Closest b radix (hx*ty)%R x1y2).
Hypothesis C3: (Closest b radix (tx*hy)%R x2y1).
Hypothesis C4: (Closest b radix (tx*ty)%R x2y2).

Hypothesis D1: (Closest b radix (x*y)%R r).
Hypothesis D2: (Closest b radix (-r+x1y1)%R t1).
Hypothesis D3: (Closest b radix (t1+x1y2)%R t2).
Hypothesis D4: (Closest b radix (t2+x2y1)%R t3).
Hypothesis D5: (Closest b radix (t3+x2y2)%R t4).

Hypothesis dExpPos: ~(Z_of_N (dExp b)=0)%Z.

Theorem Dekker: (radix=2)%Z \/ (even t) ->
  ((-dExp b <= Fexp x+Fexp y)%Z -> (x*y=r+t4)%R) /\
    (Rabs (x*y-(r+t4)) <= (7/2)*powerRZ radix (-(dExp b)))%R.
intros.
case (Zle_or_lt (-dExp b) (Fexp x+Fexp y)); intros.
cut (x * y = r + t4)%R; [intros; split; auto|idtac].
rewrite H1; ring_simplify ( (r + t4) - (r + t4))%R; rewrite Rabs_R0.
apply Rlt_le; apply Rmult_lt_0_compat; auto with real zarith.
unfold Rdiv; apply Rmult_lt_0_compat; auto with real zarith.
apply powerRZ_lt, Rlt_IZR; omega.
apply trans_eq with (r-(Fopp t4))%R;[idtac|unfold FtoRradix; rewrite Fopp_correct; ring].
unfold FtoRradix; apply Dekker1 with b t p q hx tx p' q' hy ty x1y1 x1y2 x2y1 x2y2
  (Fopp t1) (Fopp t2) (Fopp t3); auto; try rewrite Fopp_correct; fold FtoRradix.
replace (r-x1y1)%R with (-(-r+x1y1))%R;[apply ClosestOpp; auto|ring].
replace (-t1-x1y2)%R with (-(t1+x1y2))%R;[apply ClosestOpp; auto|ring].
replace (-t2-x2y1)%R with (-(t2+x2y1))%R;[apply ClosestOpp; auto|ring].
replace (-t3-x2y2)%R with (-(t3+x2y2))%R;[apply ClosestOpp; auto|ring].
split.
intros; absurd (Fexp x + Fexp y < - dExp b)%Z; auto with zarith.
replace (r+t4)%R with (r-(Fopp t4))%R;[idtac|unfold FtoRradix; rewrite Fopp_correct; ring].
unfold FtoRradix; apply Dekker2 with t p q hx tx p' q' hy ty x1y1 x1y2 x2y1 x2y2
  (Fopp t1) (Fopp t2) (Fopp t3); auto; try rewrite Fopp_correct; fold FtoRradix.
replace (r-x1y1)%R with (-(-r+x1y1))%R;[apply ClosestOpp; auto|ring].
replace (-t1-x1y2)%R with (-(t1+x1y2))%R;[apply ClosestOpp; auto|ring].
replace (-t2-x2y1)%R with (-(t2+x2y1))%R;[apply ClosestOpp; auto|ring].
replace (-t3-x2y2)%R with (-(t3+x2y2))%R;[apply ClosestOpp; auto|ring].
Qed.

End AlgoT.

This proof file has been written by Sylvie Boldo(1), following a proof presented by Pr William Kahan (2), and adapted to Coq proof checker with the help of Guillaume Melquiond(1) and Marc Daumas(1). This work has been partially supported by the CNRS grant PICS 2533.
(1) LIP Computer science laboratory UMR 5668 CNRS - ENS de Lyon - INRIA Lyon, France
(2) University of California at Berkeley Berkeley, California

Section Discriminant1.
Variable bo : Fbound.
Variable precision : nat.

Let radix := 2%Z.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ TwoMoreThanOne).

Hypothesis precisionGreaterThanOne : 1 < precision.
Hypothesis pGivesBound : Zpos (vNum bo) = Zpower_nat radix precision.

Variables a b b' c p q d:float.

Let delta := (Rabs (d-(b*b'-a*c)))%R.

Hypothesis Fa : (Fbounded bo a).
Hypothesis Fb : (Fbounded bo b).
Hypothesis Fb': (Fbounded bo b').
Hypothesis Fc : (Fbounded bo c).
Hypothesis Fp : (Fbounded bo p).
Hypothesis Fq : (Fbounded bo q).
Hypothesis Fd : (Fbounded bo d).

There is no underflow
Hypothesis U1:(- dExp bo <= Fexp d - 1)%Z.
Hypothesis Nd:(Fnormal radix bo d).
Hypothesis Nq:(Fnormal radix bo q).
Hypothesis Np:(Fnormal radix bo p).

Hypothesis Square:(0 <=b*b')%R.

Hypothesis Roundp : (EvenClosest bo radix precision (b*b')%R p).
Hypothesis Roundq : (EvenClosest bo radix precision (a*c)%R q).

Hypothesis Firstcase : (p+q <= 3*(Rabs (p-q)))%R.
Hypothesis Roundd : (EvenClosest bo radix precision (p-q)%R d).

Theorem delta_inf: (delta <= (/2)*(Fulp bo radix precision d)+
   ((/2)*(Fulp bo radix precision p)+(/2)*(Fulp bo radix precision q)))%R.
unfold delta; rewrite <- Rabs_Ropp.
replace (-(d - (b * b' - a * c)))%R with (((p-q)-d)+((b*b'-p)+-(a*c-q)))%R;[idtac|ring].
apply Rle_trans with ((Rabs ((p-q)-d))+(Rabs (b * b' - p + - (a * c - q))))%R;
  [apply Rabs_triang|idtac].
apply Rplus_le_compat.
apply Rmult_le_reg_l with (S (S O)); auto with arith real.
apply Rle_trans with (Fulp bo radix precision d).
unfold FtoRradix; apply ClosestUlp;auto with zarith.
elim Roundd; auto.
right; simpl; field; auto with real.
apply Rle_trans with ((Rabs (b*b'-p))+(Rabs (-(a*c-q))))%R;
  [apply Rabs_triang|idtac].
apply Rplus_le_compat.
apply Rmult_le_reg_l with (S (S O)); auto with arith real.
apply Rle_trans with (Fulp bo radix precision p).
unfold FtoRradix; apply ClosestUlp;auto with zarith.
elim Roundp; auto.
right; simpl; field; auto with real.
rewrite Rabs_Ropp; apply Rmult_le_reg_l with (S (S O)); auto with arith real.
apply Rle_trans with (Fulp bo radix precision q).
unfold FtoRradix; apply ClosestUlp;auto with zarith.
elim Roundq; auto.
right; simpl; field; auto with real.
Qed.

Theorem P_positive: (Rle 0 p)%R.
unfold FtoRradix; apply RleRoundedR0 with (b:=bo) (precision:=precision) (P:=(Closest bo radix)) (r:=(b*b')%R); auto.
apply ClosestRoundedModeP with precision; auto.
elim Roundp; auto.
Qed.

Theorem Fulp_le_twice_l: forall x y:float, (0 <= x)%R ->
   (Fnormal radix bo x) -> (Fbounded bo y) -> (2*x<=y)%R ->
   (2*(Fulp bo radix precision x) <= (Fulp bo radix precision y))%R.
intros.
assert (2*x=(Float (Fnum x) (Z.succ (Fexp x))))%R.
unfold FtoRradix, FtoR, Z.succ; simpl; rewrite powerRZ_add; auto with real zarith.
rewrite powerRZ_1; unfold radix; ring.
apply Rle_trans with (Fulp bo radix precision (Float (Fnum x) (Z.succ (Fexp x)))).
right; rewrite CanonicFulp; auto; [rewrite CanonicFulp|left]; auto.
unfold FtoR, Z.succ; simpl; rewrite powerRZ_add; auto with real zarith.
rewrite powerRZ_1; unfold radix; ring.
elim H0; intros H4 H5; elim H4; intros.
left; split; auto.
split; simpl; auto with zarith.
apply LeFulpPos; auto with real.
elim H0; intros H4 H5; elim H4; intros;split; simpl; auto with zarith.
fold FtoRradix; rewrite <- H3; apply Rmult_le_pos; auto with real.
fold FtoRradix; rewrite <- H3; auto with real.
Qed.

Theorem Fulp_le_twice_r: forall x y:float, (0 <= x)%R ->
   (Fnormal radix bo y) -> (Fbounded bo x) -> (x<=2*y)%R ->
   ((Fulp bo radix precision x) <= 2*(Fulp bo radix precision y))%R.
intros.
assert (2*y=(Float (Fnum y) (Z.succ (Fexp y))))%R.
unfold FtoRradix, FtoR, Z.succ; simpl; rewrite powerRZ_add; auto with real zarith.
rewrite powerRZ_1; unfold radix; ring.
apply Rle_trans with (Fulp bo radix precision (Float (Fnum y) (Z.succ (Fexp y)))).
2:right; rewrite CanonicFulp; auto; [rewrite CanonicFulp|left]; auto.
2:unfold FtoR, Z.succ; simpl; rewrite powerRZ_add; auto with real zarith.
2: rewrite powerRZ_1; unfold radix ; ring.
2:left; auto.
2:elim H0; intros H6 H5; elim H6; intros.
2:split; auto with zarith.
2:split; simpl; auto with zarith.
apply LeFulpPos; auto with real.
elim H0; intros H6 H5; elim H6; intros;split; simpl; auto with zarith.
fold FtoRradix; rewrite <- H3; auto with real.
Qed.

Theorem Half_Closest_Round: forall (x:float) (r:R),
   (- dExp bo <= Z.pred (Fexp x))%Z -> (Closest bo radix r x)
  -> (Closest bo radix (r/2)%R (Float (Fnum x) (Z.pred (Fexp x)))).
intros x r L H.
assert (x/2=(Float (Fnum x) (Z.pred (Fexp x))))%R.
unfold FtoRradix, FtoR, Z.pred; simpl; rewrite powerRZ_add; auto with real zarith.
unfold powerRZ at 3; simpl; unfold radix; field.
elim H; intros H2 H3.
split; [split; simpl; auto with zarith; apply H2|idtac].
intros.
fold FtoRradix; rewrite <- H0.
replace (x/2-r/2)%R with (/2*(x-r))%R;[idtac|unfold Rdiv; ring].
rewrite Rabs_mult; rewrite Rabs_right; auto with real.
2: apply Rle_ge; auto with real.
replace (f-r/2)%R with (/2*((Float (Fnum f) (Z.succ (Fexp f)))-r))%R.
rewrite Rabs_mult; rewrite Rabs_right with (/2)%R.
2: apply Rle_ge; auto with real.
apply Rmult_le_compat_l; auto with real.
unfold FtoRradix; apply H3.
destruct H1; split; simpl; auto with zarith.
unfold FtoRradix, FtoR, Z.succ; simpl; rewrite powerRZ_add; auto with real zarith.
rewrite powerRZ_1; unfold radix; field.
Qed.

Theorem Twice_EvenClosest_Round: forall (x:float) (r:R),
   (-(dExp bo) <= (Fexp x)-1)%Z -> (Fnormal radix bo x)
  -> (EvenClosest bo radix precision r x)
  -> (EvenClosest bo radix precision (2*r)%R (Float (Fnum x) (Z.succ (Fexp x)))).
intros x r U Nx H.
assert (x*2=(Float (Fnum x) (Z.succ (Fexp x))))%R.
unfold FtoRradix, FtoR, Z.succ; simpl; rewrite powerRZ_add; auto with real zarith.
rewrite powerRZ_1; unfold radix; ring.
elim H; intros H2 H3; elim H2; intros H'1 H'2; split.
split; [split; simpl; auto with zarith; apply H'1|idtac].
intros.
fold FtoRradix; rewrite <- H0.
replace (x*2-2*r)%R with (2*(x-r))%R;[idtac|unfold Rdiv; ring].
rewrite Rabs_mult; rewrite Rabs_right; auto with real.
case (Zle_lt_or_eq (-(dExp bo))%Z (Fexp f)); try apply H1; intros L.
replace (f-2*r)%R with (2*((Float (Fnum f) (Z.pred (Fexp f)))-r))%R.
rewrite Rabs_mult; rewrite Rabs_right with (2)%R.
2: apply Rle_ge; auto with real.
apply Rmult_le_compat_l; auto with real.
unfold FtoRradix; apply H'2.
split; simpl; auto with zarith; apply H1.
unfold FtoRradix, FtoR, Z.pred; simpl; rewrite powerRZ_add; auto with real zarith.
unfold powerRZ at 2; simpl; unfold radix; field.
replace (f-2*r)%R with (-((2*r)-f))%R;[rewrite Rabs_Ropp|ring].
apply Rle_trans with (2:=Rabs_triang_inv (2*r)%R f).
rewrite Rabs_mult; rewrite (Rabs_right 2%R); try apply Rle_ge;auto with real.
pattern r at 2 in |-*; replace r with (x-(x-r))%R;[idtac|ring].
apply Rle_trans with (2*(Rabs (x)-Rabs (x-r))-Rabs f)%R;[idtac|unfold Rminus; apply Rplus_le_compat_r; apply Rmult_le_compat_l; auto with real].
2: generalize (Rabs_triang_inv x (x-r)%R); unfold Rminus; auto with real.
apply Rplus_le_reg_l with (Rabs f -2*(Rabs (x-r)))%R.
apply Rle_trans with (Rabs f);[right;ring|idtac].
apply Rle_trans with (2*(Rabs x)-4*Rabs (x-r))%R;[idtac|right;ring].
apply Rle_trans with (((powerRZ radix precision)-1)*(powerRZ radix ((Fexp x)-1)))%R.
unfold FtoRradix; rewrite <- Fabs_correct; auto;unfold Fabs, FtoR; simpl.
apply Rmult_le_compat; auto with real zarith.
apply Rle_IZR; auto with zarith.
apply powerRZ_le, IZR_lt; auto with zarith.
apply Rle_trans with (Z.pred (Zpower_nat radix precision));[rewrite <- pGivesBound|idtac].
apply Rle_IZR;apply Zle_Zpred; auto with zarith; apply H1.
unfold Z.pred; rewrite plus_IZR; rewrite Zpower_nat_Z_powerRZ; auto with real zarith.
rewrite <- L; apply Rle_powerRZ; auto with real zarith.
apply Rle_trans with (2*(powerRZ radix (Z.pred precision))*(powerRZ radix (Fexp x))-2*(powerRZ radix (Fexp x)))%R.
apply Rle_trans with (((powerRZ radix (precision+1))-4)*(powerRZ radix (Fexp x-1)))%R;[apply Rmult_le_compat_r; auto with real zarith|idtac].
apply powerRZ_le, IZR_lt; auto with zarith.
rewrite powerRZ_add; auto with real zarith; simpl.
apply Rplus_le_reg_l with (-(powerRZ radix precision)+4)%R.
unfold radix at 5; ring_simplify.
apply Rle_trans with (powerRZ radix 2)%R; auto with real zarith.
simpl; unfold radix; ring_simplify (2*(2*1))%R; auto with real zarith.
apply Rle_powerRZ; auto with real zarith.
replace (2*powerRZ radix (Fexp x))%R with (4*powerRZ radix (Fexp x -1))%R.
replace 2%R with (powerRZ radix 1%Z);[idtac|apply powerRZ_1].
repeat rewrite <- powerRZ_add; auto with real zarith.
replace (1+Z.pred precision+Fexp x)%Z with ((precision+1)+(Fexp x-1))%Z;[idtac|unfold Z.pred; ring].
rewrite powerRZ_add with (n:=(precision+1)%Z);auto with real zarith; right;ring.
unfold Zminus; rewrite powerRZ_add; auto with real zarith.
unfold powerRZ at 2; simpl; unfold radix; field.
unfold Rminus; apply Rplus_le_compat;[rewrite Rmult_assoc; apply Rmult_le_compat_l; auto with real|apply Ropp_le_contravar].
unfold FtoRradix; rewrite <- Fabs_correct; auto; unfold FtoR, Fabs; simpl.
apply Rmult_le_compat_r; auto with real zarith.
apply powerRZ_le, IZR_lt; auto with zarith.
apply Rmult_le_reg_l with radix; auto with real zarith.
pattern (IZR radix) at 1 in |-*; replace (IZR radix) with (powerRZ radix 1%Z);[idtac|simpl; ring].
rewrite <- powerRZ_add; auto with real zarith; elim Nx; intros.
replace (1+Z.pred precision)%Z with (Z_of_nat precision)%Z;[idtac|unfold Z.pred; ring].
apply Rle_trans with (IZR (Zpos (vNum bo)));[rewrite pGivesBound; rewrite Zpower_nat_Z_powerRZ; auto with real zarith|idtac].
apply Rle_trans with (IZR (Z.abs (radix * Fnum x))); auto with real zarith.
apply Rle_IZR; auto.
rewrite Zabs_Zmult; rewrite Z.abs_eq; auto with real zarith.
rewrite mult_IZR; auto with real.
unfold radix; auto with zarith.
replace 4%R with (2*2%nat)%R; [rewrite Rmult_assoc; apply Rmult_le_compat_l; auto with real|simpl;ring].
replace (x+-r)%R with (-(r-x))%R;[rewrite Rabs_Ropp|ring].
apply Rle_trans with (Fulp bo radix precision x).
unfold FtoRradix; apply ClosestUlp; auto.
rewrite CanonicFulp; auto with real zarith.
right; unfold FtoR, radix; simpl; ring.
left; auto.
case H3; intros V.
left; generalize V; unfold FNeven; rewrite FcanonicFnormalizeEq; auto with zarith.
rewrite FcanonicFnormalizeEq; auto with zarith.
elim Nx; intros; left; split; auto with zarith.
elim H1; intros; split; simpl; auto with zarith.
left; auto.
right; intros.
apply trans_eq with (2*(FtoR radix (Float (Fnum q0) (Z.pred (Fexp q0)))))%R.
unfold FtoR, Z.pred; simpl; rewrite powerRZ_add; auto with real zarith.
unfold powerRZ at 3; simpl; unfold radix; field.
apply trans_eq with (2*(FtoR radix x))%R.
2: unfold FtoR, Z.succ; simpl; rewrite powerRZ_add; auto with real zarith.
2: rewrite powerRZ_1; unfold radix; ring.
apply Rmult_eq_compat_l; apply V.
replace r with ((2*r)/2)%R;[idtac|field; auto with real].
apply Half_Closest_Round; auto.
apply Z.le_trans with (1:=U).
fold (Z.pred (Fexp x)); cut (Fexp x <= Fexp q0)%Z; auto with zarith.
apply Z.le_trans with (Fexp (Fnormalize radix bo precision q0)).
apply Fcanonic_Rle_Zle with radix bo precision; auto with zarith.
left; auto.
apply FnormalizeCanonic; auto with zarith.
elim H1; auto.
generalize ClosestMonotone; unfold MonotoneP; intros.
repeat rewrite <- Fabs_correct; auto.
apply H4 with bo (Rabs r) (Rabs (2*r))%R.
rewrite Rabs_mult; rewrite (Rabs_right 2%R); try apply Rle_ge; auto with real.
apply Rle_lt_trans with (1*Rabs r)%R;[right;ring|apply Rmult_lt_compat_r; auto with real].
apply Rabs_pos_lt;unfold not;intros.
absurd (is_Fzero x).
apply FnormalNotZero with radix bo ; auto.
apply is_Fzero_rep2 with radix; auto.
cut (0 <= FtoR radix x)%R; intros.
cut (FtoR radix x <= 0)%R; intros; auto with real.
apply RleRoundedLessR0 with bo precision (Closest bo radix) r; auto with real.
apply ClosestRoundedModeP with precision; auto.
apply RleRoundedR0 with bo precision (Closest bo radix) r; auto with real.
apply ClosestRoundedModeP with precision; auto.
apply ClosestFabs with precision; auto.
apply ClosestFabs with precision; auto.
generalize ClosestCompatible; unfold CompatibleP; intros T.
apply T with (2*r)%R q0; auto with real zarith.
apply sym_eq;apply FnormalizeCorrect; auto.
apply FnormalizeBounded; auto with zarith.
elim H1; auto.
apply FcanonicLeastExp with radix bo precision; auto with zarith.
apply sym_eq; apply FnormalizeCorrect; auto.
elim H1; auto.
apply FnormalizeCanonic; auto with zarith;elim H1; auto.
Qed.

Theorem EvenClosestMonotone2: forall (p q : R) (p' q' : float),
  (p <= q)%R -> (EvenClosest bo radix precision p p') ->
  (EvenClosest bo radix precision q q') -> (p' <= q')%R.
intros.
case H; intros H2.
generalize EvenClosestMonotone; unfold MonotoneP.
intros W; unfold FtoRradix.
apply W with bo precision p0 q0; auto.
generalize EvenClosestUniqueP; unfold UniqueP.
intros W; unfold FtoRradix.
right; apply W with bo precision p0; auto with real.
rewrite H2; auto.
Qed.

Theorem Fulp_le_twice_r_round: forall (x y:float) (r:R), (0 <= x)%R ->
   (Fbounded bo x) -> (Fnormal radix bo y) -> (- dExp bo <= Fexp y - 1)%Z
     -> (x<=2*r)%R ->
   (EvenClosest bo radix precision r y) ->
   ((Fulp bo radix precision x) <= 2*(Fulp bo radix precision y))%R.
intros x y r H H0 H1 U H2 H3.
assert (2*y=(Float (Fnum y) (Z.succ (Fexp y))))%R.
unfold FtoRradix, FtoR, Z.succ; simpl; rewrite powerRZ_add; auto with real zarith.
rewrite powerRZ_1; unfold radix; ring.
apply Rle_trans with (Fulp bo radix precision (Float (Fnum y) (Z.succ (Fexp y)))).
2:right; rewrite CanonicFulp; auto; [rewrite CanonicFulp|left]; auto.
2:unfold FtoR, Z.succ; simpl; rewrite powerRZ_add; auto with real zarith.
2:rewrite powerRZ_1; unfold radix; ring.
2:left; auto.
2:elim H1; intros H6 H5; elim H6; intros.
2:split; simpl; auto with zarith.
2:split; simpl; auto with zarith.
apply LeFulpPos; auto with real.
elim H1; intros H6 H5; elim H6; intros;split; simpl; auto with zarith.
apply EvenClosestMonotone2 with x (2*r)%R; auto.
unfold FtoRradix; apply RoundedModeProjectorIdem with (b:=bo) (P:=(EvenClosest bo radix precision)); auto.
apply EvenClosestRoundedModeP; auto.
apply Twice_EvenClosest_Round; auto.
Qed.

Theorem discri1: (delta <= 2*(Fulp bo radix precision d))%R.
apply Rle_trans with (1:=delta_inf).
case (Rle_or_lt q p); intros H1.
case (Rle_or_lt 0%R q); intros H2.
cut (2*(Fulp bo radix precision q)<=(Fulp bo radix precision p))%R; try intros H3.
cut ((Fulp bo radix precision p)<=2*(Fulp bo radix precision d))%R; try intros H4.
apply Rle_trans with ((/ 2 * Fulp bo radix precision d +
    (/ 2 * (2*Fulp bo radix precision d) + / 2 * Fulp bo radix precision d)))%R.
apply Rplus_le_compat; auto with real.
apply Rplus_le_compat; auto with real.
apply Rmult_le_compat_l; auto with real.
apply Rmult_le_reg_l with 2%R; auto with real.
apply Rle_trans with (1:=H3); auto with real.
right; field; auto with real.
apply Fulp_le_twice_r_round with (p-q)%R; auto.
apply P_positive.
apply Rplus_le_reg_l with (2*q-p)%R.
ring_simplify.
apply Rmult_le_reg_l with 2%R; auto with real.
apply Rplus_le_reg_l with (p-3*q)%R.
ring_simplify.
apply Rle_trans with (1:=Firstcase); rewrite Rabs_right.
right; ring.
apply Rle_ge; apply Rplus_le_reg_l with q; ring_simplify; auto with real.
apply Fulp_le_twice_l; auto.
apply Rmult_le_reg_l with 2%R; auto with real.
apply Rplus_le_reg_l with (p-3*q)%R.
ring_simplify.
apply Rle_trans with (1:=Firstcase); rewrite Rabs_right.
right; ring.
apply Rle_ge; apply Rplus_le_reg_l with q; ring_simplify; auto with real.
apply Rle_trans with ((/ 2 * Fulp bo radix precision d +
    (/ 2 * (Fulp bo radix precision d) + / 2 * Fulp bo radix precision d)))%R.
apply Rplus_le_compat; auto with real.
apply Rplus_le_compat; auto with real.
apply Rmult_le_compat; auto with real.
unfold Fulp; auto with real zarith.
apply powerRZ_le, IZR_lt; auto with zarith.
apply LeFulpPos; auto with real.
fold FtoRradix; apply P_positive.
fold FtoRradix; apply EvenClosestMonotone2 with p (p-q)%R; auto.
apply Rle_trans with (p-0)%R; unfold Rminus; auto with real; right;ring.
unfold FtoRradix; apply RoundedModeProjectorIdem with (b:=bo) (P:=(EvenClosest bo radix precision)); auto.
apply EvenClosestRoundedModeP; auto.
apply Rmult_le_compat; auto with real.
unfold Fulp; auto with real zarith.
apply powerRZ_le, IZR_lt; auto with zarith.
rewrite FulpFabs; auto.
apply LeFulpPos; auto with real.
apply absFBounded; easy.
rewrite Fabs_correct; auto with real; apply Rabs_pos.
fold FtoRradix; apply EvenClosestMonotone2 with (-q)%R (p-q)%R; auto.
generalize P_positive; intros; auto with real.
apply Rle_trans with (0-q)%R; unfold Rminus; auto with real; right;ring.
replace (-q)%R with (FtoRradix (Fabs q)).
unfold FtoRradix; apply RoundedModeProjectorIdem with (b:=bo) (P:=(EvenClosest bo radix precision)); auto.
apply EvenClosestRoundedModeP; auto.
apply absFBounded; easy.
unfold FtoRradix;rewrite Fabs_correct; auto with real; rewrite Rabs_left; auto with real.
apply Rle_trans with ((3*/2)*(Fulp bo radix precision d))%R.
right; field; auto with real.
apply Rmult_le_compat_r;auto with zarith real.
unfold Fulp; auto with zarith real.
apply powerRZ_le, IZR_lt; auto with zarith.
apply Rmult_le_reg_l with 2%R;auto with real.
apply Rle_trans with 3%R; auto with real.
right; field; auto with real.
rewrite <- mult_IZR; auto with real zarith.
cut (2*(Fulp bo radix precision p)<=(Fulp bo radix precision q))%R; try intros H3.
cut ((Fulp bo radix precision q)<=2*(Fulp bo radix precision d))%R; try intros H4.
apply Rle_trans with ((/ 2 * Fulp bo radix precision d +
    (/ 2 * (Fulp bo radix precision d) + / 2 * (2*Fulp bo radix precision d))))%R.
apply Rplus_le_compat; auto with real.
apply Rplus_le_compat; auto with real.
apply Rmult_le_compat_l; auto with real.
apply Rmult_le_reg_l with 2%R; auto with real.
apply Rle_trans with (1:=H3); auto with real.
right; field; auto with real.
assert (p-q <=0)%R.
apply Rplus_le_reg_l with q.
ring_simplify; auto with real.
rewrite FulpFabs with bo radix precision d; auto.
apply Fulp_le_twice_r_round with (Rabs (p-q))%R; auto.
apply Rle_trans with p; auto with real; apply P_positive.
apply FnormalFabs; auto.
rewrite Rabs_left; auto with real.
apply Rmult_le_reg_l with 2%R; auto with real.
apply Rplus_le_reg_l with (p-q)%R.
ring_simplify.
apply Rle_trans with (1:=Firstcase); rewrite Rabs_left1; auto.
right; ring.
generalize EvenClosestSymmetric; unfold SymmetricP; intros.
rewrite Rabs_left1; auto with real.
replace (Fabs d) with (Fopp d).
apply H0; auto.
unfold Fabs, Fopp; replace (Z.abs (Fnum d)) with (-(Fnum d))%Z; auto.
rewrite <- Zabs_Zopp; rewrite Z.abs_eq; auto with zarith.
cut (Fnum d <= 0)%Z; auto with zarith.
apply R0LeFnum with radix; auto.
apply RleRoundedLessR0 with bo precision (EvenClosest bo radix precision) (p-q)%R; auto with real zarith.
apply EvenClosestRoundedModeP; auto.
apply Fulp_le_twice_l; auto.
apply P_positive.
apply Rmult_le_reg_l with 2%R; auto with real.
apply Rplus_le_reg_l with (-3*p+q)%R.
ring_simplify.
apply Rle_trans with (1:=Firstcase); rewrite Rabs_left1; auto.
right; ring.
apply Rplus_le_reg_l with q.
ring_simplify; auto with real.
Qed.

End Discriminant1.

Section Discriminant2.
Variable bo : Fbound.
Variable precision : nat.

Let radix := 2%Z.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ TwoMoreThanOne).

Hypothesis precisionGreaterThanOne : 1 < precision.
Hypothesis pGivesBound : Zpos (vNum bo) = Zpower_nat radix precision.

Variables a b b' c p q t dp dq s d:float.

Let delta := (Rabs (d-(b*b'-a*c)))%R.

Hypothesis Fa : (Fbounded bo a).
Hypothesis Fb : (Fbounded bo b).
Hypothesis Fb': (Fbounded bo b').
Hypothesis Fc : (Fbounded bo c).
Hypothesis Fp : (Fbounded bo p).
Hypothesis Fq : (Fbounded bo q).
Hypothesis Fd : (Fbounded bo d).
Hypothesis Ft : (Fbounded bo t).
Hypothesis Fs : (Fbounded bo s).
Hypothesis Fdp: (Fbounded bo dp).
Hypothesis Fdq: (Fbounded bo dq).
Hypothesis Cs:(Fcanonic radix bo s).

There is no underflow
Hypothesis U1: (- dExp bo <= (Fexp t)-1)%Z.
Hypothesis U2: (powerRZ radix (-dExp bo+2*precision-1) <= Rabs (b*b'))%R.
Hypothesis U3: (powerRZ radix (-dExp bo+2*precision-1) <= Rabs (a*c))%R.

Hypothesis Np:(Fnormal radix bo p).
Hypothesis Nq:(Fnormal radix bo q).
Hypothesis Nd:(Fnormal radix bo d).

Hypothesis Square:(0 <=b*b')%R.

Hypothesis Roundp : (EvenClosest bo radix precision (b*b')%R p).
Hypothesis Roundq : (EvenClosest bo radix precision (a*c)%R q).

Hypothesis Secondcase : (3*(Rabs (p-q)) < p+q)%R.

Hypothesis Roundt : (EvenClosest bo radix precision (p-q)%R t).
Hypothesis dpEq : (FtoRradix dp=b*b'-p)%R.
Hypothesis dqEq : (FtoRradix dq=a*c-q)%R.
Hypothesis Rounds : (EvenClosest bo radix precision (dp-dq)%R s).
Hypothesis Roundd : (EvenClosest bo radix precision (t+s)%R d).

Hypothesis p_differ_q:~(p=q)%R.

Theorem Q_positive: (0 < q)%R.
case (Rle_or_lt q 0%R); auto; intros.
absurd (3*(Rabs (p-q)) < (Rabs (p-q)))%R.
apply Rle_not_lt; apply Rle_trans with (1*(Rabs (p-q)))%R; auto with real.
apply Rmult_le_compat_r; auto with real.
apply Rabs_pos.
apply Rlt_le_trans with (1:=Secondcase).
apply Rle_trans with (2:=Rabs_triang_inv p q).
right; rewrite Rabs_right.
rewrite Rabs_left1; auto with real; ring.
apply Rle_ge; apply P_positive with bo precision b b'; auto.
Qed.

Theorem Q_le_two_P: (q <= 2*p)%R.
fold FtoRradix; apply Rmult_le_reg_l with 2%R; auto with real; simpl.
apply Rle_trans with (3*q-q)%R; [right; ring|idtac].
pattern (FtoRradix q) at 1; rewrite <- (Rabs_right q).
2: apply Rle_ge; generalize Q_positive; auto with real.
pattern (FtoRradix q) at 1; replace (FtoRradix q) with (-(p-q)+p)%R;[idtac|ring].
apply Rle_trans with (3*(Rabs (-(p-q))+(Rabs p))-q)%R.
unfold Rminus; apply Rplus_le_compat_r.
apply Rmult_le_compat_l; auto with real.
apply Rabs_triang.
rewrite Rabs_Ropp.
rewrite (Rabs_right p).
2:apply Rle_ge; apply P_positive with bo precision b b'; auto.
apply Rle_trans with (3 * (Rabs (p - q)) + (3*p - q))%R;[right;ring|idtac].
apply Rle_trans with (p+q+(3*p-q))%R; auto with real.
right;ring.
Qed.

Theorem P_le_two_Q: (p <= 2*q)%R.
fold FtoRradix; apply Rmult_le_reg_l with 2%R; auto with real; simpl.
apply Rle_trans with (3*p-p)%R; [right; ring|idtac].
pattern (FtoRradix p) at 1; rewrite <- (Rabs_right p).
2: apply Rle_ge; apply P_positive with bo precision b b'; auto with real.
pattern (FtoRradix p) at 1; replace (FtoRradix p) with ((p-q)+q)%R;[idtac|ring].
apply Rle_trans with (3*(Rabs (p-q)+(Rabs q))-p)%R.
unfold Rminus; apply Rplus_le_compat_r.
apply Rmult_le_compat_l; auto with real.
apply Rabs_triang.
rewrite (Rabs_right q).
2: apply Rle_ge; generalize Q_positive; auto with real.
apply Rle_trans with (3 * (Rabs (p - q)) + (3*q - p))%R;[right;ring|idtac].
apply Rle_trans with (p+q+(3*q-p))%R; auto with real.
right;ring.
Qed.

Theorem t_exact: (FtoRradix t=p-q)%R.
unfold FtoRradix; rewrite <- Fminus_correct; auto with zarith.
apply sym_eq; apply RoundedModeProjectorIdemEq with (b:=bo) (P:=(EvenClosest bo radix precision)) (precision:=precision); auto.
apply EvenClosestRoundedModeP; auto.
2: rewrite Fminus_correct; auto with zarith.
apply Sterbenz; auto.
fold FtoRradix; apply Rmult_le_reg_l with 2%R; auto with real.
apply Rle_trans with (FtoRradix q);[simpl; right; field; auto with real|idtac].
apply Q_le_two_P.
fold FtoRradix; simpl; apply P_le_two_Q.
Qed.

Theorem dp_dq_le: (Rabs (dp-dq) <= (3/2)*(Rmin
    (Fulp bo radix precision p) (Fulp bo radix precision q)))%R.
unfold Rminus; apply Rle_trans with (1:=Rabs_triang dp (-dq)).
rewrite Rabs_Ropp;apply Rmult_le_reg_l with (S (S O))%R; auto with real.
apply Rle_trans with (S 1 * Rabs dp + S 1*Rabs dq)%R;[right;ring|idtac].
apply Rle_trans with ((Fulp bo radix precision p)+(Fulp bo radix precision q))%R.
apply Rplus_le_compat.
rewrite dpEq; unfold FtoRradix; apply ClosestUlp; auto.
elim Roundp; auto.
rewrite dqEq; unfold FtoRradix; apply ClosestUlp; auto.
elim Roundq; auto.
rewrite <- Rmult_assoc.
apply Rle_trans with (3*(Rmin (Fulp bo radix precision p) (Fulp bo radix precision q)))%R;[idtac|apply Rmult_le_compat_r].
2: unfold Rmin; case (Rle_dec (Fulp bo radix precision p) (Fulp bo radix precision q)); intros H1; unfold Fulp; auto with real zarith.
2: apply powerRZ_le, IZR_lt; auto with zarith.
2: apply powerRZ_le, IZR_lt; auto with zarith.
2: right; simpl; unfold Rdiv; field; auto with real.
unfold Rmin; case (Rle_dec (Fulp bo radix precision p) (Fulp bo radix precision q)); intros H1.
apply Rle_trans with (Fulp bo radix precision p+2*Fulp bo radix precision p)%R;[apply Rplus_le_compat_l|right;ring].
apply Fulp_le_twice_r; auto with real; fold radix FtoRradix.
generalize Q_positive; auto with real.
apply Q_le_two_P.
apply Rle_trans with (2*Fulp bo radix precision q+Fulp bo radix precision q)%R;[apply Rplus_le_compat_r|right;ring].
apply Fulp_le_twice_r; auto with real; fold radix FtoRradix.
apply P_positive with bo precision b b'; auto with real.
apply P_le_two_Q.
Qed.

Theorem EvenClosestFabs :
 forall (f : float) (r : R), (Fcanonic radix bo f)
    -> EvenClosest bo radix precision r f ->
    EvenClosest bo radix precision (Rabs r) (Fabs f).
intros.
case (Rle_or_lt 0%R r); intros.
rewrite Rabs_right; auto with real.
unfold Fabs; rewrite Z.abs_eq; auto with zarith.
apply LeR0Fnum with (radix := radix); auto with zarith.
apply RleRoundedR0 with bo precision (EvenClosest bo radix precision) r; auto with zarith.
apply EvenClosestRoundedModeP; auto with zarith.
rewrite Rabs_left; auto with real.
replace (Fabs f) with (Fopp f).
generalize EvenClosestSymmetric; unfold SymmetricP; auto.
unfold Fabs, Fopp; rewrite <- Zabs_Zopp; rewrite Z.abs_eq; auto.
assert (Fnum f <= 0)%Z; auto with zarith.
apply R0LeFnum with (radix:=radix); auto with zarith.
apply RleRoundedLessR0 with bo precision (EvenClosest bo radix precision) r; auto with zarith real.
apply EvenClosestRoundedModeP; auto with zarith.
Qed.

Theorem discri2: (3*(Rmin (Fulp bo radix precision p) (Fulp bo radix precision q))
  <= (Rabs (p-q)))%R -> (delta <= 2*(Fulp bo radix precision d))%R.
intros H1; unfold delta.
apply Rle_trans with (1 * Fulp bo radix precision d)%R;[ring_simplify (1 * Fulp bo radix precision d)%R | unfold Fulp; auto with real zarith].
replace (d - (b * b' - a * c))%R with ((d-(t+s))+(t+s-b*b'+a*c))%R;[idtac|ring].
apply Rle_trans with (1:=Rabs_triang (d-(t+s))%R (t + s - b * b' + a * c)%R).
apply Rmult_le_reg_l with 2%R; auto with real.
rewrite Rmult_plus_distr_l.
apply Rle_trans with (Fulp bo radix precision d+Fulp bo radix precision d)%R;[idtac|right;ring].
apply Rplus_le_compat.
rewrite <- Rabs_Ropp; replace (- (d - (t + s)))%R with ((t+s)-d)%R;[idtac|ring].
replace 2%R with (INR 2); auto with real.
unfold FtoRradix; apply ClosestUlp; auto.
elim Roundd; auto.
rewrite t_exact.
replace (p - q + s - b * b' + a * c)%R with (-((dp-dq) - s))%R;[idtac|rewrite dpEq; rewrite dqEq; ring].
rewrite Rabs_Ropp; apply Rle_trans with (Fulp bo radix precision s).
unfold FtoRradix; apply ClosestUlp; auto.
elim Rounds; auto.
rewrite FulpFabs; auto; rewrite FulpFabs with (f:=d); auto.
apply LeFulpPos; auto with real zarith.
apply absFBounded; auto.
apply absFBounded; auto.
rewrite Fabs_correct; auto with real.
apply Rabs_pos.
apply EvenClosestMonotone2 with bo precision (Rabs (dp-dq)) (Rabs (t+s))%R; auto.
2: apply EvenClosestFabs; auto; left; auto.
2: apply EvenClosestFabs; auto; left; auto.
cut (Rabs (dp - dq) <= (Rabs (p-q))/2)%R.
intros H2; cut ((Rabs s) <= (Rabs t)/2)%R.
intros H3; apply Rle_trans with (1:=H2).
rewrite <- t_exact; apply Rle_trans with ((Rabs t)-(Rabs t)/2)%R.
right; unfold Rdiv; field; auto with real.
apply Rle_trans with ((Rabs t)-(Rabs s))%R; auto with real.
unfold Rminus; apply Rplus_le_compat_l; auto with real.
replace (t+s)%R with (t-(-s))%R; [idtac|ring].
apply Rle_trans with ((Rabs t)-(Rabs (-s)))%R;[idtac|apply Rabs_triang_inv].
rewrite Rabs_Ropp; auto with real.
assert (t/2=(Float (Fnum t) (Z.pred (Fexp t))))%R.
unfold FtoRradix, FtoR, Z.pred; simpl; rewrite powerRZ_add; auto with real zarith.
unfold powerRZ at 3; unfold radix; simpl; field.
unfold Rdiv; rewrite <- (Rabs_right (/2)%R); auto with real.
2: apply Rle_ge; apply Rlt_le; auto with real.
rewrite <- Rabs_mult; fold (Rdiv t 2%R).
rewrite H; unfold FtoRradix; rewrite <- Fabs_correct; auto.
rewrite <- Fabs_correct; auto.
apply EvenClosestMonotone2 with bo precision (Rabs (dp-dq))%R (Rabs (p-q)/2)%R; auto.
apply EvenClosestFabs; auto; left; auto.
replace (Rabs (p - q) / 2)%R with (FtoRradix (Fabs (Float (Fnum t) (Z.pred (Fexp t))))).
unfold FtoRradix; apply RoundedModeProjectorIdem with (b:=bo) (P:=(EvenClosest bo radix precision)); auto.
apply EvenClosestRoundedModeP; auto.
split; simpl; auto with zarith.
rewrite Z.abs_eq; auto with zarith.
apply Roundt.
unfold FtoRradix; rewrite Fabs_correct; auto; fold FtoRradix; rewrite <- H.
rewrite t_exact; unfold Rdiv; rewrite Rabs_mult; auto with real.
rewrite (Rabs_right (/2)%R); auto with real.
apply Rle_ge; apply Rlt_le; auto with real.
apply Rle_trans with (1:=dp_dq_le).
apply Rmult_le_reg_l with 2%R; auto with real; unfold Rdiv.
rewrite <- Rmult_assoc.
replace (2*(3*/2))%R with 3%R;[idtac|field; auto with real].
apply Rle_trans with (1:=H1).
right; field; auto with real.
apply Rmult_le_compat_r; auto with real.
apply powerRZ_le, Rlt_IZR; auto with zarith.
Qed.

Theorem discri3: (exists f:float, (Fbounded bo f) /\ (FtoRradix f)=(dp-dq)%R)
    -> (delta <= 2*(Fulp bo radix precision d))%R.
intros T; elim T; intros f T1; elim T1; intros H1 H2; clear T T1.
unfold delta.
replace (d - (b * b' - a * c))%R with (-((t+s)-d))%R.
apply Rmult_le_reg_l with (INR 2); auto with arith real.
apply Rle_trans with (Fulp bo radix precision d).
rewrite Rabs_Ropp; unfold FtoRradix; apply ClosestUlp; auto.
elim Roundd; auto.
simpl; apply Rle_trans with (1*(1*(Fulp bo radix precision d)))%R; unfold Fulp; auto with real zarith.
right; ring.
apply Rmult_le_compat; auto with real zarith.
ring_simplify (1 * powerRZ radix (Fexp (Fnormalize radix bo precision d)))%R; auto with real zarith.
apply powerRZ_le, Rlt_IZR; auto with zarith.
apply Rmult_le_compat_r; auto with real.
apply powerRZ_le, Rlt_IZR; auto with zarith.
replace (FtoRradix s) with (dp-dq)%R.
rewrite dpEq; rewrite dqEq; rewrite t_exact; ring.
rewrite <- H2.
unfold FtoRradix; apply RoundedModeProjectorIdemEq with (b:=bo) (P:=(EvenClosest bo radix precision)) (precision:=precision); auto with real.
apply EvenClosestRoundedModeP; auto.
fold FtoRradix; rewrite H2; auto.
Qed.

Theorem errorBoundedMultClosest_Can:
       forall f1 f2 g : float,
       Fbounded bo f1 ->
       Fbounded bo f2 ->
       Closest bo radix (f1* f2) g ->
       (powerRZ radix (-dExp bo+2*precision-1) <= Rabs (f1*f2))%R ->
       Fcanonic radix bo g ->
         (exists s : float,
            Fbounded bo s /\
           (FtoRradix s = f1*f2 - g)%R /\
            Fexp s = (Fexp g - precision)%Z /\
            (Rabs (Fnum s) <= powerRZ radix (Z.pred precision))%R).
intros.
generalize errorBoundedMultClosest; intros T.
elim T with (b:=bo) (radix:=radix) (precision:=precision) (p:=f1) (q:=f2) (pq:=g); auto with zarith real; clear T; fold FtoRradix.
intros g' T1; elim T1; intros dg T2; elim T2; intros H5 T3; elim T3; intros H6 T4; elim T4; intros H7 T5; elim T5; intros H8 T6; elim T6; intros H9 H10; clear T1 T2 T3 T4 T5 T6.
exists dg; split; auto; split.
rewrite <- H8; auto with real.
split; [replace g with g'; auto with zarith|idtac].
apply FcanonicUnique with radix bo precision; auto with zarith.
apply Rmult_le_reg_l with (powerRZ radix (Fexp dg)); auto with zarith real.
apply powerRZ_lt, Rlt_IZR; auto with zarith.
apply Rle_trans with (Rabs dg);[right; unfold FtoRradix, FtoR|idtac].
rewrite Rabs_mult;rewrite (Rabs_right (powerRZ radix (Fexp dg)));auto with real.
apply Rle_ge; auto with real zarith.
apply powerRZ_le, Rlt_IZR; auto with zarith.
rewrite H9; rewrite <- powerRZ_add; auto with real zarith.
apply Rmult_le_reg_l with (INR 2); auto with real zarith.
apply Rle_trans with (Fulp bo radix precision g').
unfold FtoRradix; apply ClosestUlp; auto.
replace g' with g; auto.
apply FcanonicUnique with radix bo precision; auto with zarith.
rewrite CanonicFulp; auto.
right; apply trans_eq with (powerRZ radix (Fexp g'));[unfold FtoR; simpl; ring|idtac].
apply trans_eq with ((powerRZ radix 1%Z)*(powerRZ radix (Fexp dg+Z.pred precision)))%R;[rewrite <- powerRZ_add; auto with zarith real|idtac].
2: rewrite powerRZ_1; unfold radix; simpl; ring.
rewrite H10; unfold Z.pred; auto with zarith real.
ring_simplify (1 + (Fexp g' - precision + (precision + -1)))%Z; auto with real.
assert (- dExp bo + 2 * precision - 1 < 2*precision+Fexp f1+Fexp f2)%Z; auto with zarith.
apply Zlt_powerRZ with radix; auto with real zarith.
apply Rle_lt_trans with (1:=H2).
rewrite Rabs_mult; unfold FtoRradix; repeat rewrite <- Fabs_correct; auto.
replace (2*precision)%Z with (precision+precision)%Z; auto with zarith.
unfold FtoR, Fabs; simpl.
repeat rewrite powerRZ_add; auto with real zarith.
repeat rewrite <- Rmult_assoc; apply Rmult_lt_compat_r; auto with real zarith.
apply powerRZ_lt, Rlt_IZR; auto with zarith.
apply Rle_lt_trans with (Z.abs (Fnum f1)*Z.abs (Fnum f2) * powerRZ 2 (Fexp f1))%R;
  [unfold radix; right; ring| apply Rmult_lt_compat_r; auto with real zarith].
apply powerRZ_lt, Rlt_IZR; auto with zarith.
apply Rle_lt_trans with (Z.abs (Fnum f1)* powerRZ 2 precision)%R.
apply Rmult_le_compat_l; auto with real zarith.
apply Rle_IZR; auto with zarith.
apply Rle_trans with (Zpos (vNum bo)).
apply Rle_IZR, Zlt_le_weak; apply H0.
rewrite pGivesBound; rewrite Zpower_nat_Z_powerRZ; auto with real.
apply Rmult_lt_compat_r; auto with real zarith.
apply powerRZ_lt, Rlt_IZR; auto with zarith.
apply Rlt_le_trans with (Zpos (vNum bo)).
apply Rlt_IZR; apply H.
rewrite pGivesBound; rewrite Zpower_nat_Z_powerRZ; auto with real.
rewrite FcanonicFnormalizeEq; auto with zarith.
assert (powerRZ radix (- dExp bo + 2 * precision - 1) <= Rabs g)%R.
cut (exists f:float, Fbounded bo f /\
  (FtoRradix f=(powerRZ radix (- dExp bo + 2 * precision - 1)))%R).
intros T; elim T; intros f T'; elim T'; intros; clear T T'.
rewrite <- H5; unfold FtoRradix.
apply RoundAbsMonotonel with bo precision
  (Closest bo radix) (f1*f2)%R; auto with zarith real.
apply ClosestRoundedModeP with precision; auto with zarith.
fold FtoRradix; rewrite H5; auto.
exists (Float 1 (-dExp bo+2*precision-1)).
split;[split|idtac].
simpl; apply vNumbMoreThanOne with radix precision; auto with zarith.
apply Z.le_trans with (- dExp bo + 2 * precision - 1)%Z; auto with zarith.
unfold FtoRradix, FtoR; simpl; ring.
assert (- dExp bo + 2 * precision - 1 < precision+Fexp g)%Z; auto with zarith.
apply Zlt_powerRZ with radix; auto with zarith real.
apply Rle_lt_trans with (1:=H4).
unfold FtoRradix; rewrite <- Fabs_correct; auto.
rewrite powerRZ_add; auto with real zarith.
unfold FtoR, Fabs; simpl.
apply Rmult_lt_compat_r; auto with real zarith.
apply powerRZ_lt, Rlt_IZR; auto with zarith.
apply Rlt_le_trans with (Zpos (vNum bo)).
apply Rlt_IZR, H1.
rewrite pGivesBound; rewrite Zpower_nat_Z_powerRZ; auto with real zarith.
Qed.

Theorem discri4: (Fexp p)=(Fexp q) -> (delta <= 2*(Fulp bo radix precision d))%R.
intros H1; apply discri3.
generalize errorBoundedMultClosest_Can; intros T.
elim T with (f1:=b) (f2:=b') (g:=p); auto with zarith real; clear T.
intros dp' T2; elim T2; intros H2 T3; elim T3; intros H3 T4; elim T4; intros H4 H5; clear T2 T3 T4.
2: elim Roundp; auto.
generalize errorBoundedMultClosest_Can; intros T.
elim T with (f1:=a) (f2:=c) (g:=q); auto with zarith real; clear T.
intros dq' T2; elim T2; intros H2' T3; elim T3; intros H3' T4; elim T4; intros H4' H5'; clear T2 T3 T4.
2: elim Roundq; auto.
2: left; auto.
2: left; auto.
assert ((Rabs (Fnum dp'-Fnum dq') < (powerRZ radix precision))%R \/
 (((Rabs dp')= (powerRZ radix (Z.pred (Fexp p))))%R /\ ((Rabs dq')= (powerRZ radix (Z.pred (Fexp p))))%R)).
case H5; intros.
left; unfold Rminus; apply Rle_lt_trans with (1:=Rabs_triang (Fnum dp') (-(Fnum dq'))%R).
rewrite Rabs_Ropp.
apply Rlt_le_trans with ((powerRZ radix (Z.pred precision)) +(Rabs (Fnum dq')))%R; auto with real zarith.
apply Rle_trans with ((powerRZ radix (Z.pred precision))+ (powerRZ radix (Z.pred precision)))%R; auto with real zarith.
right; unfold Z.pred; repeat rewrite powerRZ_add; auto with real zarith.
unfold powerRZ at 2 4; simpl; unfold radix; field.
case H5'; intros.
left; unfold Rminus; apply Rle_lt_trans with (1:=Rabs_triang (Fnum dp') (-(Fnum dq'))%R); rewrite Rabs_Ropp.
apply Rle_lt_trans with ((powerRZ radix (Z.pred precision)) +(Rabs (Fnum dq')))%R; auto with real zarith.
apply Rlt_le_trans with ((powerRZ radix (Z.pred precision))+ (powerRZ radix (Z.pred precision)))%R; auto with real zarith.
right; unfold Z.pred; repeat rewrite powerRZ_add; auto with real zarith.
unfold powerRZ at 2 4; simpl; unfold radix; field.
right; unfold FtoRradix, FtoR;repeat rewrite Rabs_mult.
rewrite (Rabs_right (powerRZ radix (Fexp dp'))); try apply Rle_ge; auto with real zarith.
rewrite (Rabs_right (powerRZ radix (Fexp dq'))); try apply Rle_ge; auto with real zarith.
rewrite H; rewrite H0.
repeat rewrite <- powerRZ_add; auto with real zarith.
rewrite H4'; rewrite H4; unfold Z.pred.
ring_simplify (precision + -1 + (Fexp p - precision))%Z; ring_simplify (precision + -1 + (Fexp q - precision))%Z;
  ring_simplify (Fexp p+-1)%Z; rewrite <- H1; auto with zarith real.
apply powerRZ_le, Rlt_IZR; auto with zarith.
apply powerRZ_le, Rlt_IZR; auto with zarith.
case H; clear H; intros H.
exists (Float ((Fnum dp')-(Fnum dq'))%Z (Fexp dq')).
split; [split; auto with zarith|idtac].
simpl; apply Zlt_Rlt.
rewrite pGivesBound;rewrite Zpower_nat_Z_powerRZ; auto.
rewrite <- Rabs_Zabs; unfold Zminus; rewrite plus_IZR; rewrite Ropp_Ropp_IZR; auto with real zarith.
simpl; auto with zarith.
apply H2'.
rewrite dpEq; rewrite dqEq; rewrite <- H3; rewrite <- H3'.
unfold FtoRradix, FtoR; simpl.
unfold Zminus; rewrite plus_IZR; rewrite Ropp_Ropp_IZR; replace (Fexp dp') with (Fexp dq');[ring|idtac].
rewrite H4'; rewrite <- H1; auto with zarith.
rewrite dpEq; rewrite dqEq; rewrite <- H3; rewrite <- H3'.
elim H; unfold Rabs; case (Rcase_abs dp'); case (Rcase_abs dq'); intros.
exists (Float 0%Z 0%Z); split;[split; auto with zarith|idtac].
simpl; case (dExp bo); auto with zarith.
apply trans_eq with (-(-dp')+-dq')%R;[rewrite H0; rewrite H6; unfold FtoRradix, FtoR;simpl|idtac];ring.
exists (Float (-2)%Z (Z.pred (Fexp p))); split;[split; simpl; auto with zarith|idtac].
rewrite pGivesBound; apply Z.le_lt_trans with (Zpower_nat radix 1); auto with zarith.
apply Zpower_nat_monotone_lt; auto with zarith.
apply Z.le_trans with (Fexp dp'); auto with zarith; apply H2.
apply trans_eq with (-(-dp')+-dq')%R;[rewrite H0; rewrite H6; unfold FtoRradix, FtoR; simpl|idtac];ring.
exists (Float 2%Z (Z.pred (Fexp p))); split;[split;simpl;auto with zarith|idtac].
rewrite pGivesBound; apply Z.le_lt_trans with (Zpower_nat radix 1); auto with zarith.
apply Zpower_nat_monotone_lt; auto with zarith.
apply Z.le_trans with (Fexp dp'); auto with zarith; apply H2.
unfold Rminus;rewrite H0; rewrite H6; unfold FtoRradix, FtoR;simpl; ring.
exists (Float 0%Z 0%Z); split;[split; auto with zarith|idtac].
simpl; case (dExp bo); auto with zarith.
rewrite H0; rewrite H6; unfold FtoRradix, FtoR; simpl;ring.
Qed.

End Discriminant2.

Section Discriminant3.
Variable bo : Fbound.
Variable precision : nat.

Let radix := 2%Z.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ TwoMoreThanOne).

Hypothesis precisionGreaterThanOne : 1 < precision.
Hypothesis pGivesBound : Zpos (vNum bo) = Zpower_nat radix precision.

Variables a b b' c p q t dp dq s d:float.

Let delta := (Rabs (d-(b*b'-a*c)))%R.

Hypothesis Fa : (Fbounded bo a).
Hypothesis Fb : (Fbounded bo b).
Hypothesis Fb': (Fbounded bo b').
Hypothesis Fc : (Fbounded bo c).
Hypothesis Fp : (Fbounded bo p).
Hypothesis Fq : (Fbounded bo q).
Hypothesis Fd : (Fbounded bo d).
Hypothesis Ft : (Fbounded bo t).
Hypothesis Fs : (Fbounded bo s).
Hypothesis Fdp: (Fbounded bo dp).
Hypothesis Fdq: (Fbounded bo dq).
Hypothesis Cs:(Fcanonic radix bo s).

There is no underflow
Hypothesis U1: (- dExp bo <= (Fexp d)-1)%Z.
Hypothesis U2: (powerRZ radix (-dExp bo+2*precision-1) <= Rabs (b*b'))%R.
Hypothesis U3: (powerRZ radix (-dExp bo+2*precision-1) <= Rabs (a*c))%R.

Hypothesis Np:(Fnormal radix bo p).
Hypothesis Nq:(Fnormal radix bo q).
Hypothesis Nd:(Fnormal radix bo d).

Hypothesis Square:(0 <=b*b')%R.

Hypothesis Roundp : (EvenClosest bo radix precision (b*b')%R p).
Hypothesis Roundq : (EvenClosest bo radix precision (a*c)%R q).

Hypothesis p_pos:(0 <= p)%R.
Hypothesis q_pos:(0 <= q)%R.

Hypothesis Secondcase : (3*(Rabs (p-q)) < p+q)%R.

Hypothesis Roundt : (EvenClosest bo radix precision (p-q)%R t).
Hypothesis dpEq : (FtoRradix dp=b*b'-p)%R.
Hypothesis dqEq : (FtoRradix dq=a*c-q)%R.
Hypothesis Rounds : (EvenClosest bo radix precision (dp-dq)%R s).
Hypothesis Roundd : (EvenClosest bo radix precision (t+s)%R d).

Hypothesis p_differ_q:~(p=q)%R.

Variable e:Z.
Hypothesis p_eqF : p=(Float (Zpower_nat radix (pred precision)) (Z.succ e)).
Hypothesis p_eqR : (FtoRradix p)=(powerRZ radix (precision+e)%Z).
Hypothesis q_eqExp : (Fexp q)=e.

Theorem discri5: (0 < dp*dq)%R -> (delta <= 2*(Fulp bo radix precision d))%R.
intros.
unfold FtoRradix, delta; apply discri3 with p q t dp dq s; auto.
assert (forall f1 f2 g : float,
      Fbounded bo f1 ->
      Fbounded bo f2 ->
      Closest bo 2 (FtoR 2 f1 * FtoR 2 f2) g ->
      (powerRZ (Zpos 2) (- dExp bo + 2 * precision - 1) <=
        Rabs (FtoR 2 f1 * FtoR 2 f2))%R ->
      Fcanonic 2 bo g ->
      exists s : float,
        Fbounded bo s /\
        FtoR 2 s = (FtoR 2 f1 * FtoR 2 f2 - FtoR 2 g)%R /\
        Fexp s = (Fexp g - precision)%Z /\
        (Rabs (Fnum s) <= powerRZ (Zpos 2) (Z.pred precision))%R).
apply errorBoundedMultClosest_Can; auto.
fold radix in H0; fold FtoRradix in H0.
elim H0 with (f1:=b) (f2:=b') (g:=p); auto with zarith real.
intros dp' T2; elim T2; intros H2 T3; elim T3; intros H3 T4; elim T4; intros H4 H5; clear T2 T3 T4.
2: elim Roundp; auto.
elim H0 with (f1:=a) (f2:=c) (g:=q); auto with zarith real; clear H0.
intros dq' T2; elim T2; intros H2' T3; elim T3; intros H3' T4; elim T4; intros H4' H5'; clear T2 T3 T4.
2: elim Roundq; auto.
2: left; auto.
2: left; auto.
fold radix; fold FtoRradix; rewrite dpEq; rewrite dqEq; rewrite <- H3; rewrite <- H3'.
exists (Fminus radix dp' dq'); split.
2: unfold FtoRradix; rewrite Fminus_correct; auto with real.
unfold Fminus, Fopp, Fplus; simpl.
repeat rewrite H4'; repeat rewrite q_eqExp; repeat rewrite H4.
replace (Fexp p) with (Z.succ e); [idtac|rewrite p_eqF; auto].
rewrite Zmin_le2; auto with zarith.
split; auto with zarith.
simpl; unfold Z.succ.
ring_simplify (e + 1 - precision - (e - precision))%Z; ring_simplify (e - precision - (e - precision))%Z.
simpl.
unfold nat_of_P, Zpower_nat; simpl.
replace ( - Fnum dq' * 1)%Z with (- Fnum dq')%Z; [idtac|ring].
apply Zlt_Rlt.
rewrite pGivesBound;rewrite Zpower_nat_Z_powerRZ; auto.
rewrite <- Rabs_Zabs; rewrite plus_IZR;rewrite mult_IZR;rewrite Ropp_Ropp_IZR.
assert (forall (x y z:R), (0 < x*y)%R -> (Rabs x <= z)%R ->
   (Rabs y <= z)%R -> (Rabs (2*x-y) < 2*z)%R).
intros.
unfold Rabs; case (Rcase_abs (2*x-y)%R); case (Rle_or_lt 0%R x); intros.
case H7; intros; replace (- (2 * x - y))%R with (-(2)*x+y)%R by ring.
apply Rlt_le_trans with (-(2)*0+y)%R; auto with real.
apply Rplus_lt_compat_r; repeat rewrite Ropp_mult_distr_l_reverse.
apply Ropp_lt_contravar; auto with real.
ring_simplify (-(2)*0+y)%R; apply Rle_trans with z; auto with real.
apply Rle_trans with (2:=H6); apply RRle_abs.
apply Rle_trans with (1*z)%R; auto with real.
apply Rmult_le_compat_r; auto with real.
apply Rle_trans with (2:=H1); auto with real.
apply Rabs_pos.
contradict H0; rewrite <- H8; auto with real.
ring_simplify (0*y)%R; auto with real.
ring_simplify (- (2 * x - y))%R.
apply Rlt_le_trans with (-2*x+0)%R;[apply Rplus_lt_compat_l|idtac].
apply Rmult_lt_reg_l with (-x)%R; auto with real.
apply Rle_lt_trans with (-(x*y))%R; auto with real.
apply Rlt_le_trans with (-0)%R; auto with real; right;ring.
apply Rle_trans with (2*(-x))%R;[right;ring|apply Rmult_le_compat_l; auto with real].
apply Rle_trans with (2:=H1); rewrite <- Rabs_Ropp; apply RRle_abs.
apply Rlt_le_trans with (2*x-0)%R;[unfold Rminus; apply Rplus_lt_compat_l|idtac].
apply Ropp_lt_contravar; apply Rmult_lt_reg_l with x; auto with real.
case H7; auto with real.
intros H8; contradict H0; rewrite <- H8; ring_simplify (0*y)%R; auto with real.
ring_simplify (x*0)%R; auto with real.
apply Rle_trans with (2*x)%R;[right;ring|apply Rmult_le_compat_l; auto with real].
apply Rle_trans with (2:=H1); apply RRle_abs.
apply Rlt_le_trans with (2*0-y)%R; [unfold Rminus; apply Rplus_lt_compat_r; apply Rmult_lt_compat_l; auto with real|idtac].
apply Rle_trans with (-y)%R;[right;ring|apply Rle_trans with z].
apply Rle_trans with (2:=H6); rewrite <- Rabs_Ropp; apply RRle_abs.
apply Rle_trans with (1*z)%R;[right;ring|apply Rmult_le_compat_r; auto with real].
apply Rle_trans with (2:=H1); auto with real.
apply Rabs_pos.
replace (Fnum dp' * Zpos 2+-Fnum dq')%R with (2*(Fnum dp')-Fnum dq')%R; auto with real zarith.
apply Rlt_le_trans with (2*powerRZ radix (Z.pred precision))%R.
apply H0; auto.
apply Rmult_lt_reg_l with (powerRZ radix (Fexp dq')); auto with real zarith.
apply powerRZ_lt, Rlt_IZR; auto with zarith.
apply Rmult_lt_reg_l with (powerRZ radix (Fexp dp')); auto with real zarith.
apply powerRZ_lt, Rlt_IZR; auto with zarith.
apply Rle_lt_trans with 0%R;[right;ring|apply Rlt_le_trans with (1:=H)].
rewrite dpEq; rewrite dqEq; rewrite <- H3; rewrite <- H3'.
unfold FtoRradix, FtoR; right; ring.
right; unfold Z.pred, Zminus; rewrite powerRZ_add; auto with real zarith.
unfold powerRZ at 2; simpl; unfold radix; field.
simpl; ring.
simpl;rewrite <-q_eqExp; rewrite <- H4'; auto with zarith.
apply H2'.
Qed.

Theorem discri6: (0< dp)%R -> (dq < 0)%R
    -> (delta <= 2*(Fulp bo radix precision d))%R.
intros;unfold delta.
replace (d - (b * b' - a * c))%R with (-((t+s)-d)+-((dp-dq)-s))%R.
2: rewrite dpEq; rewrite dqEq; unfold FtoRradix, radix; rewrite t_exact with bo precision b b' p q t; auto; ring.
apply Rle_trans with (1:=Rabs_triang (-(t+s-d))%R (-(dp-dq-s))%R).
apply Rmult_le_reg_l with (INR 2); auto with real zarith;rewrite Rmult_plus_distr_l.
apply Rle_trans with ((Fulp bo radix precision d)+(Fulp bo radix precision s))%R;[apply Rplus_le_compat|idtac].
rewrite Rabs_Ropp; unfold FtoRradix; apply ClosestUlp; auto.
elim Roundd; auto.
rewrite Rabs_Ropp; unfold FtoRradix; apply ClosestUlp; auto.
elim Rounds; auto.
apply Rle_trans with ((Fulp bo radix precision d+ 3* Fulp bo radix precision d))%R;[apply Rplus_le_compat_l|simpl;right;ring].
apply Rle_trans with (2*Fulp bo radix precision d)%R;[idtac|unfold Fulp; auto with real zarith].
rewrite FulpFabs; auto; rewrite FulpFabs with bo radix precision d; auto.
assert (2*(Fabs d)=(Float (Fnum (Fabs d)) (Z.succ (Fexp (Fabs d)))))%R.
unfold FtoRradix, FtoR, Z.succ; simpl; rewrite powerRZ_add; auto with real zarith.
rewrite powerRZ_1; unfold radix; ring.
apply Rle_trans with (Fulp bo radix precision (Float (Fnum (Fabs d)) (Z.succ (Fexp (Fabs d))))).
2:assert (Fnormal radix bo (Fabs d));[apply FnormalFabs; auto|idtac].
2:right; rewrite CanonicFulp; auto; [rewrite CanonicFulp|left]; auto.
2:unfold FtoR, Z.succ; simpl; rewrite powerRZ_add; auto with real zarith.
2:rewrite powerRZ_1; unfold radix; ring.
2:left; auto.
2:elim H2; intros H6 H5; elim H6; intros.
2:split; simpl; auto with zarith.
2:split; simpl; auto with zarith.
apply LeFulpPos; auto with real.
apply absFBounded; auto.
assert (Fnormal radix bo (Fabs d));[apply FnormalFabs; auto|idtac].
elim H2; intros H6 H5; elim H6; intros;split; simpl; auto with zarith.
rewrite Fabs_correct; auto with real zarith; apply Rabs_pos.
2: apply Rmult_le_compat_r; auto with real.
2: apply powerRZ_le, Rlt_IZR; auto with zarith.
apply EvenClosestMonotone2 with bo precision (Rabs (dp-dq))%R (2*Rabs (t+s))%R; auto.
2: apply EvenClosestFabs; auto; left; auto.
2: apply Twice_EvenClosest_Round; auto.
2: apply FnormalFabs; auto.
2: apply EvenClosestFabs; auto; left; auto.
unfold Rminus; apply Rle_trans with (1:=Rabs_triang dp (-dq)%R).
apply Rmult_le_reg_l with (INR 2); auto with real zarith; rewrite Rmult_plus_distr_l.
apply Rle_trans with (Fulp bo radix precision p+Fulp bo radix precision q)%R;[apply Rplus_le_compat|idtac].
rewrite dpEq; unfold FtoRradix; apply ClosestUlp; auto.
elim Roundp; auto.
rewrite Rabs_Ropp; rewrite dqEq; unfold FtoRradix; apply ClosestUlp; auto.
elim Roundq; auto.
rewrite CanonicFulp; auto;[idtac|left; auto].
rewrite CanonicFulp; auto;[idtac|left; auto].
apply Rle_trans with (3*(powerRZ radix e))%R;[right|idtac].
unfold FtoRradix, FtoR; simpl; rewrite q_eqExp; rewrite p_eqF; simpl.
unfold Z.succ; rewrite powerRZ_add; auto with real zarith.
rewrite powerRZ_1; unfold radix ;ring.
assert ((powerRZ radix e <= t))%R.
unfold FtoRradix, radix; rewrite t_exact with bo precision b b' p q t; auto.
fold radix; fold FtoRradix; rewrite p_eqR.
apply Rle_trans with (powerRZ radix (precision + e) - ((powerRZ radix precision - 1) * powerRZ radix e))%R; auto with real.
rewrite powerRZ_add; auto with real zarith; right;ring.
unfold Rminus; apply Rplus_le_compat_l; auto with real.
apply Ropp_le_contravar.
unfold FtoRradix, FtoR; rewrite q_eqExp; apply Rmult_le_compat_r; auto with real zarith.
apply powerRZ_le, Rlt_IZR; auto with zarith.
apply Rle_trans with (1:=(RRle_abs (Fnum q))).
assert (Z.abs (Fnum q) < Zpower_nat radix precision)%Z; auto with real zarith.
rewrite <- pGivesBound; auto with zarith; apply Nq.
rewrite Rabs_Zabs; apply Rle_trans with (Z.pred (Zpower_nat radix precision)); auto with real zarith.
apply Rle_IZR; auto with zarith.
unfold Z.pred; rewrite plus_IZR.
rewrite Zpower_nat_Z_powerRZ; right; simpl; ring.
assert (0<=s)%R.
unfold FtoRradix; apply RleRoundedR0 with bo precision (EvenClosest bo radix precision) (dp-dq)%R; auto with real.
apply EvenClosestRoundedModeP; auto.
apply Rle_trans with (0-0)%R; unfold Rminus; auto with real.
apply Rplus_le_compat; auto with real.
rewrite Rabs_right; auto with real.
2: apply Rle_ge; apply Rle_trans with (0+0)%R; auto with real.
2: apply Rplus_le_compat; auto with real zarith.
2: apply Rle_trans with (2:=H2); auto with real zarith.
2: apply powerRZ_le, Rlt_IZR; auto with zarith.
apply Rle_trans with (4*powerRZ radix e)%R;[apply Rmult_le_compat_r; auto with real zarith|idtac].
apply powerRZ_le, Rlt_IZR; auto with zarith.
apply Rle_trans with (4*(t+s))%R;[apply Rmult_le_compat_l; auto with real|simpl; right; ring].
apply Rle_trans with (powerRZ radix e+0)%R;[idtac|apply Rplus_le_compat];auto with real.
Qed.

Theorem discri7: (dp < 0)%R -> (0 < dq)%R
    -> (delta <= 2*(Fulp bo radix precision d))%R.
intros L1 L2.
unfold delta, FtoRradix; apply discri3 with p q t dp dq s; auto.
assert (H0:forall f1 f2 g : float,
      Fbounded bo f1 ->
      Fbounded bo f2 ->
      Closest bo 2 (FtoR 2 f1 * FtoR 2 f2) g ->
      (powerRZ (Zpos 2) (- dExp bo + 2 * precision - 1) <=
        Rabs (FtoR 2 f1 * FtoR 2 f2))%R ->
      Fcanonic 2 bo g ->
      exists s : float,
        Fbounded bo s /\
        FtoR 2 s = (FtoR 2 f1 * FtoR 2 f2 - FtoR 2 g)%R /\
        Fexp s = (Fexp g - precision)%Z /\
        (Rabs (Fnum s) <= powerRZ (Zpos 2) (Z.pred precision))%R).
apply errorBoundedMultClosest_Can; auto.
fold radix in H0; fold FtoRradix in H0.
elim H0 with (f1:=b) (f2:=b') (g:=p); auto with zarith real.
intros dp' T2; elim T2; intros H2 T3; elim T3; intros H3 T4; elim T4; intros H4
H5; clear T2 T3 T4.
2: elim Roundp; auto.
elim H0 with (f1:=a) (f2:=c) (g:=q); auto with zarith real; clear H0.
intros dq' T2; elim T2; intros H2' T3; elim T3; intros H3' T4; elim T4; intros H4' H5'; clear T2 T3 T4.
2: elim Roundq; auto.
2: left; auto.
2: left; auto.
cut (exists dp'':float, (Fexp dp''=Fexp dq' /\ (FtoRradix dp''=dp')%R /\
  (Rabs (Fnum dp'') <= powerRZ radix (Z.pred precision))%R)).
intros T; elim T; intros dp'' T1; elim T1; intros H4'' T2; elim T2;
 intros H5'' H6''; clear T T1 T2.
assert ((Rabs (Fnum dp''-Fnum dq') < (powerRZ radix precision))%R \/
 (((Rabs dp'')= (powerRZ radix (Z.pred (Fexp q))))%R /\ ((Rabs dq')=
(powerRZ radix (Z.pred (Fexp q))))%R)).
case H6''; intros.
left; unfold Rminus; apply Rle_lt_trans with (1:=Rabs_triang (Fnum dp'')
(-(Fnum dq'))%R).
rewrite Rabs_Ropp.
apply Rlt_le_trans with ((powerRZ radix (Z.pred precision)) +(Rabs (Fnum
dq')))%R; auto with real zarith.
apply Rle_trans with ((powerRZ radix (Z.pred precision))+ (powerRZ radix
(Z.pred precision)))%R; auto with real zarith.
right; unfold Z.pred; repeat rewrite powerRZ_add; auto with real zarith.
unfold powerRZ at 2 4; simpl; unfold radix; field.
case H5'; intros.
left; unfold Rminus; apply Rle_lt_trans with (1:=Rabs_triang (Fnum dp'') (-(Fnum
dq'))%R); rewrite Rabs_Ropp.
apply Rle_lt_trans with ((powerRZ radix (Z.pred precision)) +(Rabs (Fnum dq')))%R
; auto with real zarith.
apply Rlt_le_trans with ((powerRZ radix (Z.pred precision))+ (powerRZ radix (Z.pred precision)))%R; auto with real zarith.
right; unfold Z.pred; repeat rewrite powerRZ_add; auto with real zarith.
unfold powerRZ at 2 4; simpl; unfold radix; field.
right; unfold FtoRradix, FtoR;repeat rewrite Rabs_mult.
rewrite (Rabs_right (powerRZ radix (Fexp dp''))); try apply Rle_ge; auto with real zarith.
2: apply powerRZ_le, Rlt_IZR; auto with zarith.
rewrite (Rabs_right (powerRZ radix (Fexp dq'))); try apply Rle_ge; auto with real zarith.
2: apply powerRZ_le, Rlt_IZR; auto with zarith.
rewrite H; rewrite H0.
repeat rewrite <- powerRZ_add; auto with real zarith.
rewrite H4''; rewrite H4'; unfold Z.pred.
ring_simplify (precision + -1 + (Fexp q - precision))%Z; ring_simplify (precision + -1 + (Fexp q -
 precision))%Z; ring_simplify (Fexp q+-1)%Z; auto with zarith real.
case H; intros V; clear H.
exists (Float (Fnum dp''-Fnum dq') (Fexp dq')).
split;[split; auto with zarith|idtac].
simpl; apply Zlt_Rlt.
rewrite pGivesBound;rewrite Zpower_nat_Z_powerRZ; auto.
rewrite <- Rabs_Zabs; unfold Zminus; rewrite plus_IZR; rewrite Ropp_Ropp_IZR; auto with real zarith.
simpl; auto with zarith; apply H2'.
fold radix; fold FtoRradix; rewrite dpEq; rewrite dqEq.
rewrite <- H3'; rewrite <- H3;rewrite <- H5''.
unfold FtoRradix, FtoR; simpl.
unfold Zminus; rewrite plus_IZR; rewrite Ropp_Ropp_IZR.
rewrite H4''; ring.
exists (Float (-1)%Z (Fexp q)).
split;[split; simpl; auto with zarith|idtac].
rewrite pGivesBound; apply Z.le_lt_trans with (Zpower_nat radix 0); auto with zarith.
apply Zpower_nat_monotone_lt; auto with zarith.
apply Roundq.
fold radix; fold FtoRradix; elim V; intros.
replace (FtoRradix dp) with (-(-dp))%R;[idtac|ring].
rewrite <- (Rabs_left dp); auto with real.
rewrite <- (Rabs_right dq); auto with real.
2: apply Rle_ge; auto with real.
rewrite dpEq; rewrite <- H3; rewrite <- H5''; rewrite H.
rewrite dqEq; rewrite <- H3'; rewrite H0.
unfold FtoRradix, FtoR, Z.pred; simpl.
repeat rewrite powerRZ_add; auto with real zarith.
unfold powerRZ at 3 5; simpl; unfold radix; field.
assert (FtoRradix dp'=(Float (2*Fnum dp') (Z.pred (Fexp dp'))))%R.
unfold FtoRradix, FtoR, Z.pred.
apply trans_eq with ((2 * Fnum dp')%Z*(powerRZ radix (Fexp dp'+-1)))%R;[auto|idtac].
rewrite mult_IZR;rewrite powerRZ_add; auto with real zarith.
unfold powerRZ at 3; simpl; unfold radix; field.
simpl; auto with real.
exists (Float (2*Fnum dp') (Z.pred (Fexp dp'))); split.
simpl; rewrite H4'; rewrite H4.
rewrite q_eqExp; rewrite p_eqF; unfold Z.pred, Z.succ;simpl; auto with zarith.
split; auto with real.
apply Rmult_le_reg_l with (powerRZ radix (Z.pred (Fexp dp'))); auto with real zarith.
apply powerRZ_lt, Rlt_IZR; auto with zarith.
rewrite <- powerRZ_add; auto with real zarith.
rewrite <- (Rabs_right (powerRZ radix (Z.pred (Fexp dp'))));auto with real.
2: apply Rle_ge; auto with real zarith.
rewrite <- Rabs_mult.
replace (powerRZ radix (Z.pred (Fexp dp')) *
  Fnum (Float (2 * Fnum dp') (Z.pred (Fexp dp'))))%R with (FtoRradix dp'); auto with real.
2: rewrite H; unfold FtoRradix, FtoR; simpl; auto with real.
rewrite H3; rewrite <- dpEq.
rewrite H4; unfold Z.pred;ring_simplify (Fexp p - precision + -1 + (precision + -1))%Z.
rewrite Rabs_left; auto with real.
apply Rmult_le_reg_l with 2%R; auto with real.
apply Rplus_le_reg_l with (FtoRradix dp).
ring_simplify (dp+2*(-dp))%R.
rewrite <- Rabs_left; auto with real.
assert (Fbounded bo (Float (Z.pred (Zpower_nat radix precision)) e)).
split; auto with zarith.
simpl; rewrite pGivesBound; auto with zarith.
simpl; rewrite <- q_eqExp; apply Roundq; auto with zarith.
rewrite <- Rabs_Ropp.
replace (-dp)%R with (p-b*b')%R; [idtac|rewrite dpEq;ring].
elim Roundp; intros K1 K2; elim K1; intros K3 K4.
apply Rle_trans with (Rabs ((Float (Z.pred (Zpower_nat radix precision)) e)-b*b')).
unfold FtoRradix; apply K4; auto.
clear K1 K2 K3 K4; rewrite Rabs_left1.
rewrite dpEq; rewrite p_eqR.
apply Rle_trans with (b*b'-(powerRZ radix precision -1)*(powerRZ radix e))%R.
unfold FtoRradix, FtoR, Z.pred, radix; simpl.
rewrite plus_IZR; simpl; right; ring_simplify.
rewrite Zpower_nat_Z_powerRZ; auto with real zarith; simpl; ring.
unfold Rminus; rewrite Rplus_assoc; apply Rplus_le_compat_l.
replace (Fexp p) with (Z.succ e);[unfold Z.succ|rewrite p_eqF; simpl; auto with zarith].
ring_simplify (e+1-2)%Z; unfold Zminus.
repeat rewrite powerRZ_add; auto with real zarith.
unfold powerRZ at 6; simpl; unfold radix; right; field.
apply Rplus_le_reg_l with (p-(Float (Z.pred (Zpower_nat radix precision)) e))%R.
apply Rle_trans with (-(b*b'-p))%R;[right;ring|idtac].
rewrite <- dpEq; rewrite <- Rabs_left; auto with real.
rewrite dpEq; apply Rmult_le_reg_l with (INR 2); auto with real zarith.
apply Rle_trans with (Fulp bo radix precision p).
unfold FtoRradix; apply ClosestUlp; auto.
elim Roundp; auto.
rewrite CanonicFulp; auto;[idtac|left; auto].
replace (Fexp p) with (Z.succ e);[unfold Z.succ|rewrite p_eqF; simpl; auto with zarith].
rewrite p_eqR; unfold FtoRradix, FtoR, Z.pred; simpl.
rewrite plus_IZR; rewrite Zpower_nat_Z_powerRZ; auto with real zarith.
repeat rewrite powerRZ_add; auto with real zarith.
rewrite powerRZ_1; unfold radix;right; ring.
apply powerRZ_le, Rlt_IZR; auto with zarith.
Qed.

Theorem discri8: (delta <= 2*(Fulp bo radix precision d))%R.
case (Rle_or_lt 0%R dp); intros H;[case H; clear H; intros H|idtac].
case (Rle_or_lt 0%R dq); intros H';[case H'; clear H'; intros H'|idtac].
apply discri5; auto with real.
apply Rle_lt_trans with (dp*0)%R;[right;ring|auto with real].
unfold FtoRradix, delta; apply discri3 with p q t dp dq s; auto.
exists dp; split; auto.
fold radix; fold FtoRradix; rewrite <- H'; ring.
apply discri6; auto.
unfold FtoRradix, delta; apply discri3 with p q t dp dq s; auto.
exists (Fopp dq); split; auto with zarith.
apply oppBounded; easy.
rewrite Fopp_correct; fold radix; fold FtoRradix; rewrite <- H; ring.
case (Rle_or_lt 0%R dq); intros H';[case H'; clear H'; intros H'|idtac].
apply discri7; auto.
unfold FtoRradix, delta; apply discri3 with p q t dp dq s; auto.
exists dp; split; auto.
fold radix; fold FtoRradix; rewrite <- H'; ring.
apply discri5; auto.
apply Rle_lt_trans with (-dp*0)%R;[right;ring|idtac].
apply Rlt_le_trans with ((-dp)*(-dq))%R;[auto with real|right;ring].
Qed.

End Discriminant3.

Section Discriminant4.
Variable bo : Fbound.
Variable precision : nat.

Let radix := 2%Z.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ TwoMoreThanOne).

Hypothesis precisionGreaterThanOne : 1 < precision.
Hypothesis pGivesBound : Zpos (vNum bo) = Zpower_nat radix precision.

Variables a b c p q t dp dq s d:float.

Let delta := (Rabs (d-(b*b-a*c)))%R.

Hypothesis Fa : (Fbounded bo a).
Hypothesis Fb : (Fbounded bo b).
Hypothesis Fc : (Fbounded bo c).
Hypothesis Fp : (Fbounded bo p).
Hypothesis Fq : (Fbounded bo q).
Hypothesis Fd : (Fbounded bo d).
Hypothesis Ft : (Fbounded bo t).
Hypothesis Fs : (3*(Rabs (p-q)) < p+q)%R -> (Fbounded bo s).
Hypothesis Fdp: (3*(Rabs (p-q)) < p+q)%R -> (Fbounded bo dp).
Hypothesis Fdq: (3*(Rabs (p-q)) < p+q)%R -> (Fbounded bo dq).
Hypothesis Cs:(3*(Rabs (p-q)) < p+q)%R -> (Fcanonic radix bo s).

There is no underflow
Hypothesis U0: (- dExp bo <= (Fexp d)-1)%Z.
Hypothesis U1: (- dExp bo <= (Fexp t)-1)%Z.
Hypothesis U2: (powerRZ radix (-dExp bo+2*precision-1) <= Rabs (b*b))%R.
Hypothesis U3: (powerRZ radix (-dExp bo+2*precision-1) <= Rabs (a*c))%R.

Hypothesis Np:(Fnormal radix bo p).
Hypothesis Nq:(Fnormal radix bo q).
Hypothesis Nd:(Fnormal radix bo d).

Hypothesis Roundp : (EvenClosest bo radix precision (b*b)%R p).
Hypothesis Roundq : (EvenClosest bo radix precision (a*c)%R q).

Hypothesis Firstcase : (p+q <= 3*(Rabs (p-q)))%R ->
   (EvenClosest bo radix precision (p-q)%R d).

Hypothesis SRoundt : (3*(Rabs (p-q)) < p+q)%R -> (EvenClosest bo radix precision (p-q)%R t).
Hypothesis SdpEq : (3*(Rabs (p-q)) < p+q)%R -> (FtoRradix dp=b*b-p)%R.
Hypothesis SdqEq : (3*(Rabs (p-q)) < p+q)%R -> (FtoRradix dq=a*c-q)%R.
Hypothesis SRounds : (3*(Rabs (p-q)) < p+q)%R -> (EvenClosest bo radix precision (dp-dq)%R s).
Hypothesis SRoundd : (3*(Rabs (p-q)) < p+q)%R -> (EvenClosest bo radix precision (t+s)%R d).

Theorem discri9: (delta <= 2*(Fulp bo radix precision d))%R.
assert (Square:(0<=b*b)%R).
apply Rle_trans with (Rsqr b); auto with real.
case (Rle_or_lt (p + q)%R (3 * Rabs (p - q))%R); intros.
unfold delta;apply discri1 with p q; auto.
case (Rle_or_lt (3*(Rmin (Fulp bo radix precision p) (Fulp bo radix precision q)))%R (Rabs (p-q))%R); intros.
unfold delta; apply discri2 with p q t dp dq s; auto.
case (Zle_or_lt (Fexp q) (Fexp p)); intros.
case (Zle_lt_or_eq (Fexp q) (Fexp p)); auto;intros.
assert (Fexp q = Z.pred (Fexp p))%Z.
cut (Z.le (Fexp q) (Z.pred (Fexp p))); auto with zarith.
cut (Z.le (Z.pred (Fexp p)) (Fexp q)); auto with zarith.
clear H1;apply Zle_powerRZ with radix; auto with real zarith.
apply Rle_trans with (Fulp bo radix precision (Float (Fnum p) (Z.pred (Fexp p)))).
rewrite CanonicFulp; auto.
unfold FtoR; simpl; right; ring.
left; split; auto with zarith.
split; simpl; auto with zarith; try apply Np.
apply Z.le_trans with (Fexp q); auto with zarith; apply Nq.
simpl; elim Np; auto with zarith.
apply Rle_trans with (Fulp bo radix precision q).
apply LeFulpPos; auto with real zarith.
split; simpl; auto with zarith; try apply Np.
apply Z.le_trans with (Fexp q); auto with zarith; apply Nq.
apply LeFnumZERO; simpl; auto with real zarith.
apply LeR0Fnum with radix; auto with real zarith.
apply P_positive with bo precision b b; auto.
apply Rmult_le_reg_l with 2%R; auto with real.
apply Rle_trans with (FtoRradix p).
unfold FtoRradix, FtoR, Z.pred; simpl; rewrite powerRZ_add; auto with real zarith.
unfold powerRZ at 2; simpl; unfold radix; right; field.
apply P_le_two_Q with bo precision b b; auto.
rewrite CanonicFulp; auto with zarith.
unfold FtoR; simpl; right; ring.
left; auto.
clear H1 H2; assert (FtoRradix p=powerRZ radix (precision+Z.pred (Fexp p)))%R.
case (Zle_lt_or_eq (Zpower_nat radix (pred precision)) (Fnum p)).
elim Np; intros.
apply Zmult_le_reg_r with radix; auto with zarith; try (apply Z.lt_gt; auto with zarith).
apply Z.le_trans with (Zpower_nat radix precision).
pattern radix at 2 in |-*; replace radix with (Zpower_nat radix 1).
rewrite <- Zpower_nat_is_exp; auto with zarith.
apply Zpower_nat_monotone_le; auto with zarith.
simpl; auto with zarith.
rewrite <- pGivesBound; apply Z.le_trans with (1:=H2).
rewrite Z.abs_eq; auto with zarith.
apply Z.mul_nonneg_nonneg; try unfold radix; auto with zarith.
apply LeR0Fnum with radix; auto.
apply P_positive with bo precision b b; auto.
intros H1; contradict H0.
apply Rle_not_lt.
rewrite CanonicFulp; auto; [idtac|left; auto].
rewrite CanonicFulp; auto; [idtac|left; auto].
replace (FtoR radix (Float (S 0) (Fexp p))) with (powerRZ radix (Fexp p));[idtac|unfold FtoR; simpl; ring].
replace (FtoR radix (Float (S 0) (Fexp q))) with (powerRZ radix (Fexp q));[idtac|unfold FtoR; simpl; ring].
rewrite H3; unfold Rmin.
case (Rle_dec (powerRZ radix (Fexp p)) (powerRZ radix (Z.pred (Fexp p)))); auto with real zarith; intros J.
contradict J; apply Rlt_not_le; auto with real zarith.
apply Rlt_powerRZ; try apply Rlt_IZR; auto with zarith.
clear J; rewrite Rabs_right.
unfold FtoRradix, FtoR, Rminus.
apply Rle_trans with ((Z.succ (Zpower_nat radix (pred precision))*(powerRZ radix (Fexp p))+-((Z.pred (Zpower_nat radix precision))*(powerRZ radix (Fexp q)))))%R.
unfold Z.pred, Z.succ; rewrite plus_IZR; rewrite plus_IZR; repeat rewrite Zpower_nat_Z_powerRZ.
rewrite inj_pred; auto with zarith.
rewrite H3; unfold Z.pred; simpl; right;ring_simplify.
repeat rewrite powerRZ_add; auto with real zarith.
unfold powerRZ at 2 4 7 9; simpl; unfold radix; field.
apply Rplus_le_compat;[apply Rmult_le_compat_r; auto with real zarith|idtac].
apply powerRZ_le, Rlt_IZR; auto with zarith.
apply Rle_IZR; auto with zarith.
apply Ropp_le_contravar; apply Rmult_le_compat_r; auto with real zarith.
apply powerRZ_le, Rlt_IZR; auto with zarith.
apply Rle_trans with (1:=RRle_abs (Fnum q)).
rewrite Rabs_Zabs; rewrite <- pGivesBound;auto with zarith.
destruct Nq as ((T1,T2),T3); apply Rle_IZR; auto with zarith.
apply Rle_ge; apply Rlt_le; apply Rplus_lt_reg_r with q.
ring_simplify.
unfold FtoRradix; apply FcanonicPosFexpRlt with bo precision; auto with zarith.
apply Rlt_le; apply Q_positive with bo precision b b p; auto.
apply P_positive with bo precision b b; auto.
left; auto.
left; auto.
intros H1; unfold FtoRradix, FtoR; rewrite <- H1.
rewrite Zpower_nat_Z_powerRZ; rewrite <- powerRZ_add; auto with real zarith.
rewrite inj_pred; auto with zarith.
replace (Z.pred precision+Fexp p)%Z with (precision + Z.pred (Fexp p))%Z;[auto with real|unfold Z.pred; ring].
unfold delta; apply discri8 with p q t dp dq s (Z.pred (Fexp p)); auto.
apply FcanonicUnique with radix bo precision; auto with zarith real.
left; auto.
left; split; [split;simpl|idtac]; auto with zarith.
rewrite Z.abs_eq; auto with zarith.
rewrite pGivesBound; auto with zarith.
apply Zpower_nat_monotone_lt; auto with zarith.
apply Zpower_NR0; auto with zarith.
apply Z.le_trans with (Fexp p); auto with zarith; apply Np.
simpl (Fnum (Float (Zpower_nat 2 (pred precision)) (Z.succ (Z.pred (Fexp p))))).
replace radix with (Zpower_nat radix 1);[idtac|simpl; auto with zarith].
rewrite <-Zpower_nat_is_exp; rewrite pGivesBound; auto with zarith.
rewrite Z.abs_eq; auto with zarith.
apply Zpower_nat_monotone_le; auto with zarith.
apply Zpower_NR0; auto with zarith.
fold FtoRradix; rewrite H1; unfold FtoRradix, FtoR, Z.pred, Z.succ; simpl.
rewrite Zpower_nat_Z_powerRZ; rewrite <- powerRZ_add; auto with real zarith.
rewrite inj_pred; auto with zarith; unfold Z.pred.
replace (precision + -1 + (Fexp p + -1 + 1))%Z with (precision+(Fexp p+-1))%Z; auto with real; ring.
unfold delta;apply discri4 with p q t dp dq s; auto.
assert (Fexp p = Z.pred (Fexp q))%Z.
cut (Z.le (Fexp p) (Z.pred (Fexp q))); auto with zarith.
cut (Z.le (Z.pred (Fexp q)) (Fexp p)); auto with zarith.
clear H1;apply Zle_powerRZ with radix; auto with real zarith.
apply Rle_trans with (Fulp bo radix precision (Float (Fnum q) (Z.pred (Fexp q)))).
rewrite CanonicFulp; auto.
unfold FtoR; simpl; right; ring.
left; split; auto with zarith.
split; simpl; auto with zarith.
apply Nq.
assert (exists s : float,
         Fbounded bo s /\
         FtoR 2 s = (FtoR 2 a * FtoR 2 c - FtoR 2 q)%R /\
         Fexp s = (Fexp q - precision)%Z /\
         (Rabs (Fnum s) <= powerRZ (Zpos 2) (Z.pred precision))%R).
apply errorBoundedMultClosest_Can; auto.
elim Roundq; auto.
left; auto.
elim H1; intros s' T; elim T; intros T1 T2; elim T2; intros T3 T4; elim T4; intros.
apply Z.le_trans with (Fexp q - precision)%Z; auto with zarith.
rewrite <- H2; auto with zarith.
apply T.
simpl; elim Nq; auto with zarith.
apply Rle_trans with (Fulp bo radix precision p).
apply LeFulpPos; auto with real zarith.
split; simpl; auto with zarith; try apply Nq.
assert (exists s : float,
         Fbounded bo s /\
         FtoR 2 s = (FtoR 2 a * FtoR 2 c - FtoR 2 q)%R /\
         Fexp s = (Fexp q - precision)%Z /\
         (Rabs (Fnum s) <= powerRZ (Zpos 2) (Z.pred precision))%R).
apply errorBoundedMultClosest_Can; auto.
elim Roundq; auto.
left; auto.
elim H1; intros s' T; elim T; intros T1 T2; elim T2; intros T3 T4; elim T4; intros.
apply Z.le_trans with (Fexp q - precision)%Z; auto with zarith.
rewrite <- H2; auto with zarith; apply T.
apply LeFnumZERO; simpl; auto with real zarith.
apply LeR0Fnum with radix; auto with real zarith.
apply Rlt_le; apply Q_positive with bo precision b b p; auto.
apply Rmult_le_reg_l with 2%R; auto with real.
apply Rle_trans with (FtoRradix q).
unfold FtoRradix, FtoR, Z.pred; simpl; rewrite powerRZ_add; auto with real zarith.
unfold powerRZ at 2; simpl; unfold radix; right; field.
apply Q_le_two_P with bo precision b b; auto.
rewrite CanonicFulp; auto with zarith.
unfold FtoR; simpl; right; ring.
left; auto.
clear H1; assert (FtoRradix q=powerRZ radix (precision+Z.pred (Fexp q)))%R.
case (Zle_lt_or_eq (Zpower_nat radix (pred precision)) (Fnum q)).
elim Nq; intros.
apply Zmult_le_reg_r with radix; auto with zarith; try (apply Z.lt_gt; auto with zarith).
apply Z.le_trans with (Zpower_nat radix precision).
pattern radix at 2 in |-*; replace radix with (Zpower_nat radix 1).
rewrite <- Zpower_nat_is_exp; auto with zarith.
apply Zpower_nat_monotone_le; auto with zarith.
simpl; auto with zarith.
rewrite <- pGivesBound; apply Z.le_trans with (1:=H3).
rewrite Z.abs_eq; auto with zarith.
apply Z.mul_nonneg_nonneg; try unfold radix; auto with zarith.
apply LeR0Fnum with radix; auto.
apply Rlt_le; apply Q_positive with bo precision b b p; auto.
intros H1; contradict H0.
apply Rle_not_lt.
rewrite CanonicFulp; auto; [idtac|left; auto].
rewrite CanonicFulp; auto; [idtac|left; auto].
replace (FtoR radix (Float (S 0) (Fexp p))) with (powerRZ radix (Fexp p));[idtac|unfold FtoR; simpl; ring].
replace (FtoR radix (Float (S 0) (Fexp q))) with (powerRZ radix (Fexp q));[idtac|unfold FtoR; simpl; ring].
rewrite H2; unfold Rmin.
case (Rle_dec (powerRZ radix (Z.pred (Fexp q))) (powerRZ radix (Fexp q))); auto with real zarith; intros J.
clear J; rewrite Rabs_left1.
unfold FtoRradix, FtoR, Rminus.
apply Rle_trans with (- ((Z.pred (Zpower_nat radix precision)) * powerRZ radix (Fexp p) + - ((Z.succ (Zpower_nat radix (pred precision))) * powerRZ radix (Fexp q))))%R.
unfold Z.pred, Z.succ; rewrite plus_IZR; rewrite plus_IZR; repeat rewrite Zpower_nat_Z_powerRZ.
rewrite inj_pred; auto with zarith.
rewrite H2; unfold Z.pred; simpl; right.
repeat rewrite powerRZ_add; auto with real zarith.
unfold powerRZ at 2 5 7; simpl; unfold radix; field.
apply Ropp_le_contravar.
apply Rplus_le_compat;[apply Rmult_le_compat_r; auto with real zarith|idtac].
apply powerRZ_le, Rlt_IZR; auto with zarith.
2:apply Ropp_le_contravar; apply Rmult_le_compat_r; auto with real zarith.
2: apply powerRZ_le, Rlt_IZR; auto with zarith.
apply Rle_trans with (1:=RRle_abs (Fnum p)).
rewrite Rabs_Zabs; rewrite <- pGivesBound;auto with zarith.
apply Rle_IZR, Zle_Zpred, Np.
apply Rle_IZR; auto with zarith.
apply Rlt_le; apply Rplus_lt_reg_r with q.
ring_simplify.
unfold FtoRradix; apply FcanonicPosFexpRlt with bo precision; auto with zarith.
2:apply Rlt_le; apply Q_positive with bo precision b b p; auto.
apply P_positive with bo precision b b; auto.
left; auto.
left; auto.
contradict J; auto with real zarith.
apply Rle_powerRZ; try apply Rle_IZR; try unfold radix; auto with zarith.
intros H1; unfold FtoRradix, FtoR; rewrite <- H1.
rewrite Zpower_nat_Z_powerRZ; rewrite <- powerRZ_add; auto with real zarith.
rewrite inj_pred; auto with zarith.
replace (Z.pred precision+Fexp q)%Z with (precision + Z.pred (Fexp q))%Z;[auto with real|unfold Z.pred; ring].
unfold delta; rewrite <-Rabs_Ropp.
replace (- (d - (b * b - a * c)))%R with (Fopp d-(a*c-b*b))%R;[idtac|unfold FtoRradix; rewrite Fopp_correct; ring].
replace (Fulp bo radix precision d) with (Fulp bo radix precision (Fopp d)); auto with zarith.
2: unfold Fulp; rewrite Fnormalize_Fopp; unfold Fopp; simpl; auto with real zarith.
apply discri8 with q p (Fopp t) dq dp (Fopp s) (Z.pred (Fexp q)); auto with zarith.
apply oppBounded; auto.
apply FcanonicFopp; auto.
apply FnormalFop; auto.
fold radix; fold FtoRradix; case (Rle_or_lt 0%R (a*c)%R); auto.
intros H3; absurd (0 < q)%R.
apply Rle_not_lt; unfold FtoRradix; apply RleRoundedLessR0 with bo precision (EvenClosest bo radix precision) (a*c)%R; auto with real.
apply EvenClosestRoundedModeP; auto.
apply Q_positive with bo precision b b p; auto.
fold radix; fold FtoRradix; rewrite (Rplus_comm q p).
replace (q-p)%R with (-(p-q))%R; auto with real; rewrite Rabs_Ropp; auto.
fold radix; fold FtoRradix; replace (q-p)%R with (-(p-q))%R; auto with real.
generalize EvenClosestSymmetric; unfold SymmetricP;intros T; apply T; auto.
fold radix; fold FtoRradix; replace (dq-dp)%R with (-(dp-dq))%R;auto with real.
generalize EvenClosestSymmetric; unfold SymmetricP;intros T; apply T; auto.
fold radix; fold FtoRradix; replace (Fopp t+Fopp s)%R with (-(t+s))%R;[idtac|unfold FtoRradix; repeat rewrite Fopp_correct; ring].
generalize EvenClosestSymmetric; unfold SymmetricP;intros T; apply T; auto.
apply FcanonicUnique with radix bo precision; auto with zarith real.
left; auto.
left; split; [split;simpl|idtac]; auto with zarith.
rewrite Z.abs_eq; auto with zarith.
rewrite pGivesBound; auto with zarith.
apply Zpower_nat_monotone_lt; auto with zarith.
apply Zpower_NR0; auto with zarith.
apply Z.le_trans with (Fexp q); auto with zarith; apply Fq.
simpl (Fnum (Float (Zpower_nat 2 (pred precision)) (Z.succ (Z.pred (Fexp q))))).
replace radix with (Zpower_nat radix 1);[idtac|simpl; auto with zarith].
rewrite <-Zpower_nat_is_exp; rewrite pGivesBound; auto with zarith.
rewrite Z.abs_eq; auto with zarith.
apply Zpower_nat_monotone_le; auto with zarith.
apply Zpower_NR0; auto with zarith.
fold FtoRradix; rewrite H1; unfold FtoRradix, FtoR, Z.pred, Z.succ; simpl.
rewrite Zpower_nat_Z_powerRZ; rewrite <- powerRZ_add; auto with real zarith.
rewrite inj_pred; auto with zarith; unfold Z.pred.
replace (precision + -1 + (Fexp q + -1 + 1))%Z with (precision+(Fexp q+-1))%Z; auto with real; ring.
Qed.

End Discriminant4.


Section Discriminant1A.
Variable bo : Fbound.
Variable precision : nat.

Let radix := 2%Z.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ TwoMoreThanOne).

Hypothesis precisionGreaterThanOne : 1 < precision.
Hypothesis pGivesBound : Zpos (vNum bo) = Zpower_nat radix precision.
Hypothesis precisionGreaterThanThree : 3 <= precision.

Theorem RoundLeNormal: forall f:float, forall r:R,
  Closest bo radix r f -> (Fnormal radix bo f) ->
  (Rabs f <= Rabs r / (1 - powerRZ radix (- precision)))%R.
intros.
assert (0 < (1 - powerRZ radix (- precision)))%R.
apply Rplus_lt_reg_r with (powerRZ radix (- precision)).
ring_simplify.
replace 1%R with (powerRZ radix 0);
    [ auto with real zarith | simpl in |- *; auto ].
apply Rlt_powerRZ; try apply Rlt_IZR; auto with zarith.
apply Rmult_le_reg_l with ((1 - powerRZ radix (- precision)))%R; auto with real.
apply Rle_trans with (Rabs r);[idtac|right; field; auto with real].
apply Rplus_le_reg_l with (-Rabs r+Rabs f*powerRZ radix (-precision))%R.
apply Rle_trans with (Rabs f-Rabs r)%R;[right; ring|idtac].
apply Rle_trans with (Rabs f * powerRZ radix (- precision) )%R;[idtac|right; ring].
apply Rle_trans with (Rabs (f-r));[apply Rabs_triang_inv|idtac].
replace (f-r)%R with (-(r-f))%R;[rewrite Rabs_Ropp|ring].
apply Rmult_le_reg_l with (INR 2); auto with real.
apply Rle_trans with (Fulp bo radix precision f).
unfold FtoRradix; apply ClosestUlp; auto.
apply Rle_trans with (Rabs f*powerRZ radix (Z.succ (-precision)))%R.
unfold FtoRradix; apply FulpLe2; auto.
elim H0; auto.
rewrite FcanonicFnormalizeEq; auto with zarith.
left; auto.
right; unfold Z.succ; rewrite powerRZ_add; auto with real zarith.
rewrite powerRZ_1; unfold radix; simpl; ring.
Qed.

Variables a b b' c p q t d u v dp dq:float.

Let delta := (Rabs (d-(b*b'-a*c)))%R.

Hypothesis Fa : (Fbounded bo a).
Hypothesis Fb : (Fbounded bo b).
Hypothesis Fb': (Fbounded bo b').
Hypothesis Fc : (Fbounded bo c).
Hypothesis Fp : (Fbounded bo p).
Hypothesis Fq : (Fbounded bo q).
Hypothesis Fd : (Fbounded bo d).
Hypothesis Fu : (Fbounded bo u).
Hypothesis Fv : (Fbounded bo v).
Hypothesis Cand : (Fcanonic radix bo d).

There is no underflow
Hypothesis Nq:(Fnormal radix bo q).
Hypothesis Np:(Fnormal radix bo p).
Hypothesis Nv:(Fnormal radix bo v).
Hypothesis Nu:(Fnormal radix bo u).
Hypothesis U0: (- dExp bo <= Fexp p - 2)%Z.

Hypothesis Square:(0 <=b*b')%R.

Hypothesis Roundp : (EvenClosest bo radix precision (b*b')%R p).
Hypothesis Roundq : (EvenClosest bo radix precision (a*c)%R q).
Hypothesis Roundt : (EvenClosest bo radix precision (p-q)%R t).
Hypothesis dDef : d=t.
Hypothesis Roundu : (EvenClosest bo radix precision (3*Rabs t)%R u).
Hypothesis Roundv : (EvenClosest bo radix precision (p+q)%R v).
Hypothesis dpEq : (FtoRradix dp=b*b'-p)%R.
Hypothesis dqEq : (FtoRradix dq=a*c-q)%R.

Hypothesis Case1 : (3*(Rabs (p-q)) < p+q )%R.
Hypothesis Case2 : (v <= u )%R.

Theorem IneqEq: (FtoRradix v=u)%R.
assert (u <= v)%R; auto with real.
generalize (EvenClosestMonotone bo radix precision); unfold MonotoneP.
intros L; apply L with (3 * Rabs (p - q))%R (p+q)%R; auto; clear L.
generalize (EvenClosestCompatible bo radix precision); unfold CompatibleP.
intros L.
apply L with (3*Rabs t)%R u; auto.
replace (FtoRradix t) with (p-q)%R; auto with real.
apply sym_eq; unfold FtoRradix, radix; apply t_exact with bo precision b b'; auto.
Qed.

Theorem dexact: (FtoRradix d=p-q)%R.
unfold FtoRradix, radix; apply t_exact with bo precision b b'; auto.
rewrite dDef; auto with real zarith.
Qed.

Theorem discri10: (q <= p)%R -> (delta <= 2*(Fulp bo radix precision d))%R.
intros G.
case (Zle_lt_or_eq (Fexp d) (Fexp q)).
apply Fcanonic_Rle_Zle with radix bo precision; auto with real zarith.
left; auto.
fold FtoRradix; rewrite Rabs_right.
rewrite Rabs_right.
rewrite dexact; apply Rplus_le_reg_l with q.
ring_simplify; apply P_le_two_Q with bo precision b b'; auto.
apply Rle_ge; left; apply Q_positive with bo precision b b' p; auto.
apply Rle_ge; rewrite dexact; apply Rplus_le_reg_l with q.
ring_simplify; auto with real.
assert ( 0 < 2+(powerRZ radix (-precision)))%R.
apply Rle_lt_trans with (0+0)%R; auto with real.
apply Rplus_lt_compat; auto with real zarith.
apply powerRZ_lt, Rlt_IZR; auto with zarith.
assert ((q <= p * (1 + powerRZ radix (1 - precision)) /
    (2 + powerRZ radix (- precision)))%R).
apply Rmult_le_reg_l with (2 + powerRZ radix (- precision))%R; auto with real.
apply Rle_trans with (p * (1 + powerRZ radix (1 - precision)))%R;[idtac|right; field; auto with real].
apply Rmult_le_reg_l with 2%R; auto with real.
apply Rplus_le_reg_l with (-2*p+p*(powerRZ radix (1-precision))-3*q*(powerRZ radix (1-precision)))%R.
apply Rle_trans with ((powerRZ radix (1-precision))*(3*(p-q)))%R;[idtac|right; ring].
apply Rle_trans with ((4*q-2*p)*(1-powerRZ radix (-precision)))%R.
right; unfold Zminus; rewrite powerRZ_add; auto with real zarith.
rewrite powerRZ_1; unfold radix ; ring.
assert (0 < (1 - powerRZ radix (- precision)))%R.
apply Rplus_lt_reg_r with (powerRZ radix (- precision)).
ring_simplify.
replace 1%R with (powerRZ radix 0);
    [ auto with real zarith | simpl in |- *; auto ].
apply Rlt_powerRZ; try apply Rlt_IZR; auto with zarith.
apply Rmult_le_reg_l with (/ (1 - powerRZ radix (- precision)))%R; auto with real.
apply Rle_trans with (4 * q - 2 * p)%R;[right; field; auto with real|idtac].
apply Rle_trans with (Fulp bo radix precision v).
apply Rle_trans with (((p+q)-v) +- (3*(p-q)-v))%R;[right;ring|idtac].
apply Rle_trans with (Rabs (((p+q)-v) +- (3*(p-q)-v)));[apply RRle_abs|idtac].
apply Rle_trans with (Rabs (((p+q)-v)) +Rabs (- (3*(p-q)-v)))%R;[apply Rabs_triang|idtac].
apply Rmult_le_reg_l with 2%R; auto with real.
apply Rle_trans with (2*(Rabs (((p+q)-v))) +2*Rabs (- (3*(p-q)-v)))%R;[right;ring|idtac].
apply Rle_trans with (Fulp bo radix precision v+Fulp bo radix precision v)%R;[idtac|right;ring].
rewrite Rabs_Ropp; replace 2%R with (INR 2); auto with real zarith.
apply Rplus_le_compat; unfold FtoRradix; apply ClosestUlp; auto with zarith.
elim Roundv; auto.
elim Roundu; intros Y1 Y2.
apply ClosestCompatible with (1:=Y1); fold FtoRradix; auto.
rewrite <- dDef; rewrite dexact.
rewrite Rabs_right; auto with real.
apply Rle_ge; apply Rplus_le_reg_l with q; auto with real.
ring_simplify (q+0)%R; apply Rle_trans with (1:=G); right; ring.
apply sym_eq; apply IneqEq.
apply Rle_trans with (Rabs v * powerRZ radix (Z.succ (- precision)))%R.
unfold FtoRradix; apply FulpLe2; auto.
rewrite FcanonicFnormalizeEq; auto with zarith.
left; auto.
unfold Z.succ; replace (-precision+1)%Z with (1-precision)%Z; auto with zarith.
apply Rle_trans with (((3 * (p - q))*/(1 - powerRZ radix (- precision)))* powerRZ radix (1 - precision))%R;
   [idtac|right; ring].
apply Rmult_le_compat_r; auto with real zarith.
apply powerRZ_le, Rlt_IZR; auto with zarith.
assert (0 <= p-q)%R.
apply Rplus_le_reg_l with q; auto with real.
ring_simplify (q+0)%R; apply Rle_trans with (1:=G); right; ring.
apply Rle_trans with (Rabs (3 * Rabs t) / (1 - powerRZ radix (- precision)))%R.
rewrite IneqEq.
apply RoundLeNormal; auto.
elim Roundu; auto.
rewrite <- dDef; rewrite dexact.
rewrite Rabs_right;[rewrite Rabs_right; auto with real; right; ring|idtac].
rewrite Rabs_right; auto with real.
apply Rle_ge; apply Rle_trans with (3*0)%R; auto with real.
intros W.
case (Zle_lt_or_eq (Fexp q) (Fexp p)).
apply Fcanonic_Rle_Zle with radix bo precision; auto with zarith.
left; auto.
left; auto.
assert (0 <= q)%R.
left; apply Q_positive with bo precision b b' p; auto.
fold FtoRradix; repeat rewrite Rabs_right; auto with real.
apply Rle_ge; apply Rle_trans with (1:=H1); auto with real.
intros M.
assert (exists qq:float, (FtoRradix qq=2*q)%R /\ Fbounded bo qq /\ Fnormal radix bo qq
  /\ Fnum qq=Fnum q /\ (Fexp qq=Fexp q+1)%Z).
exists (Float (Fnum q) ((Fexp q)+1)).
elim Nq; intros F1 F2; elim F1; intros.
split;[unfold FtoRradix, FtoR; simpl|idtac].
rewrite powerRZ_add; auto with real zarith.
rewrite powerRZ_1; unfold radix; ring.
repeat split; simpl; auto with zarith.
elim H1; intros qq H1'; elim H1'; intros J1 H2; elim H2; intros J2 H3; elim H3;
  intros J3 H4; elim H4; intros J4 J5;clear H1' H1 H2 H3 H4.
assert (2*q=FSucc bo radix precision p)%R.
rewrite <- J1.
assert ((FSucc bo radix precision p <= qq)%R /\ (qq <= FSucc bo radix precision p)%R).
2: elim H1; auto with real.
split.
unfold FtoRradix; apply FSuccProp; auto with real zarith.
left; auto.
left; auto.
fold FtoRradix; rewrite J1.
assert (p <= 2*q)%R.
apply P_le_two_Q with bo precision b b'; auto.
case H1; auto.
intros A.
absurd (3 * Rabs (p - q) < p + q)%R; auto.
apply Rle_not_lt; rewrite A.
ring_simplify (2*q-q)%R; rewrite Rabs_right.
right; ring.
apply Rle_ge; left; apply Q_positive with bo precision b b' p; auto.
case (Rle_or_lt qq (FSucc bo radix precision p)); auto; intros N.
absurd (FSucc bo radix precision (FSucc bo radix precision p) <= qq)%R.
2: unfold FtoRradix; apply FSuccProp; auto with zarith.
2: apply FSuccCanonic; auto with zarith.
2: left; auto.
2: left; auto.
apply Rlt_not_le.
apply Rle_lt_trans with (2*(p*(1+powerRZ radix (1-precision))/(2+(powerRZ radix (-precision)))))%R.
rewrite J1; apply Rmult_le_compat_l; auto with real.
assert (0 < (powerRZ radix precision) - 1)%R.
apply Rplus_lt_reg_r with 1%R; ring_simplify.
replace 1%R with (powerRZ radix 0); auto with real zarith.
apply Rlt_powerRZ; try apply Rlt_IZR; auto with zarith.
apply Rlt_le_trans with (p*(((powerRZ radix precision)+1)/((powerRZ radix precision) -1)))%R.
apply Rle_lt_trans with (p*((2 + 2*powerRZ radix (1 - precision)) / (2 + powerRZ radix (- precision))))%R;
  [right; unfold Rdiv; ring|idtac].
apply Rmult_lt_compat_l.
apply Rlt_le_trans with (2:=G).
apply Q_positive with bo precision b b' p; auto.
apply Rmult_lt_reg_l with (powerRZ radix precision - 1)%R; auto with real.
apply Rlt_le_trans with (powerRZ radix precision + 1)%R;[idtac|right; field; auto with real].
apply Rmult_lt_reg_l with (2+powerRZ radix (-precision))%R; auto with real.
apply Rle_lt_trans with ((powerRZ radix precision - 1) *
     ((2 + 2 * powerRZ radix (1 - precision))))%R;[right; field; auto with real|idtac].
apply Rle_lt_trans with (2*powerRZ radix precision+2+-2*powerRZ radix (1-precision))%R;
  [right; ring_simplify|idtac].
rewrite Rmult_assoc.
repeat rewrite <- powerRZ_add; auto with real zarith.
ring_simplify (precision+(1-precision))%Z; rewrite powerRZ_1; unfold radix; ring.
apply Rlt_le_trans with (2 * powerRZ radix precision + 2 +1+powerRZ radix (-precision))%R.
repeat rewrite Rplus_assoc; repeat apply Rplus_lt_compat_l.
apply Rle_lt_trans with (1+0)%R; auto with real.
ring_simplify (1+0)%R; apply Rle_trans with 0%R; auto with real.
apply Ropp_le_cancel; ring_simplify ( - (-2 * powerRZ radix (1 - precision)))%R.
apply Rle_trans with (2*0)%R; auto with real zarith.
right; ring.
apply Rmult_le_compat_l; auto with real.
apply powerRZ_le, Rlt_IZR; auto with zarith.
apply Rplus_lt_compat_l.
apply powerRZ_lt, Rlt_IZR; auto with zarith.
right; ring_simplify.
rewrite <- powerRZ_add; auto with real zarith.
ring_simplify (precision+-precision)%Z; simpl; ring.
apply Rle_trans with (p+2*Fulp bo radix precision p)%R.
apply Rmult_le_reg_l with (powerRZ radix precision - 1)%R; auto with real.
apply Rle_trans with (p * ((powerRZ radix precision + 1)))%R;[right; field; auto with real|idtac].
apply Rle_trans with ((powerRZ radix precision -1)*p+2*p)%R;[right; ring|idtac].
apply Rle_trans with ((powerRZ radix precision -1)*p
  + 2*((powerRZ radix precision - 1)* Fulp bo radix precision p))%R;[idtac|right; ring].
apply Rplus_le_compat_l.
apply Rmult_le_compat_l; auto with real.
apply Rle_trans with (Rabs p);[apply RRle_abs|idtac].
unfold FtoRradix; apply FulpGe; auto.
apply Rplus_le_reg_l with (-(FSucc bo radix precision p))%R.
apply Rle_trans with (Fminus radix (FSucc bo radix precision (FSucc bo radix precision p))
  (FSucc bo radix precision p));[idtac|unfold FtoRradix; rewrite Fminus_correct; auto; right; ring].
unfold FtoRradix; rewrite FSuccDiff1; auto with zarith.
rewrite <- CanonicFulp with bo radix precision (FSucc bo radix precision p); auto with zarith.
2:apply FSuccCanonic; auto with zarith; left; auto.
apply Rle_trans with (-(Fminus radix (FSucc bo radix precision p) p)+ 2 * Fulp bo radix precision p)%R;
  [unfold FtoRradix; rewrite Fminus_correct; auto; right; ring|idtac].
unfold FtoRradix; rewrite FSuccDiff1; auto with zarith.
rewrite <- CanonicFulp with bo radix precision p; auto with zarith.
2: left; auto.
ring_simplify.
apply LeFulpPos; auto with zarith.
apply FBoundedSuc; auto with zarith.
fold FtoRradix; apply P_positive with bo precision b b'; auto.
left; apply FSuccLt; auto with zarith.
assert (0 < nNormMin radix precision)%Z;[apply nNormPos; auto with zarith|idtac].
assert (0 <= Fnum p)%Z; auto with zarith.
apply LeR0Fnum with radix; auto with real zarith.
apply P_positive with bo precision b b'; auto.
assert (0 < nNormMin radix precision)%Z;[apply nNormPos; auto with zarith|idtac].
assert (0 <= Fnum (FSucc bo radix precision p))%Z; auto with zarith.
apply LeR0Fnum with radix; auto with real zarith.
apply Rle_trans with (FtoR radix p).
apply P_positive with bo precision b b'; auto.
left; apply FSuccLt; auto with zarith.
assert (Fexp d = Fexp q-1)%Z.
assert (Fexp q -1 <= Fexp d)%Z; auto with zarith.
apply Z.le_trans with (Fexp p-2)%Z; auto with zarith.
apply Z.le_trans with (Fexp (Float (Fnum p) (Fexp p-2))); auto with zarith.
apply Fcanonic_Rle_Zle with radix bo precision; auto with real zarith.
elim Np; intros Y1 Y2; elim Y1; intros; left; split;try split; simpl; auto with zarith.
apply Rle_trans with (p/4)%R.
unfold Rdiv, FtoRradix, FtoR; simpl.
rewrite Rabs_mult; repeat rewrite Rabs_right.
unfold Zminus; rewrite powerRZ_add; auto with real zarith.
unfold powerRZ at 2; simpl; unfold radix; right; field.
apply Rle_ge; auto with real zarith.
apply powerRZ_le, Rlt_IZR; auto with zarith.
apply Rle_ge; assert (0 <= Fnum p)%Z; auto with real zarith.
apply LeR0Fnum with radix; auto with real zarith.
fold FtoRradix; apply P_positive with bo precision b b'; auto.
apply Rle_IZR; auto.
fold FtoRradix; rewrite dexact.
rewrite Rabs_right.
2: apply Rle_ge; apply Rplus_le_reg_l with q.
2: ring_simplify; auto with real.
apply Rle_trans with (p- p * (1 + powerRZ radix (1 - precision)) /
        (2 + powerRZ radix (- precision)))%R; auto with real.
2: unfold Rminus; apply Rplus_le_compat_l; auto with real.
unfold Rdiv; apply Rle_trans with (p*(1-(1 + powerRZ radix (1 - precision)) */
    (2 + powerRZ radix (- precision))))%R;[idtac|right; ring].
apply Rmult_le_compat_l.
apply P_positive with bo precision b b'; auto.
apply Rmult_le_reg_l with 4%R; auto with real.
apply Rmult_le_reg_l with (2 + powerRZ radix (- precision))%R; auto with real.
apply Rle_trans with (2 + powerRZ radix (- precision))%R;
  [right; field|idtac].
assert (0 < 4)%R; [apply Rlt_le_trans with 2%R|idtac]; auto with real.
apply Rle_trans with (4* (2 + powerRZ radix (- precision))-
  4*(1 + powerRZ radix (1 - precision)))%R;[idtac|right;field; auto with real].
unfold Zminus; rewrite powerRZ_add; auto with real zarith; rewrite powerRZ_1.
apply Rplus_le_reg_l with (-2+4*powerRZ radix (-precision))%R.
apply Rle_trans with (5*powerRZ radix (-precision))%R;[right; ring|idtac].
apply Rle_trans with 2%R;[idtac|simpl; unfold radix; right; ring].
apply Rle_trans with (powerRZ radix 0); auto with real zarith.
apply Rle_trans with (powerRZ radix (3-precision)); auto with real zarith.
unfold Zminus; rewrite powerRZ_add; auto with real zarith.
apply Rmult_le_compat_r; auto with real zarith.
apply powerRZ_le, Rlt_IZR; auto with zarith.
unfold powerRZ; simpl; unfold radix; repeat rewrite <- mult_IZR; auto with real zarith.
apply Rle_powerRZ; auto with real zarith.
simpl; auto with real.
case (Zle_lt_or_eq (nNormMin radix precision) (Fnum q)).
elim Nq; intros.
apply Zmult_le_reg_r with radix.
unfold radix; auto with zarith.
rewrite Zmult_comm; rewrite <- PosNormMin with radix bo precision; auto with zarith.
apply Z.le_trans with (1:=H4); rewrite Z.abs_eq; auto with zarith.
apply Z.mul_nonneg_nonneg.
unfold radix; omega.
assert (0 < Fnum q)%Z; auto with zarith.
apply LtR0Fnum with radix; auto with real zarith.
apply Q_positive with bo precision b b' p; auto.
intros P.
assert (Fnum q < nNormMin radix precision+2)%Z.
apply lt_IZR; rewrite plus_IZR; simpl.
apply Rmult_lt_reg_l with (powerRZ radix (Fexp q)); auto with real zarith.
apply powerRZ_lt, Rlt_IZR; auto with zarith.
apply Rle_lt_trans with q;[unfold FtoRradix, FtoR; auto with real|idtac].
apply Rplus_lt_reg_r with (-2*powerRZ radix (Fexp q))%R.
apply Rlt_le_trans with (powerRZ radix (Fexp q) * (nNormMin radix precision))%R;
  [idtac|right; ring].
apply Rle_lt_trans with d;[rewrite dexact|idtac].
apply Rplus_le_reg_l with q.
ring_simplify (q + (p - q))%R;right.
rewrite <- FPredSuc with bo radix precision p; auto with zarith.
2: left; auto.
replace (FSucc bo radix precision p) with qq.
rewrite FPredSimpl4; auto with zarith.
unfold FtoRradix, FtoR, Z.pred; simpl; rewrite J4; rewrite J5; ring_simplify.
rewrite plus_IZR; rewrite powerRZ_add; auto with real zarith.
rewrite powerRZ_1; unfold radix; ring.
assert (- pPred (vNum bo) < Fnum qq)%Z; auto with zarith.
apply Z.lt_trans with (nNormMin radix precision); auto with zarith.
assert (0 < pPred (vNum bo))%Z.
apply pPredMoreThanOne with radix precision; auto with zarith.
apply Z.lt_trans with 0%Z; auto with zarith.
apply nNormPos; auto with zarith.
apply FcanonicUnique with radix bo precision; auto with zarith real.
left; auto.
apply FSuccCanonic; auto with zarith.
left; auto.
fold FtoRradix; rewrite J1; auto with real.
apply Rle_lt_trans with (Rabs d);[apply RRle_abs|idtac].
unfold FtoRradix; rewrite <- Fabs_correct; auto.
unfold Fabs, FtoR; simpl.
apply Rlt_le_trans with (powerRZ 2 precision* powerRZ 2 (Fexp d))%R.
apply Rmult_lt_compat_r; auto with real zarith.
apply powerRZ_lt, Rlt_IZR; auto with zarith.
assert (Fbounded bo d);[rewrite dDef; elim Roundt; intros I1 I2; elim I1; auto|idtac].
elim H3; intros.
apply Rlt_le_trans with (Zpos (vNum bo)); try apply Rlt_IZR; auto with real zarith.
right; rewrite pGivesBound; rewrite Zpower_nat_Z_powerRZ; auto with real.
unfold nNormMin; repeat rewrite <- powerRZ_add; auto with real zarith.
rewrite Zpower_nat_Z_powerRZ; repeat rewrite <- powerRZ_add; auto with real zarith.
replace (Fexp q+pred precision)%Z with (precision+Fexp d)%Z; auto with real zarith.
assert (Fnum q= nNormMin radix precision + 1)%Z; auto with zarith.
clear P H3.
assert (p=Float (nNormMin radix precision) (Fexp q+1)).
rewrite <- FPredSuc with bo radix precision p; auto with zarith.
2: left; auto.
replace (FSucc bo radix precision p) with qq.
rewrite FPredSimpl4; auto with zarith.
rewrite J4; rewrite J5; rewrite H4; unfold Z.pred; auto with zarith.
ring_simplify (nNormMin radix precision + 1 + -1)%Z; auto.
assert (- pPred (vNum bo) < Fnum qq)%Z; auto with zarith.
rewrite J4; rewrite H4.
apply Z.lt_trans with (nNormMin radix precision); auto with zarith.
assert (0 < pPred (vNum bo))%Z.
apply pPredMoreThanOne with radix precision; auto with zarith.
apply Z.lt_trans with 0%Z; auto with zarith.
apply nNormPos; auto with zarith.
apply FcanonicUnique with radix bo precision; auto with zarith real.
left; auto.
apply FSuccCanonic; auto with zarith.
left; auto.
fold FtoRradix; rewrite J1; auto with real.
contradict Case2; apply Rlt_not_le.
assert (FtoRradix p=powerRZ radix (Fexp q)*powerRZ radix precision)%R.
rewrite H3; unfold FtoRradix, FtoR; simpl.
unfold nNormMin; rewrite Zpower_nat_Z_powerRZ; rewrite inj_pred; auto with zarith.
repeat rewrite <- powerRZ_add; auto with real zarith.
replace (Fexp q+precision)%Z with (Z.pred precision+(Fexp q+1))%Z;unfold Z.pred; auto with real zarith.
assert (FtoRradix q=powerRZ radix (Fexp q)*(powerRZ radix (precision-1)+1))%R.
unfold FtoRradix, FtoR; rewrite H4; rewrite plus_IZR; simpl.
unfold nNormMin; rewrite Zpower_nat_Z_powerRZ; rewrite inj_pred; auto with zarith.
unfold Z.pred, Zminus; simpl; ring.
assert (exists w1:float, Fbounded bo w1 /\ (FtoRradix w1=powerRZ radix (Fexp q)*
  (powerRZ radix precision+powerRZ radix (precision-1)-2))%R).
exists (Float (Zpower_nat radix (pred precision)+Zpower_nat radix (pred (pred precision))-1)
    (Fexp q+1)).
split;[split; simpl|idtac].
rewrite Z.abs_eq; auto with zarith.
rewrite pGivesBound; apply Z.lt_le_trans with
 (Zpower_nat radix (pred precision) + Zpower_nat radix (pred precision))%Z.
unfold Zminus; rewrite <- Zplus_assoc; apply Zplus_lt_compat_l.
apply Z.lt_le_trans with
 (Zpower_nat radix (pred (pred precision)) + Zpower_nat radix (pred (pred precision)))%Z.
apply Zplus_lt_compat_l; auto with zarith.
apply Z.lt_trans with 0%Z; auto with zarith.
apply Zpower_nat_less; auto with zarith.
pattern (pred precision) at 3; replace (pred precision) with (1+pred (pred precision)); auto with zarith.
rewrite Zpower_nat_is_exp.
replace (Zpower_nat radix 1) with 2%Z; auto with zarith.
pattern precision at 3; replace precision with (1+pred precision); auto with zarith.
rewrite Zpower_nat_is_exp.
replace (Zpower_nat radix 1) with 2%Z; auto with zarith.
assert (1 <= Zpower_nat radix (pred precision) +
    Zpower_nat radix (pred (pred precision)))%Z; auto with zarith.
apply Z.le_trans with (Zpower_nat radix (pred precision)+0)%Z; auto with zarith.
apply Z.le_trans with (Zpower_nat radix (pred precision))%Z; auto with zarith.
apply Zpower_NR1; unfold radix; auto with zarith.
apply Zplus_le_compat_l.
apply Zpower_NR0; unfold radix; auto with zarith.
elim J2; rewrite <- J5; auto.
unfold FtoRradix, FtoR; simpl.
unfold Zminus; repeat rewrite plus_IZR; simpl.
repeat rewrite Zpower_nat_Z_powerRZ.
repeat rewrite inj_pred; auto with zarith; unfold Z.pred.
repeat rewrite powerRZ_add; auto with real zarith; simpl.
unfold radix; ring_simplify (2*1)%R; field; auto with real.
assert (0 < 4)%R;[apply Rlt_le_trans with 2%R|idtac]; auto with real.
intros.
elim H7; intros w1 T1; elim T1; intros W11 W12; clear H7 T1.
apply Rle_lt_trans with w1.
apply Rle_trans with (Rabs u);[apply RRle_abs|idtac];unfold FtoRradix.
apply RoundAbsMonotoner with bo precision (EvenClosest bo radix precision) (3*Rabs t)%R;
   auto with zarith.
apply EvenClosestRoundedModeP; auto.
fold FtoRradix; rewrite <- dDef; rewrite dexact.
rewrite Rabs_right.
rewrite Rabs_right.
rewrite H5; rewrite H6.
apply Rle_trans with (w1-powerRZ radix (Fexp q))%R.
rewrite W12; right; ring_simplify.
unfold Zminus; repeat rewrite powerRZ_add; auto with real zarith; simpl.
unfold radix; ring_simplify (2*1)%R; field; auto with real.
apply Rle_trans with (w1-0)%R; auto with real.
unfold Rminus; apply Rplus_le_compat_l; auto with real zarith.
apply Ropp_le_contravar.
apply powerRZ_le, Rlt_IZR; auto with zarith.
apply Rle_ge; apply Rplus_le_reg_l with q.
ring_simplify; auto with real.
apply Rle_ge; apply Rle_trans with (3*0)%R; auto with real.
apply Rmult_le_compat_l; auto with real.
apply Rabs_pos.
assert (exists w2:float, Fbounded bo w2 /\ (FtoRradix w2=powerRZ radix (Fexp q)*
  (powerRZ radix precision+powerRZ radix (precision-1)))%R).
exists (Float (Zpower_nat radix (pred precision)+Zpower_nat radix (pred (pred precision)))
    (Fexp q+1)).
split;[split; simpl|idtac].
rewrite Z.abs_eq; auto with zarith.
rewrite pGivesBound; apply Z.lt_le_trans with
 (Zpower_nat radix (pred precision) + Zpower_nat radix (pred precision))%Z.
apply Zplus_lt_compat_l; auto with zarith.
apply Zpower_nat_monotone_lt; auto with zarith.
pattern precision at 3; replace precision with (1+pred precision); auto with zarith.
rewrite Zpower_nat_is_exp.
replace (Zpower_nat radix 1) with 2%Z; auto with zarith.
apply Z.add_nonneg_nonneg; apply Zpower_NR0; unfold radix; auto with zarith.
elim J2; rewrite <- J5; auto.
unfold FtoRradix, FtoR; simpl.
unfold Zminus; repeat rewrite plus_IZR; simpl.
repeat rewrite Zpower_nat_Z_powerRZ.
repeat rewrite inj_pred; auto with zarith; unfold Z.pred.
repeat rewrite powerRZ_add; auto with real zarith; simpl.
ring_simplify (2*1)%R; field; auto with real.
assert (0 < 4)%R;[apply Rlt_le_trans with 2%R|idtac]; auto with real.
elim H7; intros w2 T1; elim T1; intros W21 W22; clear H7 T1.
apply Rlt_le_trans with w2.
rewrite W12; rewrite W22; apply Rmult_lt_compat_l; auto with real zarith.
apply powerRZ_lt, Rlt_IZR; auto with zarith.
unfold Rminus; rewrite Rplus_assoc.
apply Rplus_lt_compat_l; auto with real zarith.
apply Rlt_le_trans with (powerRZ radix (precision - 1) + -0)%R; auto with real.
right; ring.
assert (0 <= p+q)%R.
apply Rle_trans with (0+0)%R;auto with real; apply Rplus_le_compat.
apply P_positive with bo precision b b'; auto.
left; apply Q_positive with bo precision b b' p; auto.
apply Rle_trans with (Rabs v).
unfold FtoRradix; apply RoundAbsMonotonel with bo precision
  (EvenClosest bo radix precision) (p+q)%R;auto with zarith.
apply EvenClosestRoundedModeP; auto.
fold FtoRradix; rewrite Rabs_right; auto with real.
apply Rle_trans with (w2+0)%R; auto with real zarith.
apply Rle_trans with (w2+powerRZ radix (Fexp q))%R; auto with real zarith.
apply Rplus_le_compat_l.
apply powerRZ_le, Rlt_IZR; auto with zarith.
rewrite W22; rewrite H5; rewrite H6; right;ring.
rewrite Rabs_right; auto with real; apply Rle_ge.
unfold FtoRradix; apply RleRoundedR0 with bo precision
   (EvenClosest bo radix precision) (p+q)%R; auto.
apply EvenClosestRoundedModeP; auto.
intros P.
assert (p=Float (pPred (vNum bo)) (Fexp q)).
rewrite <- FPredSuc with bo radix precision p; auto with zarith.
2: left; auto.
replace (FSucc bo radix precision p) with qq.
rewrite FPredSimpl2; auto with zarith.
replace (Z.pred (Fexp qq)) with (Fexp q); auto.
rewrite J5; unfold Z.pred; auto with zarith.
rewrite J5; assert (-dExp bo <= Fexp q)%Z; auto with zarith; apply Fq.
apply FcanonicUnique with radix bo precision; auto with zarith real.
left; auto.
apply FSuccCanonic; auto with zarith.
left; auto.
fold FtoRradix; rewrite J1; auto with real.
contradict M.
rewrite H3; simpl; auto with zarith.
intros M.
apply Rle_trans with (Rabs (-(dp-dq)))%R.
unfold delta; replace (d - (b * b' - a * c))%R with (-(dp-dq))%R; auto with real.
rewrite dexact; rewrite dpEq; rewrite dqEq; ring.
rewrite Rabs_Ropp.
unfold Rminus; apply Rle_trans with (Rabs dp + Rabs (-dq))%R;[apply Rabs_triang|rewrite Rabs_Ropp].
apply Rmult_le_reg_l with 2%R; auto with real.
apply Rle_trans with (INR 2*Rabs dp + INR 2*Rabs dq)%R;[right; simpl; ring|idtac].
apply Rle_trans with (Fulp bo radix precision p+Fulp bo radix precision q)%R;
  [apply Rplus_le_compat|idtac].
rewrite dpEq; unfold FtoRradix; apply ClosestUlp; auto with zarith.
elim Roundp; auto.
rewrite dqEq; unfold FtoRradix; apply ClosestUlp; auto with zarith.
elim Roundq; auto.
repeat rewrite CanonicFulp; auto with zarith.
2: left; auto.
2: left; auto.
rewrite M; unfold FtoR;simpl.
apply Rle_trans with (2* powerRZ radix (Fexp p))%R;[right; ring|idtac].
apply Rmult_le_compat_l; auto with real.
apply Rle_trans with (powerRZ radix (Fexp d+1));[idtac|
   rewrite powerRZ_add; auto with real zarith; rewrite powerRZ_1; unfold radix; right; ring].
apply Rle_powerRZ; auto with real.
assert (Fexp p-1 <= Fexp d)%Z; auto with zarith.
assert (exists f:float, (Fexp f=Fexp p-1)%Z /\
  (FtoRradix f=powerRZ radix (precision-2+Fexp p))%R
     /\ Fnormal radix bo f).
exists (Float (nNormMin radix precision) (Fexp p -1)).
split;simpl; auto with zarith.
split;[unfold FtoRradix, FtoR, nNormMin; simpl|idtac].
rewrite Zpower_nat_Z_powerRZ; rewrite inj_pred; auto with zarith.
rewrite <- powerRZ_add; auto with real zarith.
replace (Z.pred precision + (Fexp p - 1))%Z with (precision - 2 + Fexp p)%Z;
  [auto with real|unfold Z.pred; ring].
split;[split;simpl; auto with zarith|idtac].
rewrite Z.abs_eq; auto with zarith.
apply ZltNormMinVnum; auto with zarith.
assert (0 < nNormMin radix precision)%Z; auto with zarith.
apply nNormPos; auto with zarith.
simpl (Fnum (Float (nNormMin radix precision) (Fexp p - 1))).
rewrite <- PosNormMin with radix bo precision; auto with zarith.
elim H1; intros f H1'; elim H1'; intros K1 H2; elim H2; intros K2 K3;clear H1' H1 H2.
rewrite <- K1.
apply Fcanonic_Rle_Zle with radix bo precision; auto with zarith.
left; auto.
fold FtoRradix; rewrite dexact; rewrite K2.
rewrite Rabs_right;[idtac|apply Rle_ge; auto with real zarith].
rewrite Rabs_right.
2:apply Rle_ge;apply Rplus_le_reg_l with q.
2: ring_simplify; auto with real.
apply Rmult_le_reg_l with 3%R; auto with real.
apply Rle_trans with ((2*powerRZ radix (precision-1+Fexp p))*(1-powerRZ radix (-precision)))%R.
replace (powerRZ radix (precision - 1 + Fexp p)) with
   (2*powerRZ radix (precision - 2 + Fexp p))%R;
  [idtac|unfold Zminus; repeat rewrite powerRZ_add; auto with real zarith; simpl].
2: unfold radix; ring_simplify (2*1)%R; field; auto with real.
apply Rle_trans with ((2*2* (1 - powerRZ radix (- precision)))*
  powerRZ radix (precision - 2 + Fexp p))%R;[idtac|right; ring].
apply Rmult_le_compat_r; auto with real zarith.
apply powerRZ_le, Rlt_IZR; auto with zarith.
apply Rplus_le_reg_l with (-3+4*powerRZ radix (- precision))%R.
ring_simplify.
apply Rle_trans with (powerRZ radix (2+-precision)).
rewrite powerRZ_add; auto with real zarith; unfold powerRZ at 2.
simpl; unfold radix; right; ring.
apply Rle_trans with (powerRZ radix 0); auto with real zarith.
apply Rle_powerRZ; try apply Rle_IZR; unfold radix; auto with zarith.
2: apply powerRZ_le; try apply Rlt_IZR; auto with zarith.
assert (0 < (1 - powerRZ radix (- precision)))%R.
apply Rplus_lt_reg_r with (powerRZ radix (- precision)).
ring_simplify.
replace 1%R with (powerRZ radix 0);
    [ auto with real zarith | simpl in |- *; auto ].
apply Rlt_powerRZ; try apply Rlt_IZR; auto with zarith.
apply Rle_trans with (v* (1 - powerRZ radix (- precision)))%R.
apply Rmult_le_compat_r; auto with real.
assert (H3: exists g:float, Fbounded bo g /\
  (FtoRradix g= 2 * powerRZ radix (precision - 1 + Fexp p))%R).
exists (Float 1%Z (Fexp p +precision))%Z.
split;[split; simpl; auto with zarith|idtac].
apply vNumbMoreThanOne with radix precision; auto with zarith.
unfold FtoRradix, FtoR; simpl;unfold Zminus; repeat rewrite powerRZ_add; auto with real zarith.
simpl; unfold radix; ring_simplify (2*1)%R; field; auto with real.
clear K1 K2 K3 f; elim H3; intros g H1'; elim H1'; intros K1 K2; clear H3 H1'.
rewrite <- K2; unfold FtoRradix.
apply EvenClosestMonotone2 with bo precision g (p+q)%R; auto with real zarith.
rewrite K2; apply Rle_trans with ( powerRZ radix (precision - 1 + Fexp p)+
 powerRZ radix (precision - 1 + Fexp p))%R;[right; ring|idtac].
assert (forall f:float, Fnormal radix bo f -> (0 <= f)%R ->
  (powerRZ radix (precision - 1 + Fexp f) <= f)%R).
intros f Hf Hf'; rewrite powerRZ_add; auto with real zarith.
unfold FtoRradix, FtoR; apply Rmult_le_compat_r; auto with real zarith.
apply powerRZ_le; try apply Rlt_IZR; auto with zarith.
elim Hf; intros.
apply Rmult_le_reg_l with (radix); auto with real zarith.
apply Rle_trans with (Zpos (vNum bo)).
rewrite pGivesBound; rewrite Zpower_nat_Z_powerRZ.
unfold Zminus; rewrite powerRZ_add; auto with real zarith; simpl.
ring_simplify (2*1)%R; right; field; auto with real.
apply Rle_trans with (Z.abs (radix * Fnum f)); try apply Rle_IZR; auto with real zarith.
rewrite Zabs_Zmult; repeat rewrite Z.abs_eq; auto with real zarith.
rewrite mult_IZR; auto with real zarith.
apply LeR0Fnum with radix; auto with real zarith.
unfold radix; auto with zarith.
apply Rplus_le_compat;[apply H2; auto|idtac].
apply P_positive with bo precision b b'; auto.
rewrite <- M; apply H2; auto.
left; apply Q_positive with bo precision b b' p; auto.
unfold FtoRradix; apply RoundedModeProjectorIdem with bo; auto.
apply EvenClosestRoundedModeP; auto with zarith.
apply Rmult_le_reg_l with (/(1 - powerRZ radix (- precision)))%R; auto with real.
apply Rle_trans with (FtoRradix v);[right; field; auto with real|idtac].
replace (3 * (p - q))%R with (Rabs (3*Rabs t)).
apply Rle_trans with (Rabs (3*Rabs t) / (1 - powerRZ radix (- precision)))%R;
  [idtac|right; unfold Rdiv; ring].
apply Rle_trans with (Rabs v);[apply RRle_abs|idtac].
rewrite IneqEq.
apply RoundLeNormal; auto.
elim Roundu; auto.
rewrite <- dDef; rewrite dexact; auto.
rewrite Rabs_right;[idtac|apply Rle_ge].
rewrite Rabs_right;[idtac|apply Rle_ge]; auto with real.
apply Rplus_le_reg_l with q.
ring_simplify; auto with real.
apply Rle_trans with (0*0)%R; auto with real; apply Rmult_le_compat; auto with real.
apply Rabs_pos.
intros M.
apply Rle_trans with (Rabs (-(dp-dq)))%R.
unfold delta; replace (d - (b * b' - a * c))%R with (-(dp-dq))%R; auto with real.
rewrite dexact; rewrite dpEq; rewrite dqEq; ring.
rewrite Rabs_Ropp.
apply Rle_trans with (3 / 2 * Rmin (Fulp bo 2 precision p) (Fulp bo 2 precision q))%R.
apply dp_dq_le with a b b' c; auto.
replace (Rmin (Fulp bo 2 precision p) (Fulp bo 2 precision q))
  with (Fulp bo 2 precision q).
apply Rmult_le_compat; auto with real zarith.
lra.
unfold Fulp; auto with real zarith.
apply powerRZ_le; try apply Rlt_IZR; auto with zarith.
lra.
repeat rewrite CanonicFulp; auto with zarith.
right; rewrite M; auto with real.
left; auto.
repeat rewrite CanonicFulp; auto with zarith.
2: left; auto.
2: left; auto.
rewrite Rmin_comm.
unfold Rmin.
case (Rle_dec (FtoR 2 (Float (S 0) (Fexp q)))
         (FtoR 2 (Float (S 0) (Fexp p)))); auto with real.
intros N.
absurd (Fexp q <= Fexp p)%Z.
assert (Fexp p < Fexp q)%Z; auto with zarith.
apply Zlt_powerRZ with radix; auto with real zarith.
apply Rle_lt_trans with (FtoR 2 (Float (S 0) (Fexp p)));
  [right; unfold FtoR, radix; simpl; ring|idtac].
apply Rlt_le_trans with (FtoR 2 (Float (S 0) (Fexp q))); auto with real.
right; unfold FtoR, radix; simpl; ring.
apply Fcanonic_Rle_Zle with radix bo precision; auto with real zarith.
left; auto.
left; auto.
fold FtoRradix; rewrite Rabs_right.
rewrite Rabs_right; auto with real.
apply Rle_ge; apply P_positive with bo precision b b'; auto.
apply Rle_ge; left; apply Q_positive with bo precision b b' p; auto.
Qed.
End Discriminant1A.

Section Discriminant2A.
Variable bo : Fbound.
Variable precision : nat.

Let radix := 2%Z.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ TwoMoreThanOne).

Hypothesis precisionGreaterThanOne : 1 < precision.
Hypothesis pGivesBound : Zpos (vNum bo) = Zpower_nat radix precision.
Hypothesis precisionGreaterThanThree : 3 <= precision.

Variables a b b' c p q t d u v dp dq:float.

Let delta := (Rabs (d-(b*b'-a*c)))%R.

Hypothesis Fa : (Fbounded bo a).
Hypothesis Fb : (Fbounded bo b).
Hypothesis Fb': (Fbounded bo b').
Hypothesis Fc : (Fbounded bo c).
Hypothesis Fp : (Fbounded bo p).
Hypothesis Fq : (Fbounded bo q).
Hypothesis Fd : (Fbounded bo d).
Hypothesis Fu : (Fbounded bo u).
Hypothesis Fv : (Fbounded bo v).
Hypothesis Cand : (Fcanonic radix bo d).

There is no underflow
Hypothesis Nq:(Fnormal radix bo q).
Hypothesis Np:(Fnormal radix bo p).
Hypothesis Nv:(Fnormal radix bo v).
Hypothesis Nu:(Fnormal radix bo u).
Hypothesis U0: (- dExp bo <= Fexp p - 2)%Z.
Hypothesis U1: (- dExp bo <= Fexp q - 2)%Z.

Hypothesis Square:(0 <=b*b')%R.

Hypothesis Roundp : (EvenClosest bo radix precision (b*b')%R p).
Hypothesis Roundq : (EvenClosest bo radix precision (a*c)%R q).
Hypothesis Roundt : (EvenClosest bo radix precision (p-q)%R t).
Hypothesis dDef : d=t.
Hypothesis Roundu : (EvenClosest bo radix precision (3*Rabs t)%R u).
Hypothesis Roundv : (EvenClosest bo radix precision (p+q)%R v).
Hypothesis dpEq : (FtoRradix dp=b*b'-p)%R.
Hypothesis dqEq : (FtoRradix dq=a*c-q)%R.

Hypothesis Case1 : (3*(Rabs (p-q)) < p+q )%R.
Hypothesis Case2 : (v <= u )%R.

Theorem discri11: (delta <= 2*(Fulp bo radix precision d))%R.
case (Rle_or_lt q p); intros.
unfold delta, FtoRradix; apply discri10 with p q t u v dp dq; auto.
case (Rle_or_lt 0%R (a*c)%R)%R; intros.
unfold delta.
replace (d - (b * b' - a * c))%R with (-(Fopp d-(a*c-b*b')))%R;unfold FtoRradix;
  [idtac|rewrite Fopp_correct; ring].
rewrite Rabs_Ropp; apply Rle_trans with (2 * Fulp bo radix precision (Fopp d))%R.
generalize (EvenClosestSymmetric bo radix precision); unfold SymmetricP; intros L.
apply discri10 with q p (Fopp t) u v dq dp; auto with real.
apply FcanonicFopp; auto with zarith.
fold radix; fold FtoRradix; replace (q-p)%R with (-(p-q))%R;[apply L; auto|ring].
rewrite dDef; auto.
rewrite Fopp_correct; rewrite Rabs_Ropp; auto.
rewrite Rplus_comm; auto.
fold radix; fold FtoRradix; rewrite Rplus_comm with q p; apply Rle_lt_trans with (2:=Case1).
replace (q-p)%R with (-(p-q))%R;[rewrite Rabs_Ropp|ring]; auto with real.
unfold Fulp; rewrite Fnormalize_Fopp; unfold Fopp; simpl; auto with real zarith.
absurd (0<q)%R.
apply Rle_not_lt; unfold FtoRradix.
apply RleRoundedLessR0 with bo precision (EvenClosest bo radix precision) (a*c)%R;
  auto with real zarith.
apply EvenClosestRoundedModeP; auto with zarith.
apply Q_positive with bo precision b b' p; auto.
Qed.

End Discriminant2A.

Section Discriminant3A.
Variable bo : Fbound.
Variable precision : nat.

Let radix := 2%Z.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ TwoMoreThanOne).

Hypothesis precisionGreaterThanOne : 1 < precision.
Hypothesis pGivesBound : Zpos (vNum bo) = Zpower_nat radix precision.
Hypothesis precisionGreaterThanFour : 4 <= precision.

Variables a b b' c p q t dp dq s d u v:float.

Let delta := (Rabs (d-(b*b'-a*c)))%R.

Hypothesis Fa : (Fbounded bo a).
Hypothesis Fb : (Fbounded bo b).
Hypothesis Fb': (Fbounded bo b').
Hypothesis Fc : (Fbounded bo c).
Hypothesis Fp : (Fbounded bo p).
Hypothesis Fq : (Fbounded bo q).
Hypothesis Fd : (Fbounded bo d).
Hypothesis Ft : (Fbounded bo t).
Hypothesis Fs : (Fbounded bo s).
Hypothesis Fdp: (u<v)%R -> (Fbounded bo dp).
Hypothesis Fdq: (u<v)%R -> (Fbounded bo dq).

Hypotheses Cv: Fcanonic radix bo v.
Hypothesis Cs:(Fcanonic radix bo s).

There is no underflow
Hypothesis U1: (- dExp bo <= (Fexp t)-1)%Z.
Hypothesis U2: (powerRZ radix (-dExp bo+2*precision-1) <= Rabs (b*b'))%R.
Hypothesis U3: (powerRZ radix (-dExp bo+2*precision-1) <= Rabs (a*c))%R.

Hypothesis Np:(Fnormal radix bo p).
Hypothesis Nq:(Fnormal radix bo q).
Hypothesis Nd:(Fnormal radix bo d).
Hypothesis Nt:(Fnormal radix bo t).
Hypothesis Nu:(Fnormal radix bo u).
Hypothesis Nv:(Fnormal radix bo v).

Hypothesis Square:(0 <=b*b')%R.

Hypothesis Roundp : (EvenClosest bo radix precision (b*b')%R p).
Hypothesis Roundq : (EvenClosest bo radix precision (a*c)%R q).
Hypothesis Roundu : (EvenClosest bo radix precision (3*Rabs t)%R u).
Hypothesis Roundv : (EvenClosest bo radix precision (p+q)%R v).

Hypothesis Case1 : (p+q <= 3*(Rabs (p-q)))%R.
Hypothesis Case2 : (u < v )%R.

Hypothesis Roundt : (EvenClosest bo radix precision (p-q)%R t).
Hypothesis dpEq : (FtoRradix dp=b*b'-p)%R.
Hypothesis dqEq : (FtoRradix dq=a*c-q)%R.
Hypothesis Rounds : (EvenClosest bo radix precision (dp-dq)%R s).
Hypothesis Roundd : (EvenClosest bo radix precision (t+s)%R d).

Theorem RoundGeNormal: forall f:float, forall r:R,
  Closest bo radix r f -> (Fnormal radix bo f) ->
  (Rabs r <= Rabs f * (1 + powerRZ radix (- precision)))%R.
intros.
apply Rplus_le_reg_l with (-Rabs f)%R.
apply Rle_trans with (Rabs r-Rabs f)%R;[right; ring|idtac].
apply Rle_trans with (Rabs f * powerRZ radix (- precision) )%R;[idtac|right; ring].
apply Rle_trans with (Rabs (r-f));[apply Rabs_triang_inv|idtac].
apply Rmult_le_reg_l with (INR 2); auto with real.
apply Rle_trans with (Fulp bo radix precision f).
unfold FtoRradix; apply ClosestUlp; auto.
apply Rle_trans with (Rabs f*powerRZ radix (Z.succ (-precision)))%R.
unfold FtoRradix; apply FulpLe2; auto.
elim H0; auto.
rewrite FcanonicFnormalizeEq; auto with zarith.
left; auto.
right; unfold Z.succ; rewrite powerRZ_add; auto with real zarith.
rewrite powerRZ_1; unfold radix; simpl; ring.
Qed.

Theorem discri12: (q <= p)%R -> (delta <= 2*(Fulp bo radix precision d))%R.
intros M.
assert (0 < 3)%R.
apply Rlt_le_trans with 2%R; auto with real.
assert (0 <= p-q)%R.
apply Rplus_le_reg_l with q.
ring_simplify; auto.
assert (0 <= p)%R.
apply P_positive with bo precision b b'; auto.
assert (0 < (1 - powerRZ radix (- precision)))%R; auto with real.
apply Rplus_lt_reg_r with (powerRZ radix (- precision)).
ring_simplify.
replace 1%R with (powerRZ radix 0);
    [ auto with real zarith | simpl in |- *; auto ].
apply Rlt_powerRZ; try apply Rlt_IZR; auto with zarith.
assert (p-q <= Rsqr (1+powerRZ radix (-precision))/(1-powerRZ radix (-precision))*/3*(p+Rabs q))%R.
apply Rle_trans with (Rabs (p-q));[apply RRle_abs|idtac].
apply Rle_trans with (Rabs t * (1 + powerRZ radix (- precision)))%R.
apply RoundGeNormal; auto.
elim Roundt; auto.
apply Rle_trans with (((((p + Rabs q)/(1 - powerRZ radix (- precision)))
 * (1 + powerRZ radix (- precision)))*/3)* (1 + powerRZ radix (- precision)))%R;
  [idtac|right; unfold Rdiv, Rsqr; ring].
apply Rmult_le_compat_r; auto with real.
apply Rle_trans with (0+0)%R; try apply Rplus_le_compat; auto with real zarith.
apply powerRZ_le; try apply Rlt_IZR; auto with zarith.
apply Rle_trans with ((3*Rabs t)*/3)%R;
  [right; field; auto with real| apply Rmult_le_compat_r; auto with real].
apply Rle_trans with (Rabs (3*Rabs t));[apply RRle_abs|idtac].
apply Rle_trans with (Rabs u * (1 + powerRZ radix (- precision)))%R.
apply RoundGeNormal; auto; elim Roundu; auto.
apply Rmult_le_compat_r; auto with real.
apply Rle_trans with (0+0)%R; try apply Rplus_le_compat; auto with real zarith.
apply powerRZ_le; try apply Rlt_IZR; auto with zarith.
apply Rle_trans with (Rabs v).
cut (0 <= u)%R.
intros; repeat rewrite Rabs_right; auto with real.
apply Rle_ge; apply Rle_trans with (FtoRradix u); auto with real.
unfold FtoRradix; apply RleRoundedR0 with bo precision
  (EvenClosest bo radix precision) (3 * Rabs t)%R; auto with zarith.
apply EvenClosestRoundedModeP; auto with zarith.
apply Rle_trans with(3*0)%R; auto with real; apply Rmult_le_compat; auto with real.
apply Rabs_pos.
apply Rle_trans with (Rabs (p+q) / (1 - powerRZ radix (- precision)))%R.
unfold FtoRradix; apply RoundLeNormal with bo; auto with zarith real.
elim Roundv; auto.
unfold Rdiv; apply Rmult_le_compat_r; auto with real.
apply Rle_trans with (Rabs p+Rabs q)%R;[apply Rabs_triang|idtac].
apply Rplus_le_compat_r; rewrite Rabs_right; auto with real.
case (Rle_or_lt q 0); intros.
absurd (p <= q)%R.
apply Rlt_not_le; case M; auto with real.
intros; absurd (0 < v)%R.
apply Rle_not_lt.
unfold FtoRradix; apply RleRoundedLessR0 with bo precision
  (EvenClosest bo radix precision) (p+q)%R; auto with zarith.
apply EvenClosestRoundedModeP; auto with zarith.
apply Rle_trans with (1:=Case1); rewrite H5.
right; ring_simplify (p-p)%R; rewrite Rabs_R0; ring.
apply Rle_lt_trans with (FtoRradix u); auto with real.
unfold FtoRradix; apply RleRoundedR0 with bo precision
  (EvenClosest bo radix precision) (3 * Rabs t)%R; auto with zarith.
apply EvenClosestRoundedModeP; auto with zarith.
apply Rle_trans with(3*0)%R; auto with real; apply Rmult_le_compat; auto with real.
apply Rabs_pos.
apply Rmult_le_reg_l with (1- Rsqr (1 + powerRZ radix (- precision)) /
        (1 - powerRZ radix (- precision)) * / 3)%R.
apply Rplus_lt_reg_r with (Rsqr (1 + powerRZ radix (- precision)) /
    (1 - powerRZ radix (- precision)) * / 3)%R.
ring_simplify.
apply Rmult_lt_reg_l with 3%R; auto with real.
apply Rmult_lt_reg_l with (1 - powerRZ radix (- precision))%R; auto with real.
apply Rle_lt_trans with (Rsqr (1 + powerRZ radix (- precision)));
  [right; field; auto with real|ring_simplify (3*1)%R].
apply Rplus_lt_reg_r with (-1+3*powerRZ radix (- precision))%R.
apply Rle_lt_trans with (5* powerRZ radix (- precision)
  + powerRZ radix (- precision)* powerRZ radix (- precision))%R;[right; unfold Rsqr; ring|idtac].
apply Rlt_le_trans with 2%R;[idtac|right; ring].
apply Rlt_le_trans with (5 * powerRZ radix (- precision) +
    3*powerRZ radix (- precision))%R;
   [apply Rplus_lt_compat_l; apply Rmult_lt_compat_r; auto with real zarith|idtac].
apply powerRZ_lt; try apply Rlt_IZR; auto with zarith.
apply Rle_lt_trans with (powerRZ radix 1); auto with real zarith.
apply Rle_powerRZ; try apply Rle_IZR; try unfold radix; auto with zarith.
simpl; unfold radix; ring_simplify (2*1)%R; auto with real.
apply Rle_trans with (8 * powerRZ radix (- precision))%R;[right; ring|idtac].
replace 8%R with (powerRZ radix 3);[idtac|unfold powerRZ, radix; simpl; ring].
replace 2%R with (powerRZ radix 1);[idtac|unfold powerRZ, radix; simpl; ring].
rewrite <- powerRZ_add; auto with real zarith.
apply Rle_powerRZ; try apply Rle_IZR; try unfold radix; auto with zarith.
apply Rplus_le_reg_l with (-q+ ((Rsqr (1 + powerRZ radix (- precision)) /
     (1 - powerRZ radix (- precision)) * / 3) * p))%R.
apply Rle_trans with (p-q)%R;[right; ring|idtac].
apply Rle_trans with (1:=H3).
apply Rle_trans with (Rsqr (1 + powerRZ radix (- precision)) /
    (1 - powerRZ radix (- precision)) * / 3 * (p - q))%R;
    [idtac|right; ring].
replace (Rabs q) with (-q)%R;[right; ring|rewrite Rabs_left1; auto with real].
assert (forall f1:float, forall f2:float, forall i:Z,
      (0 <= i)%Z
   -> Fbounded bo f1 -> Fnormal radix bo f2
   -> (Rabs f1 <= powerRZ radix i*Rabs f2)%R
   -> (Fulp bo radix precision f1 <= powerRZ radix i*Fulp bo radix precision f2)%R).
intros.
rewrite CanonicFulp with bo radix precision f2; auto with zarith.
2: left; auto.
unfold FtoR; simpl; ring_simplify ((1 * powerRZ radix (Fexp f2)))%R.
unfold Fulp; rewrite <- powerRZ_add; auto with real zarith.
apply Rle_powerRZ; auto with real zarith.
apply Z.le_trans with (Fexp (Float (Fnum f2) (i + Fexp f2))); auto with zarith.
apply Fcanonic_Rle_Zle with radix bo precision; auto with real zarith.
apply FnormalizeCanonic; auto with real zarith.
elim H7; intros Y1 Y2; elim Y1; intros Y3 Y4.
left; split; try split; simpl; auto with zarith.
rewrite FnormalizeCorrect; auto with real zarith.
fold FtoRradix; apply Rle_trans with (1:=H8).
unfold FtoRradix; repeat rewrite <- Fabs_correct; auto.
unfold FtoR, Rabs; simpl.
rewrite powerRZ_add; auto with real zarith; right; ring.
assert (2*q <= p)%R.
apply Rmult_le_reg_l with 2%R; auto with real.
apply Rplus_le_reg_l with (-3*q+p)%R.
apply Rle_trans with (p+q)%R;[right; ring|apply Rle_trans with (1:=Case1)].
rewrite Rabs_right;[right; ring|apply Rle_ge; auto with real].
assert (q <= t)%R.
apply Rle_trans with (Rabs t);[idtac|rewrite Rabs_right; auto with real].
2:unfold FtoRradix; apply Rle_ge; apply RleRoundedR0 with bo precision
  (EvenClosest bo radix precision) (p-q)%R; auto with zarith.
2: apply EvenClosestRoundedModeP; auto with zarith.
unfold FtoRradix; apply RoundAbsMonotonel with bo precision
   (EvenClosest bo radix precision) (p-q)%R; auto with zarith.
apply EvenClosestRoundedModeP; auto with zarith.
fold FtoRradix; rewrite Rabs_right;[idtac|apply Rle_ge]; auto with real.
apply Rplus_le_reg_l with q; ring_simplify (q+(p-q))%R.
apply Rle_trans with (2*q)%R; auto with real; right; ring.
assert (Rabs s <= 3*Fulp bo radix precision q)%R.
assert (exists f:float, Fbounded bo f /\ (FtoRradix f= 3 * Fulp bo radix precision q)%R).
exists (Float 3 (Fexp q)).
split;[split; simpl; auto with zarith|idtac].
apply Z.lt_le_trans with (Zpower_nat radix 2); auto with zarith.
rewrite pGivesBound; auto with zarith.
apply Zpower_nat_monotone_le; auto with zarith.
apply Fq.
rewrite CanonicFulp; auto with real zarith.
2: left; auto.
unfold FtoRradix, FtoR; simpl; ring.
elim H8; intros f T; elim T; intros; clear T H8.
rewrite <- H10.
unfold FtoRradix; apply RoundAbsMonotoner with bo precision
   (EvenClosest bo radix precision) (dp-dq)%R; auto with zarith.
apply EvenClosestRoundedModeP; auto with zarith.
fold FtoRradix; unfold Rminus; apply Rle_trans with (Rabs dp+Rabs (-dq))%R;
   [apply Rabs_triang|rewrite Rabs_Ropp].
apply Rmult_le_reg_l with (INR 2); auto with real.
apply Rle_trans with ((INR 2)*Rabs dp+(INR 2)*Rabs dq)%R;[right; ring|idtac].
apply Rle_trans with (Fulp bo radix precision p+Fulp bo radix precision q)%R;
   [apply Rplus_le_compat|idtac].
rewrite dpEq; unfold FtoRradix; apply ClosestUlp; auto with zarith.
elim Roundp; auto.
rewrite dqEq; unfold FtoRradix; apply ClosestUlp; auto with zarith.
elim Roundq; auto.
apply Rle_trans with (5*Fulp bo radix precision q+Fulp bo radix precision q)%R;
  [apply Rplus_le_compat_r|rewrite H10; simpl; right; ring].
apply Rle_trans with (powerRZ radix 2*Fulp bo radix precision q)%R;
  [idtac|apply Rmult_le_compat_r; auto with real zarith].
2: unfold Fulp; auto with real zarith.
2: apply powerRZ_le; try apply Rlt_IZR; auto with zarith.
2: unfold powerRZ, radix; simpl; lra.
apply H5; auto with zarith.
apply Rle_trans with (3*Rabs q)%R;
  [idtac|apply Rmult_le_compat_r; auto with real zarith].
2: apply Rabs_pos.
2: unfold powerRZ, radix; simpl; lra.
apply Rmult_le_reg_l with (/2)%R; auto with real.
apply Rplus_le_reg_l with (/2*Rabs p)%R.
apply Rle_trans with (Rabs p);[right; field; auto with real|idtac].
rewrite Rabs_right;auto with real.
apply Rle_trans with (/2*(p+Rabs q) +Rabs q)%R;[idtac|right; field; auto with real].
apply Rle_trans with ((p-q)+Rabs q)%R.
rewrite Rabs_right; [right; ring|apply Rle_ge;auto with real].
apply Rplus_le_compat_r.
apply Rle_trans with (1:=H3).
apply Rmult_le_compat_r; auto with real.
apply Rle_trans with (0+0)%R; try apply Rplus_le_compat; auto with real.
apply Rabs_pos.
apply Rmult_le_reg_l with 2%R; auto with real.
apply Rmult_le_reg_l with 3%R; auto with real.
apply Rmult_le_reg_l with (1 - powerRZ radix (- precision))%R; auto with real.
apply Rle_trans with (Rsqr (1 + powerRZ radix (- precision))*2)%R;
  [right; field; auto with real|idtac].
apply Rle_trans with (3*(1 - powerRZ radix (- precision)))%R;[idtac|right; field; auto with real].
apply Rplus_le_reg_l with (-2+3*powerRZ radix (- precision))%R.
apply Rle_trans with (7*powerRZ radix (- precision)
  +2*powerRZ radix (- precision)*powerRZ radix (- precision))%R;[right; unfold Rsqr; ring|idtac].
apply Rle_trans with 1%R;[idtac|right; ring].
apply Rle_trans with (7 * powerRZ radix (- precision) + powerRZ radix (- precision))%R;
   [apply Rplus_le_compat_l|idtac].
replace 2%R with (powerRZ radix 1); auto with real zarith.
repeat rewrite <- powerRZ_add; auto with real zarith.
apply Rle_powerRZ; try apply Rle_IZR; try unfold radix; auto with zarith.
simpl; unfold radix; ring.
apply Rle_trans with (powerRZ radix 3 * powerRZ radix (- precision))%R;
  [unfold powerRZ at 3; unfold radix; right; simpl; ring | idtac].
rewrite <- powerRZ_add; auto with real zarith.
apply Rle_trans with (powerRZ radix 0); auto with real zarith.
apply Rle_powerRZ; try apply Rle_IZR; try unfold radix; auto with zarith.
assert (t-3*Fulp bo radix precision t <= Rabs d)%R.
assert (exists f:float, Fbounded bo f /\
   (FtoRradix f= t-3 * Fulp bo radix precision t)%R).
exists (Float (Fnum t-3) (Fexp t)).
split;[split; simpl; auto with zarith|idtac].
assert (0 < Fnum t)%Z;[apply LtR0Fnum with radix; auto with real zarith|idtac].
fold FtoRradix; apply Rlt_le_trans with q; auto.
case (Zle_or_lt 0 (Fnum t-3)%Z); intros.
rewrite Z.abs_eq; auto.
apply Z.le_lt_trans with (Z.abs (Fnum t)); auto with zarith.
apply Ft.
rewrite <- Zabs_Zopp; rewrite Z.abs_eq; auto with zarith.
apply Z.lt_le_trans with 3%Z; auto with zarith.
apply Z.le_trans with (nNormMin radix precision); auto with zarith.
unfold nNormMin; apply Z.le_trans with (Zpower_nat radix 2); auto with zarith.
simpl; auto with zarith.
apply Zpower_nat_monotone_le; auto with zarith.
apply nNrMMimLevNum; auto with zarith.
rewrite CanonicFulp; auto with zarith.
2: left; auto.
unfold FtoRradix, FtoR, Zminus; simpl; rewrite plus_IZR; simpl; ring.
elim H9; intros f T; elim T; intros; clear H9 T.
rewrite <- H11.
unfold FtoRradix; apply RoundAbsMonotonel with bo precision
   (EvenClosest bo radix precision) (t+s)%R; auto with zarith.
apply EvenClosestRoundedModeP; auto with zarith.
fold FtoRradix; replace (t+s)%R with (t-(-s))%R;[idtac|ring].
apply Rle_trans with (Rabs t -Rabs (-s))%R;[idtac|apply Rabs_triang_inv].
rewrite Rabs_Ropp; rewrite Rabs_right.
rewrite H11; unfold Rminus; apply Rplus_le_compat_l; auto with real.
apply Ropp_le_contravar; apply Rle_trans with (1:=H8).
apply Rmult_le_compat_l; auto with real.
apply LeFulpPos; auto with real.
apply Rle_ge; apply Rle_trans with q; auto with real.
assert (Fulp bo radix precision t <= 2*Fulp bo radix precision d)%R.
replace 2%R with (powerRZ radix 1); auto with real zarith.
apply H5; auto with zarith.
apply Rmult_le_reg_l with (/2)%R; auto with real.
apply Rle_trans with (Rabs d);[idtac|right; unfold powerRZ, radix; simpl; field; auto with real].
apply Rle_trans with (2:=H9).
rewrite Rabs_right;[idtac|apply Rle_ge; apply Rle_trans with q; auto with real].
apply Rle_trans with (t-t/2)%R;[right; field; auto with real|idtac].
unfold Rminus; apply Rplus_le_compat_l; apply Ropp_le_contravar.
apply Rle_trans with (3*(Rabs (FtoR radix t) * powerRZ radix (Z.succ (- precision))))%R.
apply Rmult_le_compat_l; auto with real.
apply FulpLe2; auto with real zarith.
rewrite FcanonicFnormalizeEq; auto with zarith.
left; auto.
fold FtoRradix; apply Rle_trans with (Rabs t*(3*powerRZ radix (Z.succ (- precision))))%R;
   [right; ring|idtac].
rewrite Rabs_right;[idtac|apply Rle_ge; apply Rle_trans with q; auto with real].
unfold Rdiv; apply Rmult_le_compat_l; auto with real.
apply Rle_trans with q; auto with real.
apply Rle_trans with (4*powerRZ radix (Z.succ (- precision)))%R;
  [apply Rmult_le_compat_r; auto with real zarith|idtac].
apply powerRZ_le; try apply Rlt_IZR; auto with zarith.
replace 4%R with (powerRZ radix 2).
2: unfold powerRZ, radix; simpl; ring.
replace (/2)%R with (powerRZ radix (-1)).
2: unfold powerRZ, radix; ring_simplify (2*1)%R; auto with real.
2: unfold powerRZ, radix; simpl; ring.
rewrite <- powerRZ_add; auto with real zarith.
unfold Z.succ; apply Rle_powerRZ; auto with real zarith.
unfold delta.
replace (d-(b*b'-a*c))%R with
   (-((t+s)-d)+-((p-q)-t)+-((dp-dq)-s))%R;[idtac|rewrite dpEq; rewrite dqEq; ring].
apply Rle_trans with (Rabs (- (t + s - d) + (- (p - q - t))) + Rabs (- (dp - dq - s)))%R;
  [apply Rabs_triang|rewrite Rabs_Ropp].
apply Rle_trans with (Rabs (- (t + s - d)) + Rabs (- (p - q - t)) + Rabs ((dp - dq - s)))%R;
  [apply Rplus_le_compat_r; apply Rabs_triang|repeat rewrite Rabs_Ropp].
apply Rmult_le_reg_l with (INR 2); auto with real.
apply Rle_trans with ((INR 2)*Rabs ((t + s - d)) + (INR 2)*Rabs ((p - q - t))
   + (INR 2)*Rabs ((dp - dq - s)))%R;[right; ring|idtac].
apply Rle_trans with (Fulp bo radix precision d+Fulp bo radix precision t
  +Fulp bo radix precision s)%R;[apply Rplus_le_compat|idtac].
apply Rplus_le_compat.
unfold FtoRradix; apply ClosestUlp; auto with zarith; elim Roundd; auto.
unfold FtoRradix; apply ClosestUlp; auto with zarith; elim Roundt; auto.
unfold FtoRradix; apply ClosestUlp; auto with zarith; elim Rounds; auto.
apply Rle_trans with (Fulp bo radix precision d + 2*Fulp bo radix precision d +
    Fulp bo radix precision s)%R; auto with real.
apply Rle_trans with (Fulp bo radix precision d + 2 * Fulp bo radix precision d +
    Fulp bo radix precision d)%R;[idtac|right; simpl; ring].
apply Rplus_le_compat_l.
rewrite FulpFabs with bo radix precision s; auto.
rewrite FulpFabs with bo radix precision d; auto.
apply LeFulpPos; auto with zarith real.
apply absFBounded; auto.
apply absFBounded; auto.
rewrite Fabs_correct; auto with real zarith.
apply Rabs_pos.
repeat rewrite Fabs_correct; auto; fold FtoRradix.
apply Rle_trans with (1:=H8); apply Rle_trans with (2:=H9).
apply Rle_trans with (3 * Fulp bo radix precision t)%R.
apply Rmult_le_compat_l; auto with real.
apply LeFulpPos; auto with zarith real.
apply Rplus_le_reg_l with (3* Fulp bo radix precision t)%R.
ring_simplify (3 * Fulp bo radix precision t + (t - 3 * Fulp bo radix precision t))%R.
apply Rle_trans with (6*Fulp bo radix precision t)%R;[right; ring|idtac].
apply Rle_trans with (6*(Rabs (FtoR radix t) * powerRZ radix (Z.succ (- precision))))%R.
apply Rmult_le_compat_l; auto with real.
apply FulpLe2; auto with real zarith.
rewrite FcanonicFnormalizeEq; auto with zarith.
left; auto.
fold FtoRradix; apply Rle_trans with (Rabs t*(6*powerRZ radix (Z.succ (- precision))))%R;
   [right; ring|idtac].
rewrite Rabs_right;[idtac|apply Rle_ge; apply Rle_trans with q; auto with real].
apply Rle_trans with (t*1)%R; auto with real.
apply Rmult_le_compat_l; auto with real.
apply Rle_trans with q; auto with real.
apply Rle_trans with (8*powerRZ radix (Z.succ (- precision)))%R;
  [apply Rmult_le_compat_r; auto with real zarith|idtac].
apply powerRZ_le; try apply Rlt_IZR; auto with zarith.
replace 8%R with (powerRZ radix 3).
2: unfold powerRZ, radix; simpl; ring.
replace 1%R with (powerRZ radix 0).
2: simpl; auto with real.
rewrite <- powerRZ_add; auto with real zarith.
apply Rle_powerRZ; auto with real zarith.
Qed.
End Discriminant3A.

Section Discriminant4A.
Variable bo : Fbound.
Variable precision : nat.

Let radix := 2%Z.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ TwoMoreThanOne).

Hypothesis precisionGreaterThanOne : 1 < precision.
Hypothesis pGivesBound : Zpos (vNum bo) = Zpower_nat radix precision.
Hypothesis precisionGreaterThanFour : 4 <= precision.

Variables a b b' c p q t dp dq s d u v:float.

Let delta := (Rabs (d-(b*b'-a*c)))%R.

Hypothesis Fa : (Fbounded bo a).
Hypothesis Fb : (Fbounded bo b).
Hypothesis Fb': (Fbounded bo b').
Hypothesis Fc : (Fbounded bo c).
Hypothesis Fp : (Fbounded bo p).
Hypothesis Fq : (Fbounded bo q).
Hypothesis Fd : (Fbounded bo d).
Hypothesis Ft : (Fbounded bo t).
Hypothesis Fs : (Fbounded bo s).
Hypothesis Fdp: (u<v)%R -> (Fbounded bo dp).
Hypothesis Fdq: (u<v)%R -> (Fbounded bo dq).

Hypotheses Cv: Fcanonic radix bo v.
Hypotheses Cs: Fcanonic radix bo s.

There is no underflow
Hypothesis U1: (- dExp bo <= (Fexp t)-1)%Z.
Hypothesis U2: (powerRZ radix (-dExp bo+2*precision-1) <= Rabs (b*b'))%R.
Hypothesis U3: (powerRZ radix (-dExp bo+2*precision-1) <= Rabs (a*c))%R.

Hypothesis Np:(Fnormal radix bo p).
Hypothesis Nq:(Fnormal radix bo q).
Hypothesis Nd:(Fnormal radix bo d).
Hypothesis Nt:(Fnormal radix bo t).
Hypothesis Nu:(Fnormal radix bo u).
Hypothesis Nv:(Fnormal radix bo v).

Hypothesis Square:(0 <=b*b')%R.

Hypothesis Roundp : (EvenClosest bo radix precision (b*b')%R p).
Hypothesis Roundq : (EvenClosest bo radix precision (a*c)%R q).
Hypothesis Roundu : (EvenClosest bo radix precision (3*Rabs t)%R u).
Hypothesis Roundv : (EvenClosest bo radix precision (p+q)%R v).

Hypothesis Case1 : (p+q <= 3*(Rabs (p-q)))%R.
Hypothesis Case2 : (u < v )%R.

Hypothesis Roundt : (EvenClosest bo radix precision (p-q)%R t).
Hypothesis dpEq : (FtoRradix dp=b*b'-p)%R.
Hypothesis dqEq : (FtoRradix dq=a*c-q)%R.
Hypothesis Rounds : (EvenClosest bo radix precision (dp-dq)%R s).
Hypothesis Roundd : (EvenClosest bo radix precision (t+s)%R d).

Theorem discri13: (delta <= 2*(Fulp bo radix precision d))%R.
case (Rle_or_lt q p); intros.
unfold delta, FtoRradix; apply discri12 with p q t dp dq s u v; auto with zarith.
case (Rle_or_lt 0%R (a*c)%R)%R; intros.
unfold delta.
replace (d - (b * b' - a * c))%R with (-(Fopp d-(a*c-b*b')))%R;unfold FtoRradix;
  [idtac|rewrite Fopp_correct; ring].
rewrite Rabs_Ropp; apply Rle_trans with (2 * Fulp bo radix precision (Fopp d))%R.
generalize (EvenClosestSymmetric bo radix precision); unfold SymmetricP; intros L.
apply discri12 with q p (Fopp t) dq dp (Fopp s) u v; auto with real zarith.
apply oppBounded; auto.
apply oppBounded; auto.
apply oppBounded; auto.
apply FnormalFop; auto.
apply FnormalFop; auto.
rewrite Fopp_correct; rewrite Rabs_Ropp; auto.
rewrite Rplus_comm; auto.
rewrite Rplus_comm; fold radix; fold FtoRradix;
  replace (q-p)%R with (-(p-q))%R;[rewrite Rabs_Ropp|ring]; auto with real.
fold radix; fold FtoRradix; replace (q-p)%R with (-(p-q))%R;[apply L; auto|ring].
fold radix; fold FtoRradix; replace (dq-dp)%R with (-(dp-dq))%R;[apply L; auto|ring].
repeat rewrite Fopp_correct; fold radix; fold FtoRradix.
replace (-t+-s)%R with (-(t+s))%R;[apply L; auto|ring].
unfold Fulp; rewrite Fnormalize_Fopp; unfold Fopp; simpl; auto with real zarith.
absurd (0<q)%R.
apply Rle_not_lt; unfold FtoRradix.
apply RleRoundedLessR0 with bo precision (EvenClosest bo radix precision) (a*c)%R;
  auto with real zarith.
apply EvenClosestRoundedModeP; auto with zarith.
apply Rle_lt_trans with p; auto with real.
apply P_positive with bo precision b b'; auto.
Qed.

End Discriminant4A.

Section Discriminant5.
Variable bo : Fbound.
Variable precision : nat.

Let radix := 2%Z.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ TwoMoreThanOne).

Hypothesis precisionGreaterThanOne : 1 < precision.
Hypothesis pGivesBound : Zpos (vNum bo) = Zpower_nat radix precision.
Hypothesis precisionGreaterThanFour : 4 <= precision.

Variables a b c p q t dp dq s d u v:float.

Let delta := (Rabs (d-(b*b-a*c)))%R.

Hypothesis Fa : (Fbounded bo a).
Hypothesis Fb : (Fbounded bo b).
Hypothesis Fc : (Fbounded bo c).
Hypothesis Fp : (Fbounded bo p).
Hypothesis Fq : (Fbounded bo q).
Hypothesis Fd : (Fbounded bo d).
Hypothesis Ft : (Fbounded bo t).
Hypothesis Fs : (u<v)%R -> (Fbounded bo s).
Hypothesis Fdp: (u<v)%R -> (Fbounded bo dp).
Hypothesis Fdq: (u<v)%R -> (Fbounded bo dq).
Hypothesis Fu: (Fbounded bo u).
Hypothesis Fv: (Fbounded bo v).
Hypothesis Cs: (u < v)%R -> (Fcanonic radix bo s).

There is no underflow
Hypothesis U0: (- dExp bo <= Fexp d - 1)%Z.
Hypothesis U1: (- dExp bo <= (Fexp t)-1)%Z.
Hypothesis U2: (powerRZ radix (-dExp bo+2*precision-1) <= Rabs (b*b))%R.
Hypothesis U3: (powerRZ radix (-dExp bo+2*precision-1) <= Rabs (a*c))%R.

Hypothesis Np:(Fnormal radix bo p).
Hypothesis Nq:(Fnormal radix bo q).
Hypothesis Nd:(Fnormal radix bo d).
Hypothesis Nu:(Fnormal radix bo u).
Hypothesis Nv:(Fnormal radix bo v).
Hypothesis Nt:(Fnormal radix bo t).

Hypothesis Roundp : (EvenClosest bo radix precision (b*b)%R p).
Hypothesis Roundq : (EvenClosest bo radix precision (a*c)%R q).
Hypothesis Roundt : (EvenClosest bo radix precision (p-q)%R t).
Hypothesis Roundu : (EvenClosest bo radix precision (3*Rabs t)%R u).
Hypothesis Roundv : (EvenClosest bo radix precision (p+q)%R v).

Hypothesis FRoundd : (v <= u)%R ->
   (EvenClosest bo radix precision (p-q)%R d).

Hypothesis dpEq : (FtoRradix dp=b*b-p)%R.
Hypothesis dqEq : (FtoRradix dq=a*c-q)%R.
Hypothesis SRounds : (u < v)%R -> (EvenClosest bo radix precision (dp-dq)%R s).
Hypothesis SRoundd : (u < v)%R -> (EvenClosest bo radix precision (t+s)%R d).

Theorem discri14: (delta <= 2*(Fulp bo radix precision d))%R.
case (Rle_or_lt (p+q)%R (3*(Rabs (p-q)))%R); case (Rle_or_lt v u); intros.
unfold delta, FtoRradix.
apply discri9 with p q t dp dq s; auto;
  fold radix; fold FtoRradix; intros;
  absurd (p + q <= 3 * Rabs (p - q))%R; auto with real.
unfold delta, FtoRradix.
apply discri13 with p q t dp dq s u v; auto.
apply Rle_trans with (Rsqr (FtoR 2 b)); auto with real.
unfold delta, FtoRradix.
apply discri11 with p q t u v dp dq; auto with zarith.
left; auto.
assert (exists s : float,
         Fbounded bo s /\
         FtoR 2 s = (FtoR 2 b * FtoR 2 b - FtoR 2 p)%R /\
         Fexp s = (Fexp p - precision)%Z /\
         (Rabs (Fnum s) <= powerRZ (Zpos 2) (Z.pred precision))%R).
apply errorBoundedMultClosest_Can; auto.
elim Roundp; auto.
left; auto.
elim H1; intros s' T; elim T; intros T1 T2; elim T2; intros T3 T4; elim T4; intros.
apply Z.le_trans with (Fexp s')%Z; auto with zarith.
apply T.
assert (exists s : float,
         Fbounded bo s /\
         FtoR 2 s = (FtoR 2 a * FtoR 2 c - FtoR 2 q)%R /\
         Fexp s = (Fexp q - precision)%Z /\
         (Rabs (Fnum s) <= powerRZ (Zpos 2) (Z.pred precision))%R).
apply errorBoundedMultClosest_Can; auto.
elim Roundq; auto.
left; auto.
elim H1; intros s' T; elim T; intros T1 T2; elim T2; intros T3 T4; elim T4; intros.
apply Z.le_trans with (Fexp s')%Z; auto with zarith.
apply T.
fold FtoRradix; apply Rle_trans with (Rsqr b); auto with real.
apply FcanonicUnique with radix bo precision; auto with zarith.
left; auto.
left; auto.
generalize EvenClosestUniqueP; unfold UniqueP; intros Y.
apply Y with bo precision (p-q)%R; auto with real zarith.
unfold delta, FtoRradix.
apply discri9 with p q t dp dq s; auto.
fold radix; fold FtoRradix; intros.
absurd (p + q <= 3 * Rabs (p - q))%R; auto with real.
Qed.

End Discriminant5.

Section Discriminant6.
Variable bo : Fbound.
Variable precision : nat.

Let radix := 2%Z.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ TwoMoreThanOne).

Hypothesis precisionGreaterThanOne : 1 < precision.
Hypothesis pGivesBound : Zpos (vNum bo) = Zpower_nat radix precision.
Hypothesis precisionGreaterThanFour : 4 <= precision.

Variables a b c p q t dp dq s d u v:float.

Let delta := (Rabs (d-(b*b-a*c)))%R.

Hypothesis Fa : (Fbounded bo a).
Hypothesis Fb : (Fbounded bo b).
Hypothesis Fc : (Fbounded bo c).
Hypothesis Fdp: (u<v)%R -> (Fbounded bo dp).
Hypothesis Fdq: (u<v)%R -> (Fbounded bo dq).

There is no underflow

Hypothesis U1: (powerRZ radix (-dExp bo+2*precision-1) <= Rabs (b*b))%R.
Hypothesis U2: (powerRZ radix (-dExp bo+2*precision-1) <= Rabs (a*c))%R.
Hypothesis U4: (powerRZ radix (-dExp bo+precision) <= Rabs d)%R.
Hypothesis U5: (powerRZ radix (-dExp bo+precision-1) <= Rabs u)%R.
Hypothesis U6: (powerRZ radix (-dExp bo+precision-1) <= Rabs v)%R.
Hypothesis U7: (powerRZ radix (-dExp bo+precision) <= Rabs t)%R.

Hypothesis Roundp : (EvenClosest bo radix precision (b*b)%R p).
Hypothesis Roundq : (EvenClosest bo radix precision (a*c)%R q).
Hypothesis Roundt : (EvenClosest bo radix precision (p-q)%R t).
Hypothesis Roundu : (EvenClosest bo radix precision (3*Rabs t)%R u).
Hypothesis Roundv : (EvenClosest bo radix precision (p+q)%R v).

Hypothesis FRoundd : (v <= u)%R ->
   (EvenClosest bo radix precision (p-q)%R d).

Hypothesis dpEq : (FtoRradix dp=b*b-p)%R.
Hypothesis dqEq : (FtoRradix dq=a*c-q)%R.
Hypothesis SRounds : (u < v)%R -> (EvenClosest bo radix precision (dp-dq)%R s).
Hypothesis SRoundd : (u < v)%R -> (EvenClosest bo radix precision (t+s)%R d).

Theorem discri15: (delta <= 2*(Fulp bo radix precision d))%R.
assert (forall r:R, forall f:float,
  EvenClosest bo radix precision r f -> Fbounded bo f).
intros r f T; elim T; intros T1 T2; elim T1; auto with zarith.
assert (forall f:float, (Fbounded bo f) ->
  (powerRZ radix (- dExp bo + precision - 1) <= Rabs f)%R
  -> Fnormal radix bo (Fnormalize radix bo precision f)).
intros.
assert (Fcanonic radix bo (Fnormalize radix bo precision f)).
apply FnormalizeCanonic; auto with zarith.
case H2; auto; intros.
absurd (FtoR radix (Fabs f) < FtoR radix (firstNormalPos radix bo precision))%R.
apply Rle_not_lt; rewrite Fabs_correct; auto.
fold FtoRradix; apply Rle_trans with (2:=H1).
unfold firstNormalPos, FtoRradix, FtoR, nNormMin; simpl.
rewrite Zpower_nat_Z_powerRZ; rewrite inj_pred; auto with zarith; unfold Z.pred.
unfold Zminus; repeat rewrite powerRZ_add; auto with real zarith.
right; simpl; ring.
apply Rle_lt_trans with (FtoR radix (Fabs (Fnormalize radix bo precision f))).
repeat rewrite Fabs_correct; auto.
rewrite FnormalizeCorrect; auto with zarith real.
apply FsubnormalLtFirstNormalPos; auto with zarith.
apply FsubnormFabs; auto.
rewrite Fabs_correct; auto with zarith real.
apply Rabs_pos.
unfold delta; unfold FtoRradix.
assert (Fbounded bo d).
case (Rle_or_lt v u); intros.
elim FRoundd; auto; intros T T'; elim T; auto with zarith.
elim SRoundd; auto; intros T T'; elim T; auto with zarith.
apply Rle_trans with (2 * Fulp bo radix precision
  (Fnormalize radix bo precision d))%R.
2: unfold Fulp; rewrite FcanonicFnormalizeEq with radix bo
  precision (Fnormalize radix bo precision d); auto with zarith real.
2: apply FnormalizeCanonic; auto with zarith.
rewrite <- FnormalizeCorrect with radix bo precision d; auto with zarith.
assert (Fnormal radix bo (Fnormalize radix bo precision d)).
apply H0; auto.
apply Rle_trans with (2:=U4); auto with real zarith.
apply Rle_powerRZ; try apply Rle_IZR; try unfold radix; auto with zarith.
assert (Fnormal radix bo (Fnormalize radix bo precision t)).
apply H0.
apply H with (p-q)%R; auto.
apply Rle_trans with (2:=U7); auto with real zarith.
apply Rle_powerRZ; try apply Rle_IZR; try unfold radix; auto with zarith.
apply discri14 with (Fnormalize radix bo precision p)
  (Fnormalize radix bo precision q) (Fnormalize radix bo precision t)
  dp dq (Fnormalize radix bo precision s) (Fnormalize radix bo precision u)
  (Fnormalize radix bo precision v); auto.
apply FnormalizeBounded; auto with zarith; apply H with (b*b)%R; auto.
apply FnormalizeBounded; auto with zarith; apply H with (a*c)%R; auto.
apply FnormalizeBounded; auto with zarith.
apply FnormalizeBounded; auto with zarith; apply H with (p-q)%R; auto.
repeat rewrite FnormalizeCorrect; auto with real zarith; intros.
apply FnormalizeBounded; auto with zarith; apply H with (dp-dq)%R; auto.
repeat rewrite FnormalizeCorrect; auto with real zarith.
repeat rewrite FnormalizeCorrect; auto with real zarith.
apply FnormalizeBounded; auto with zarith; apply H with (3*Rabs t)%R; auto.
apply FnormalizeBounded; auto with zarith; apply H with (p+q)%R; auto.
intros; apply FnormalizeCanonic; auto with zarith.
apply H with (dp-dq)%R; auto.
apply SRounds.
rewrite FnormalizeCorrect in H4; auto with zarith real.
rewrite FnormalizeCorrect in H4; auto with zarith real.
assert (- dExp bo + precision < Fexp (Fnormalize radix bo precision d)
  + precision)%Z; auto with zarith.
apply Zlt_powerRZ with radix; auto with real zarith.
apply Rle_lt_trans with (1:=U4).
unfold FtoRradix; rewrite <- FnormalizeCorrect with radix bo precision d;
  auto with zarith.
rewrite <- Fabs_correct; auto; unfold FtoR, Fabs; simpl.
rewrite Rmult_comm; rewrite powerRZ_add; auto with real zarith.
apply Rmult_lt_compat_l; auto with real zarith.
apply powerRZ_lt; try apply Rlt_IZR; auto with zarith.
apply Rlt_le_trans with (Zpos (vNum bo)); auto with zarith real.
apply Rlt_IZR; apply H2.
rewrite pGivesBound; rewrite Zpower_nat_Z_powerRZ; auto with real zarith.
assert (- dExp bo + precision < Fexp (Fnormalize radix bo precision t)
  + precision)%Z; auto with zarith.
apply Zlt_powerRZ with radix; auto with real zarith.
apply Rle_lt_trans with (1:=U7).
unfold FtoRradix; rewrite <- FnormalizeCorrect with radix bo precision t;
  auto with zarith.
rewrite <- Fabs_correct; auto; unfold FtoR, Fabs; simpl.
rewrite Rmult_comm; rewrite powerRZ_add; auto with real zarith.
apply Rmult_lt_compat_l; auto with real zarith.
apply powerRZ_lt; try apply Rlt_IZR; auto with zarith.
apply Rlt_le_trans with (Zpos (vNum bo)); auto with zarith real.
apply Rlt_IZR; apply H3.
rewrite pGivesBound; rewrite Zpower_nat_Z_powerRZ; auto with real zarith.
assert (Fbounded bo (Fnormalize radix bo precision t));
  auto with zarith real.
elim H3; auto.
apply H0; auto.
apply H with (b*b)%R; auto.
cut (exists f:float, Fbounded bo f /\
  (FtoRradix f=(powerRZ radix (- dExp bo + precision - 1)))%R).
intros T; elim T; intros f T'; elim T'; intros; clear T T'.
rewrite <- H6; unfold FtoRradix.
apply RoundAbsMonotonel with bo precision
  (Closest bo radix) (b*b)%R; auto with zarith real.
apply ClosestRoundedModeP with precision; auto with zarith.
elim Roundp; auto.
fold FtoRradix; rewrite H6; auto.
apply Rle_trans with (2:=U1); auto with zarith real.
apply Rle_powerRZ; try apply Rle_IZR; unfold radix; auto with zarith.
exists (Float 1 (-dExp bo+precision-1)).
split;[split|idtac].
simpl; apply vNumbMoreThanOne with radix precision; auto with zarith.
apply Z.le_trans with (- dExp bo + precision - 1)%Z; auto with zarith.
unfold FtoRradix, FtoR; simpl; ring.
apply H0; auto.
apply H with (a*c)%R; auto.
cut (exists f:float, Fbounded bo f /\
  (FtoRradix f=(powerRZ radix (- dExp bo + precision - 1)))%R).
intros T; elim T; intros f T'; elim T'; intros; clear T T'.
rewrite <- H5; unfold FtoRradix.
apply RoundAbsMonotonel with bo precision
  (Closest bo radix) (a*c)%R; auto with zarith real.
apply ClosestRoundedModeP with precision; auto with zarith.
elim Roundq; auto.
fold FtoRradix; rewrite H5; auto.
apply Rle_trans with (2:=U2); auto with zarith real.
apply Rle_powerRZ; unfold radix; try apply Rle_IZR; auto with zarith.
exists (Float 1 (-dExp bo+precision-1)).
split;[split|idtac].
simpl; apply vNumbMoreThanOne with radix precision; auto with zarith.
apply Z.le_trans with (- dExp bo + precision - 1)%Z; auto with zarith.
unfold FtoRradix, FtoR; simpl; ring.
apply H0; auto; apply H with (3*Rabs t)%R; auto.
apply H0; auto; apply H with (p+q)%R; auto.
generalize (EvenClosestCompatible bo radix precision); unfold CompatibleP.
intros T; apply T with (b*b)%R p; auto with real zarith.
rewrite FnormalizeCorrect; auto with zarith real.
apply FnormalizeBounded; auto with zarith;apply H with (b*b)%R; auto.
generalize (EvenClosestCompatible bo radix precision); unfold CompatibleP.
intros T; apply T with (a*c)%R q; auto with real zarith.
rewrite FnormalizeCorrect; auto with zarith real.
apply FnormalizeBounded; auto with zarith;apply H with (a*c)%R; auto.
repeat rewrite FnormalizeCorrect; auto with zarith real.
generalize (EvenClosestCompatible bo radix precision); unfold CompatibleP.
intros T; apply T with (p-q)%R t; auto with real zarith.
rewrite FnormalizeCorrect; auto with zarith real.
apply FnormalizeBounded; auto with zarith;apply H with (p-q)%R; auto.
rewrite FnormalizeCorrect; auto with zarith real.
generalize (EvenClosestCompatible bo radix precision); unfold CompatibleP.
intros T; apply T with (3*Rabs t)%R u; auto with real zarith.
rewrite FnormalizeCorrect; auto with zarith real.
apply FnormalizeBounded; auto with zarith;apply H with (3*Rabs t)%R; auto.
repeat rewrite FnormalizeCorrect; auto with zarith real.
generalize (EvenClosestCompatible bo radix precision); unfold CompatibleP.
intros T; apply T with (p+q)%R v; auto with real zarith.
rewrite FnormalizeCorrect; auto with zarith real.
apply FnormalizeBounded; auto with zarith;apply H with (p+q)%R; auto.
repeat rewrite FnormalizeCorrect; auto with zarith real.
fold FtoRradix; intros.
generalize (EvenClosestCompatible bo radix precision); unfold CompatibleP.
intros T; apply T with (p-q)%R d; auto with real zarith.
rewrite FnormalizeCorrect; auto with zarith real.
apply FnormalizeBounded; auto with zarith.
rewrite FnormalizeCorrect; auto with zarith real.
rewrite FnormalizeCorrect; auto with zarith real.
repeat rewrite FnormalizeCorrect; auto with zarith real.
fold FtoRradix; intros.
generalize (EvenClosestCompatible bo radix precision); unfold CompatibleP.
intros T; apply T with (dp-dq)%R s; auto with real zarith.
rewrite FnormalizeCorrect; auto with zarith real.
apply FnormalizeBounded; auto with zarith;apply H with (dp-dq)%R; auto.
repeat rewrite FnormalizeCorrect; auto with zarith real.
fold FtoRradix; intros.
generalize (EvenClosestCompatible bo radix precision); unfold CompatibleP.
intros T; apply T with (t+s)%R d; auto with real zarith.
rewrite FnormalizeCorrect; auto with zarith real.
apply FnormalizeBounded; auto with zarith.
Qed.

End Discriminant6.

Section Discriminant7.
Variable bo : Fbound.
Variable precision : nat.

Let radix := 2%Z.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ TwoMoreThanOne).

Hypothesis precisionGreaterThanOne : 1 < precision.
Hypothesis pGivesBound : Zpos (vNum bo) = Zpower_nat radix precision.
Hypothesis precisionGreaterThanFour : 4 <= precision.

Theorem FexpGeUnderf: forall e:Z, forall f:float,
  (Fbounded bo f) ->
  ((powerRZ radix e) <= Rabs f)%R -> (e-precision+1 <= Fexp f)%Z.
intros.
assert (e < Fexp f+precision)%Z; auto with zarith.
apply Zlt_powerRZ with radix;auto with real zarith.
apply Rle_lt_trans with (1:=H0).
unfold FtoRradix; rewrite <- Fabs_correct; auto.
rewrite Zplus_comm; rewrite powerRZ_add; auto with real zarith.
unfold Fabs, FtoR; simpl.
apply Rmult_lt_compat_r; auto with real zarith.
apply powerRZ_lt; try apply Rlt_IZR; auto with zarith.
apply Rlt_le_trans with (Zpos (vNum bo)); auto with zarith real; try apply Rlt_IZR, H.
rewrite pGivesBound; rewrite Zpower_nat_Z_powerRZ; auto with real zarith.
Qed.

Theorem AddExpGeUnderf: forall f1:float ,forall f2:float, forall g:float, forall e:Z,
    Closest bo radix (f1+f2) g -> (Fbounded bo f1) -> (Fbounded bo f2)
      -> (powerRZ radix e <= Rabs f1)%R
      -> (powerRZ radix e <= Rabs f2)%R
      -> ((FtoRradix g=0)%R \/ (powerRZ radix (e-precision+1) <= Rabs g)%R).
intros.
case (Req_dec g 0); auto; intros.
right.
elim plusExactExp with bo radix precision f1 f2 g; auto with zarith.
intros s T; elim T; intros g' T'; elim T'; intros R1 T''; clear T T'.
elim T''; intros R2 T; elim T; intros R3 T'; elim T'; intros R4 TT; clear T T' T''.
elim TT; intros R5 T; elim T; intros R6 R7; clear T TT.
replace (FtoRradix g) with (FtoRradix g'); auto with real.
apply Rle_trans with (powerRZ radix (Fexp g')).
apply Rle_powerRZ; auto with real zarith.
apply Z.le_trans with (Z.min (Fexp f1) (Fexp f2)); auto with zarith.
apply Zmin_Zle; apply FexpGeUnderf; auto.
unfold FtoRradix; rewrite <- Fabs_correct; auto.
unfold Fabs, FtoR; simpl.
apply Rle_trans with (1%Z*powerRZ radix (Fexp g'))%R;
  [simpl; right; ring |apply Rmult_le_compat_r; auto with real zarith].
apply powerRZ_le; try apply Rlt_IZR; auto with zarith.
apply Rle_IZR.
case (Zle_lt_or_eq 0 (Z.abs (Fnum g'))); auto with zarith real; intros.
case (Z.eq_dec 0 (Fnum g')); intros.
absurd (FtoR radix g=0)%R; auto with real.
rewrite <- R3; unfold FtoR; simpl; rewrite <- e0; simpl; ring.
absurd (0 < Z.abs (Fnum g'))%Z; auto with zarith.
Qed.

Theorem AddExpGeUnderf2: forall f1:float ,forall f2:float, forall g:float, forall e:Z,
    Closest bo radix (f1+f2) g -> (Fbounded bo f1) -> (Fbounded bo f2)
      -> (powerRZ radix e <= Rabs f1)%R
      -> (powerRZ radix e <= Rabs f2)%R
      -> (FtoRradix g <>0)%R
      -> (powerRZ radix (e-precision+1) <= Rabs g)%R.
intros.
case (AddExpGeUnderf f1 f2 g e); auto.
intros; absurd (FtoRradix g=0); auto with real.
Qed.

Theorem AddExpGe1Underf: forall f1:float ,forall f2:float, forall g:float, forall e:Z,
    Closest bo radix (f1+f2) g -> (Fcanonic radix bo f1) -> (Fcanonic radix bo f2)
      -> (powerRZ radix e <= Rabs f1)%R
      -> (-dExp bo <= e-1)%Z
      -> ((FtoRradix g=0)%R \/ (powerRZ radix (e-precision) <= Rabs g)%R).
intros.
assert (F1:(Fbounded bo f1));[apply FcanonicBound with radix; auto|idtac].
assert (F2:(Fbounded bo f2));[apply FcanonicBound with radix; auto|idtac].
case (Req_dec g 0); auto; intros.
right.
case (Rle_or_lt (Rabs f1) (Rabs f2)); intros.
apply Rle_trans with (powerRZ radix (e-precision+1)).
apply Rle_powerRZ; auto with real zarith.
apply AddExpGeUnderf2 with f1 f2; auto with real.
apply Rle_trans with (1:=H2); auto with real.
case (Rle_or_lt (Rabs f2) ((Rabs f1)/2)); intros.
apply Rle_trans with (powerRZ radix (e-1)).
apply Rle_powerRZ; auto with real zarith.
assert (exists f:float, Fbounded bo f /\ (FtoRradix f=powerRZ radix (e - 1))%R).
exists (Float 1 (e-1)); split.
split; simpl; auto with zarith.
apply vNumbMoreThanOne with radix precision; auto with zarith.
unfold FtoRradix, FtoR; simpl; ring.
elim H7; intros f T; elim T; intros; clear H7 T.
rewrite <- H9; unfold FtoRradix.
apply RoundAbsMonotonel with bo precision (Closest bo radix) (f1+f2)%R; auto.
apply ClosestRoundedModeP with precision; auto with zarith.
fold FtoRradix; rewrite H9.
unfold Zminus; rewrite powerRZ_add; auto with real zarith; simpl.
unfold radix; ring_simplify(2*1)%R; apply Rle_trans with ((Rabs f1)/2)%R.
unfold Rdiv; apply Rmult_le_compat_r; auto with real.
apply Rle_trans with (Rabs f1 -(Rabs f1)/2)%R;[right; field; auto with real|idtac].
apply Rle_trans with (Rabs f1-Rabs f2)%R.
unfold Rminus; apply Rplus_le_compat_l; auto with real.
rewrite <- (Rabs_Ropp f2).
replace (f1+f2)%R with (f1-(-f2))%R; try ring.
apply Rabs_triang_inv.
elim plusExactExp with bo radix precision f1 f2 g; auto with zarith.
intros s T; elim T; intros g' T'; elim T'; intros R1 T''; clear T T'.
elim T''; intros R2 T; elim T; intros R3 T'; elim T'; intros R4 TT; clear T T' T''.
elim TT; intros R5 T; elim T; intros R6 R7; clear T TT.
replace (FtoRradix g) with (FtoRradix g'); auto with real.
apply Rle_trans with (powerRZ radix (Fexp g')).
apply Rle_powerRZ; auto with real zarith.
apply Z.le_trans with (Z.min (Fexp f1) (Fexp f2)); auto with zarith.
rewrite Zmin_le2.
apply Z.le_trans with ((e-1)-precision+1)%Z; auto with zarith.
apply FexpGeUnderf; auto.
unfold Zminus; rewrite powerRZ_add; auto with real zarith; simpl.
unfold radix; ring_simplify (2*1)%R; apply Rle_trans with ((Rabs f1)/2)%R; auto with real.
unfold Rdiv; apply Rmult_le_compat_r; auto with real.
apply Fcanonic_Rle_Zle with radix bo precision; auto with zarith real.
unfold FtoRradix; rewrite <- Fabs_correct; auto.
unfold Fabs, FtoR; simpl.
apply Rle_trans with (1%Z*powerRZ radix (Fexp g'))%R;
  [simpl; right; ring|apply Rmult_le_compat_r; auto with real zarith].
apply powerRZ_le; try apply Rlt_IZR; auto with zarith.
apply Rle_IZR.
case (Zle_lt_or_eq 0 (Z.abs (Fnum g'))); auto with zarith real; intros.
case (Z.eq_dec 0 (Fnum g')); intros.
absurd (FtoR radix g=0)%R; auto with real.
rewrite <- R3; unfold FtoR; simpl; rewrite <- e0; simpl; ring.
absurd (0 < Z.abs (Fnum g'))%Z; auto with zarith.
Qed.

Theorem AddExpGe1Underf2: forall f1:float ,forall f2:float, forall g:float, forall e:Z,
    Closest bo radix (f1+f2) g -> (Fbounded bo f1) -> (Fbounded bo f2)
      -> (powerRZ radix e <= Rabs f1)%R
      -> (-dExp bo <= e-1)%Z
      -> (FtoRradix g <>0)%R
      -> (powerRZ radix (e-precision) <= Rabs g)%R.
intros.
case (AddExpGe1Underf (Fnormalize radix bo precision f1)
       (Fnormalize radix bo precision f2) g e); auto.
unfold FtoRradix; repeat rewrite FnormalizeCorrect; auto with real.
apply FnormalizeCanonic; auto with zarith.
apply FnormalizeCanonic; auto with zarith.
unfold FtoRradix; rewrite FnormalizeCorrect; auto with real.
intros; absurd (FtoRradix g=0); auto with real.
Qed.

Variables a b c p q t dp dq s d u v:float.

Let delta := (Rabs (d-(b*b-a*c)))%R.

Hypothesis Fa : (Fbounded bo a).
Hypothesis Fb : (Fbounded bo b).
Hypothesis Fc : (Fbounded bo c).
Hypothesis Fdp: (u < v)%R -> (Fbounded bo dp).
Hypothesis Fdq: (u < v)%R -> (Fbounded bo dq).

There is no underflow

Hypothesis U1: (FtoRradix b=0)%R \/
    (powerRZ radix (-dExp bo+3*precision-1) <= Rabs (b*b))%R.
Hypothesis U2: (a*c=0)%R \/
  (powerRZ radix (-dExp bo+3*precision-1) <= Rabs (a*c))%R.

Hypothesis Roundp : (EvenClosest bo radix precision (b*b)%R p).
Hypothesis Roundq : (EvenClosest bo radix precision (a*c)%R q).
Hypothesis Roundt : (EvenClosest bo radix precision (p-q)%R t).
Hypothesis Roundu : (EvenClosest bo radix precision (3*Rabs t)%R u).
Hypothesis Roundv : (EvenClosest bo radix precision (p+q)%R v).

Hypothesis FRoundd : (v <= u)%R ->
   (EvenClosest bo radix precision (p-q)%R d).

Hypothesis dpEq : (FtoRradix dp=b*b-p)%R.
Hypothesis dqEq : (FtoRradix dq=a*c-q)%R.
Hypothesis SRounds : (u < v)%R -> (EvenClosest bo radix precision (dp-dq)%R s).
Hypothesis SRoundd : (u < v)%R -> (EvenClosest bo radix precision (t+s)%R d).

Theorem pGeUnderf: (FtoRradix b <> 0)%R ->
   (powerRZ radix (-dExp bo+3*precision-1) <= Rabs (p))%R.
case U1; intros.
absurd (FtoRradix b=0); auto with real.
assert (exists f:float, Fbounded bo f /\
  (powerRZ radix (- dExp bo + 3 * precision - 1)= f)%R).
exists (Float 1 (- dExp bo + 3 * precision - 1)).
split;[split|unfold FtoRradix, FtoR; simpl; ring].
simpl; apply vNumbMoreThanOne with radix precision; auto with zarith.
apply Z.le_trans with (- dExp bo + 3 * precision - 1)%Z; auto with zarith.
elim H1; intros f T; elim T; intros; clear H1 T.
rewrite H3; unfold FtoRradix.
apply RoundAbsMonotonel with bo precision (EvenClosest bo radix precision)
    (b*b)%R; auto with real zarith.
apply EvenClosestRoundedModeP; auto with zarith.
fold FtoRradix; rewrite <- H3; auto.
Qed.

Theorem qGeUnderf: (a*c <> 0)%R ->
(powerRZ radix (-dExp bo+3*precision-1) <= Rabs (q))%R.
case U2; intros.
absurd (a*c=0)%R; auto with real.
assert (exists f:float, Fbounded bo f /\
  (powerRZ radix (- dExp bo + 3 * precision - 1)= f)%R).
exists (Float 1 (- dExp bo + 3 * precision - 1)).
split;[split|unfold FtoRradix, FtoR; simpl; ring].
simpl; apply vNumbMoreThanOne with radix precision; auto with zarith.
apply Z.le_trans with (- dExp bo + 3 * precision - 1)%Z; auto with zarith.
elim H1; intros f T; elim T; intros; clear H1 T.
rewrite H3; unfold FtoRradix.
apply RoundAbsMonotonel with bo precision (EvenClosest bo radix precision)
    (a*c)%R; auto with real zarith.
apply EvenClosestRoundedModeP; auto with zarith.
fold FtoRradix; rewrite <- H3; auto.
Qed.

Theorem cases: (FtoRradix b=0)%R \/ (a*c=0)%R
   \/ (FtoRradix d=0)%R \/ (FtoRradix v=0)%R \/ (FtoRradix t=0)%R \/
    ((powerRZ radix (-dExp bo+2*precision-1) <= Rabs (b*b))%R
     /\ (powerRZ radix (-dExp bo+2*precision-1) <= Rabs (a*c))%R
     /\ (powerRZ radix (-dExp bo+precision) <= Rabs d)%R
     /\ (powerRZ radix (-dExp bo+precision-1) <= Rabs u)%R
     /\ (powerRZ radix (-dExp bo+precision-1) <= Rabs v)%R
     /\ (powerRZ radix (-dExp bo+precision) <= Rabs t)%R).
case U1; auto; intros.
case U2; auto; intros.
case (Req_dec b 0); auto with real; intros; right.
case (Req_dec (a*c) 0); auto with real; intros; right.
assert (True); auto.
case (Req_dec d 0); auto with real; intros; right.
case (Req_dec v 0); auto with real; intros; right.
case (Req_dec t 0); auto with real; intros; right.
assert (powerRZ radix (-dExp bo+3*precision-1) <= Rabs (p))%R.
apply pGeUnderf; auto with real.
assert (powerRZ radix (-dExp bo+3*precision-1) <= Rabs (q))%R.
apply qGeUnderf; auto with real.
assert (powerRZ radix (-dExp bo+2*precision) <= Rabs v)%R.
replace (-dExp bo+2*precision)%Z with
  ((-dExp bo+3*precision-1)-precision+1)%Z; auto with zarith.
apply AddExpGeUnderf2 with p q; auto.
elim Roundv; auto.
elim Roundp; intros A1 A2; elim A1; auto.
elim Roundq; intros A1 A2; elim A1; auto.
assert (powerRZ radix (-dExp bo+2*precision) <= Rabs t)%R.
replace (-dExp bo+2*precision)%Z with
  ((-dExp bo+3*precision-1)-precision+1)%Z; auto with zarith.
apply AddExpGeUnderf2 with p (Fopp q); auto.
unfold FtoRradix; rewrite Fopp_correct; elim Roundt; auto with real.
elim Roundp; intros A1 A2; elim A1; auto.
apply oppBounded; elim Roundq; intros A1 A2; elim A1; auto.
unfold FtoRradix; rewrite Fopp_correct;rewrite Rabs_Ropp; auto with real.
assert (powerRZ radix (-dExp bo+precision+1) <= Rabs u)%R.
replace (-dExp bo+precision+1)%Z with
  ((-dExp bo+2*precision)-precision+1)%Z; auto with zarith.
apply AddExpGeUnderf2 with (Fabs t) (Fabs (Float (Fnum t) (Fexp t+1))); auto.
replace (Fabs t + Fabs (Float (Fnum t) (Fexp t + 1)))%R with (3*Rabs t)%R;
  [elim Roundu; auto|idtac].
unfold FtoRradix; rewrite <- Fabs_correct; auto; unfold FtoR, Fabs; simpl.
rewrite powerRZ_add; auto with real zarith.
rewrite powerRZ_1; unfold radix; ring.
apply absFBounded; elim Roundt; intros A1 A2; elim A1; auto.
elim Roundt; intros A1 A2; elim A1; intros A4 A3; elim A4; intros.
split; unfold Fabs; simpl; auto with zarith.
unfold FtoRradix; rewrite Fabs_correct; auto.
rewrite Rabs_right; auto with real.
apply Rle_ge, Rabs_pos.
apply Rle_trans with (1:=H10).
apply Rle_trans with (1*Rabs t)%R; auto with real.
apply Rle_trans with (2*Rabs t)%R; auto with real.
apply Rmult_le_compat_r; auto with real.
apply Rabs_pos.
unfold FtoRradix; rewrite Fabs_correct; auto.
rewrite Rabs_right with (Rabs (FtoR radix (Float (Fnum t) (Fexp t + 1))));
   auto with real.
repeat rewrite <- Fabs_correct; auto; unfold FtoR, Fabs; simpl.
rewrite powerRZ_add; auto with real zarith.
rewrite powerRZ_1; unfold radix; right; ring.
apply Rle_ge, Rabs_pos.
assert (0 < u)%R; auto with real.
apply Rlt_le_trans with (Rabs t); auto with real.
assert (Rabs t <> 0)%R; auto with real.
apply Rabs_no_R0; auto with real.
assert (A:(0 <= Rabs t)%R); try apply Rabs_pos; case A; auto with real.
intros; absurd (Rabs t=0)%R; auto with real.
unfold FtoRradix; rewrite <- Fabs_correct; auto.
apply RleBoundRoundl with bo precision (EvenClosest bo radix precision)
   (3 * Rabs t)%R; auto with zarith.
apply EvenClosestRoundedModeP; auto with zarith.
apply absFBounded; elim Roundt; intros A1 A2; elim A1; auto.
apply Rle_trans with (1*Rabs t)%R; auto with real.
rewrite Fabs_correct; auto; right; unfold FtoRradix; ring.
apply Rmult_le_compat_r; auto with real; apply Rabs_pos.
split.
apply Rle_trans with (2:=H); auto with real zarith.
apply Rle_powerRZ; try apply Rle_IZR; unfold radix; auto with zarith.
split.
apply Rle_trans with (2:=H0); auto with real zarith.
apply Rle_powerRZ; try apply Rle_IZR; unfold radix; auto with zarith.
split.
2: split.
2: apply Rle_trans with (2:=H11); auto with real zarith.
2: apply Rle_powerRZ; try apply Rle_IZR; unfold radix; auto with zarith.
2:split.
2:apply Rle_trans with (2:=H9); auto with real zarith.
2: apply Rle_powerRZ; try apply Rle_IZR; unfold radix; auto with zarith.
2:apply Rle_trans with (2:=H10); auto with real zarith.
2: apply Rle_powerRZ; try apply Rle_IZR; unfold radix; auto with zarith.
case (Rle_or_lt v u); intros.
assert (powerRZ radix (-dExp bo+2*precision) <= Rabs d)%R.
replace (-dExp bo+2*precision)%Z with
  ((-dExp bo+3*precision-1)-precision+1)%Z; auto with zarith.
apply AddExpGeUnderf2 with p (Fopp q); auto.
unfold FtoRradix; rewrite Fopp_correct; elim FRoundd; auto with real.
elim Roundp; intros A1 A2; elim A1; auto.
apply oppBounded; elim Roundq; intros A1 A2; elim A1; auto.
unfold FtoRradix; rewrite Fopp_correct;rewrite Rabs_Ropp; auto with real.
apply Rle_trans with (2:=H13); auto with real zarith.
apply Rle_powerRZ; try apply Rle_IZR; unfold radix; auto with zarith.
replace (-dExp bo+precision)%Z with
  ((-dExp bo+2*precision)-precision)%Z; auto with zarith.
apply AddExpGe1Underf2 with t s; auto with zarith.
elim SRoundd; auto with real.
elim Roundt; intros A1 A2; elim A1; auto.
elim SRounds; auto; intros A1 A2; elim A1; auto.
Qed.

Theorem discri16: (FtoRradix d=0)%R \/ (delta <= 2*(Fulp bo radix precision d))%R.
assert (Fq:(Fbounded bo q)).
elim Roundq; intros A1 A2; elim A1; auto with zarith.
assert (Fp:(Fbounded bo p)).
elim Roundp; intros A1 A2; elim A1; auto with zarith.
assert (Fd:(Fbounded bo d)).
case (Rle_or_lt v u); intros.
elim FRoundd; auto; intros L; elim L;
   intros A1 A2; elim A1; auto with zarith.
elim SRoundd; auto; intros L; elim L;
   intros A1 A2; elim A1; auto with zarith.
generalize cases; intros C.
case C; clear C; intros C.
assert (FtoRradix p=0)%R; auto with real.
assert (0 <= p)%R.
unfold FtoRradix; apply RleRoundedR0 with bo precision
    (EvenClosest bo radix precision) (b*b)%R; auto with real zarith.
apply EvenClosestRoundedModeP; auto with zarith.
rewrite C; right; ring.
assert (p <= 0)%R; auto with real.
unfold FtoRradix; apply RleRoundedLessR0 with bo precision
    (EvenClosest bo radix precision) (b*b)%R; auto with real zarith.
apply EvenClosestRoundedModeP; auto with zarith.
rewrite C; right; ring.
assert (FtoRradix d=(Fopp q))%R.
apply sym_eq; unfold FtoRradix; apply RoundedModeProjectorIdemEq
  with bo precision (EvenClosest bo radix precision); auto with zarith.
apply EvenClosestRoundedModeP; auto with zarith.
apply oppBounded; auto.
replace (FtoR radix (Fopp q)) with (p-q)%R; [apply FRoundd|
   rewrite Fopp_correct; fold FtoRradix; rewrite H; ring].
unfold FtoRradix; apply EvenClosestMonotone2 with bo precision
   (p+q)%R (3*Rabs t)%R; auto.
replace (FtoRradix t) with (FtoRradix (Fopp q)).
unfold FtoRradix; rewrite Fopp_correct;auto; fold FtoRradix.
rewrite H; ring_simplify (0+q)%R.
rewrite Rabs_Ropp; apply Rle_trans with (1:=(RRle_abs q)).
apply Rle_trans with (1*(Rabs q))%R; auto with real.
apply Rmult_le_compat_r; auto with real; apply Rabs_pos.
unfold FtoRradix; apply RoundedModeProjectorIdemEq
  with bo precision (EvenClosest bo radix precision); auto with zarith.
apply EvenClosestRoundedModeP; auto with zarith.
apply oppBounded; auto.
replace (FtoR radix (Fopp q)) with (p-q)%R; [auto|
   rewrite Fopp_correct; fold FtoRradix; rewrite H; ring].
right; unfold delta; rewrite H0.
unfold FtoRradix; rewrite Fopp_correct; fold FtoRradix.
replace (-q-(b*b-a*c))%R with ((a*c-q))%R;[idtac|rewrite C; ring].
apply Rle_trans with (/2*(Fulp bo radix precision q))%R.
apply Rmult_le_reg_l with (2%nat)%R; auto with real zarith.
apply Rle_trans with (Fulp bo radix precision q);[idtac|simpl; right; field; auto with real].
unfold FtoRradix; apply ClosestUlp; auto.
elim Roundq; auto.
apply Rle_trans with (2 * Fulp bo radix precision q)%R;
  [apply Rmult_le_compat_r; auto with real|idtac].
unfold Fulp; auto with real zarith.
apply powerRZ_le; try apply Rlt_IZR; auto with zarith.
lra.
apply Rmult_le_compat_l; auto with real; right.
apply trans_eq with (Fulp bo radix precision (Fopp q)).
unfold Fulp; rewrite Fnormalize_Fopp; auto with real zarith.
apply FulpComp; auto with zarith.
apply oppBounded; auto.
case C; clear C; intros C.
assert (FtoRradix q=0)%R; auto with real.
assert (0 <= q)%R.
unfold FtoRradix; apply RleRoundedR0 with bo precision
    (EvenClosest bo radix precision) (a*c)%R; auto with real zarith.
apply EvenClosestRoundedModeP; auto with zarith.
assert (q <= 0)%R; auto with real.
unfold FtoRradix; apply RleRoundedLessR0 with bo precision
    (EvenClosest bo radix precision) (a*c)%R; auto with real zarith.
apply EvenClosestRoundedModeP; auto with zarith.
assert (FtoRradix d=p)%R.
apply sym_eq; unfold FtoRradix; apply RoundedModeProjectorIdemEq
  with bo precision (EvenClosest bo radix precision); auto with zarith.
apply EvenClosestRoundedModeP; auto with zarith.
replace (FtoR radix p) with (p-q)%R; [apply FRoundd|
  fold FtoRradix; rewrite H; ring].
unfold FtoRradix; apply EvenClosestMonotone2 with bo precision
   (p+q)%R (3*Rabs t)%R; auto.
replace (FtoRradix t) with (FtoRradix p).
rewrite H; ring_simplify (p+0)%R; apply Rle_trans with (1:=(RRle_abs p)).
apply Rle_trans with (1*(Rabs p))%R; auto with real.
apply Rmult_le_compat_r; auto with real; apply Rabs_pos.
unfold FtoRradix; apply RoundedModeProjectorIdemEq
  with bo precision (EvenClosest bo radix precision); auto with zarith.
apply EvenClosestRoundedModeP; auto with zarith.
replace (FtoR radix p) with (p-q)%R; [auto|
   fold FtoRradix; rewrite H; ring].
right; unfold delta; rewrite H0.
replace (p-(b*b-a*c))%R with ((-(b*b-p)))%R;[idtac|rewrite C; ring].
apply Rle_trans with (/2*(Fulp bo radix precision p))%R.
apply Rmult_le_reg_l with (2%nat)%R; auto with real zarith.
apply Rle_trans with (Fulp bo radix precision p);[idtac|simpl; right; field; auto with real].
rewrite Rabs_Ropp; unfold FtoRradix; apply ClosestUlp; auto.
elim Roundp; auto.
rewrite FulpComp with bo radix precision d p; auto with zarith.
apply Rmult_le_compat_r; auto with real.
unfold Fulp; auto with real zarith.
apply powerRZ_le, Rlt_IZR; auto with zarith.
lra.
case (Req_dec v 0)%R.
clear C; intros C.
assert (v <= u)%R.
rewrite C; unfold FtoRradix; apply RleRoundedR0 with bo precision
    (EvenClosest bo radix precision) (3*Rabs t)%R; auto with real zarith.
apply EvenClosestRoundedModeP; auto with zarith.
apply Rle_trans with (0*0)%R; auto with real; apply Rmult_le_compat; auto with real.
apply Rabs_pos.
assert (FtoRradix q=-p)%R.
assert (FtoRradix v=p+q)%R.
unfold FtoRradix; apply plusExactR0 with bo precision; auto with real.
elim Roundv; auto.
apply Rplus_eq_reg_l with p.
rewrite <- H0; rewrite C; ring.
assert (FtoRradix d=2*p)%R.
assert (exists f:float, Fbounded bo f /\ (FtoRradix f=2*p)%R).
exists (Float (Fnum p) (Fexp p+1)); split.
elim Fp; intros; split; simpl; auto with zarith.
unfold FtoRradix, FtoR; simpl; rewrite powerRZ_add; auto with real zarith.
rewrite powerRZ_1; unfold radix; ring.
elim H1; intros f T; elim T; intros; clear T H1.
rewrite <- H3; unfold FtoRradix; apply sym_eq.
apply RoundedModeProjectorIdemEq
  with bo precision (EvenClosest bo radix precision); auto with zarith.
apply EvenClosestRoundedModeP; auto with zarith.
replace (FtoR radix f) with (p-q)%R; auto.
fold FtoRradix; rewrite H3; rewrite H0; ring.
right; unfold delta.
replace (d - (b * b - a * c))%R with (-(b*b-p)+(a*c-q))%R;
  [idtac|rewrite H1; rewrite H0; ring].
apply Rle_trans with (Rabs (- (b * b - p)) + Rabs ((a * c - q)))%R;
  [apply Rabs_triang|idtac].
apply Rle_trans with (/2*(Fulp bo radix precision p)
  +/2*(Fulp bo radix precision q))%R.
apply Rplus_le_compat; apply Rmult_le_reg_l with (2%nat)%R; auto with real zarith.
apply Rle_trans with (Fulp bo radix precision p);
  [idtac|simpl; right; field; auto with real].
rewrite Rabs_Ropp; unfold FtoRradix; apply ClosestUlp; auto.
elim Roundp; auto.
apply Rle_trans with (Fulp bo radix precision q);
  [idtac|simpl; right; field; auto with real].
unfold FtoRradix; apply ClosestUlp; auto.
elim Roundq; auto.
replace (Fulp bo radix precision q) with (Fulp bo radix precision p).
apply Rle_trans with (Fulp bo radix precision p);
 [right; field; auto with real|idtac].
apply Rle_trans with (1*(Fulp bo radix precision d))%R;
 [idtac|apply Rmult_le_compat_r; unfold Fulp; auto with real zarith].
apply Rle_trans with (Fulp bo radix precision d); auto with real.
assert (0 <= p)%R.
apply P_positive with bo precision b b; auto.
fold FtoRradix; apply Rle_trans with (Rsqr b); auto with real.
apply LeFulpPos; auto with real zarith.
fold FtoRradix; rewrite H1; apply Rle_trans with (1*p)%R; auto with real.
apply powerRZ_le, Rlt_IZR; auto with zarith.
apply trans_eq with (Fulp bo radix precision (Fopp p)).
unfold Fulp; rewrite Fnormalize_Fopp; auto with real zarith.
apply FulpComp; auto with zarith.
apply oppBounded; auto.
rewrite Fopp_correct; auto with real zarith.
intros C'.
case (Req_dec t 0).
clear C; intros C.
assert (FtoRradix p=q)%R.
assert (FtoRradix t=p+ Fopp q)%R.
unfold FtoRradix; apply plusExactR0 with bo precision; auto with real zarith.
apply oppBounded; auto.
elim Roundt; intros; rewrite Fopp_correct; auto with real.
apply Rplus_eq_reg_l with (Fopp q); rewrite Rplus_comm.
rewrite <- H; rewrite C; unfold FtoRradix; rewrite Fopp_correct; auto; ring.
assert (FtoRradix v=2*p)%R.
assert (exists f:float, Fbounded bo f /\ (FtoRradix f=2*p)%R).
exists (Float (Fnum p) (Fexp p+1)); split.
elim Fp; intros; split; simpl; auto with zarith.
unfold FtoRradix, FtoR; simpl; rewrite powerRZ_add; auto with real zarith.
rewrite powerRZ_1; unfold radix ; ring.
elim H0; intros f T; elim T; intros; clear T H0.
rewrite <- H2; unfold FtoRradix; apply sym_eq.
apply RoundedModeProjectorIdemEq
  with bo precision (EvenClosest bo radix precision); auto with zarith.
apply EvenClosestRoundedModeP; auto with zarith.
replace (FtoR radix f) with (p+q)%R; auto.
fold FtoRradix; rewrite H2; rewrite H; ring.
assert (FtoRradix u=0)%R.
assert (0 <= u)%R.
unfold FtoRradix; apply RleRoundedR0 with bo precision
    (EvenClosest bo radix precision) (3*Rabs t)%R; auto with real zarith.
apply EvenClosestRoundedModeP; auto with zarith.
rewrite C; rewrite Rabs_R0; right; ring.
assert (u <= 0)%R; auto with real.
unfold FtoRradix; apply RleRoundedLessR0 with bo precision
    (EvenClosest bo radix precision) (3*Rabs t)%R; auto with real zarith.
apply EvenClosestRoundedModeP; auto with zarith.
rewrite C; rewrite Rabs_R0; right; ring.
assert (u < v)%R.
rewrite H1; assert (0 <= v)%R; auto with real.
rewrite H0; apply Rle_trans with (0*0)%R; auto with real.
apply Rmult_le_compat; auto with real.
apply P_positive with bo precision b b; auto.
apply Rle_trans with (Rsqr (FtoR radix b)); auto with real.
case H2; auto with real.
intros T; absurd (FtoRradix v=0)%R; auto with real.
assert (FtoRradix d=s)%R.
unfold FtoRradix; apply sym_eq.
apply RoundedModeProjectorIdemEq
  with bo precision (EvenClosest bo radix precision); auto with zarith.
apply EvenClosestRoundedModeP; auto with zarith.
elim SRounds; auto; intros A1 A2; elim A1; auto.
replace (FtoR radix s) with (t+s)%R; auto.
fold FtoRradix; rewrite C; ring.
right; unfold delta.
replace (d - (b * b - a * c))%R with (-((dp-dq)-s))%R;
  [idtac|rewrite dpEq; auto; rewrite dqEq; auto; rewrite H3; rewrite H; ring].
apply Rle_trans with (/2*(Fulp bo radix precision s))%R.
apply Rmult_le_reg_l with (2%nat)%R; auto with real zarith.
apply Rle_trans with (Fulp bo radix precision s);
  [idtac|simpl; right; field; auto with real].
rewrite Rabs_Ropp; unfold FtoRradix; apply ClosestUlp; auto.
elim SRounds; auto.
rewrite FulpComp with bo radix precision d s; auto with zarith real.
2: elim SRounds; auto; intros A1 A2; elim A1; auto.
apply Rmult_le_compat_r.
unfold Fulp; auto with real zarith.
apply powerRZ_le, Rlt_IZR; auto with zarith.
lra.
intros C''.
case (Rle_or_lt v u); intros.
case C; clear C; intros C; auto with real.
case C; clear C; intros C; auto with real.
absurd (FtoRradix v=0)%R; auto with real.
case C; clear C; intros C; auto with real.
absurd (FtoRradix t=0)%R; auto with real.
elim C; intros Y1 C1; elim C1; intros Y2 C2; elim C2; intros Y3 C3; clear C C1 C2.
elim C3; intros Y4 C4; elim C4; intros Y5 Y6; clear C3 C4.
right; unfold delta;apply discri15 with p q t dp dq s u v; auto with real zarith.
case C; clear C; intros C; auto with real.
case C; clear C; intros C; auto with real.
absurd (FtoRradix v=0)%R; auto with real.
case C; clear C; intros C; auto with real.
absurd (FtoRradix t=0)%R; auto with real.
elim C; intros Y1 C1; elim C1; intros Y2 C2; elim C2; intros Y3 C3; clear C C1 C2.
elim C3; intros Y4 C4; elim C4; intros Y5 Y6; clear C3 C4.
right; unfold delta;apply discri15 with p q t dp dq s u v; auto with real zarith.
Qed.
End Discriminant7.

This proof file has been written by Sylvie Boldo(1), following a proof presented by Pr William Kahan (2), and adapted to Coq proof checker with the help of Guillaume Melquiond(1) and Marc Daumas(1). This work has been partially supported by the CNRS grant PICS 2533.
(1) LIP Computer science laboratory UMR 5668 CNRS - ENS de Lyon - INRIA Lyon, France
(2) University of California at Berkeley Berkeley, California

Section Discriminant5B.
Variable bo : Fbound.
Variable precision : nat.

Let radix := 2%Z.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ TwoMoreThanOne).

Hypothesis precisionGreaterThanOne : 1 < precision.
Hypothesis pGivesBound : Zpos (vNum bo) = Zpower_nat radix precision.

Variables a b c p q t dp dq s d:float.

Let delta := (Rabs (d-(b*b-a*c)))%R.

Hypothesis Fa : (Fbounded bo a).
Hypothesis Fb : (Fbounded bo b).
Hypothesis Fc : (Fbounded bo c).
Hypothesis Fp : (Fbounded bo p).
Hypothesis Fq : (Fbounded bo q).
Hypothesis Fd : (Fbounded bo d).
Hypothesis Ft : (Fbounded bo t).
Hypothesis Fs : (Fbounded bo s).
Hypothesis Fdp: (Fbounded bo dp).
Hypothesis Fdq: (Fbounded bo dq).

There is no underflow
Hypothesis U0: (- dExp bo <= Fexp d - 1)%Z.
Hypothesis U1: (- dExp bo <= (Fexp t)-1)%Z.
Hypothesis U2: (powerRZ radix (-dExp bo+2*precision-1) <= Rabs (b*b))%R.
Hypothesis U3: (powerRZ radix (-dExp bo+2*precision-1) <= Rabs (a*c))%R.

Hypothesis Np:(FtoRradix p=0)%R \/ (Fnormal radix bo p).
Hypothesis Nq:(FtoRradix q=0)%R \/ (Fnormal radix bo q).
Hypothesis Ns: (3*(Rabs (p-q)) < p+q)%R -> (FtoRradix s=0)%R \/ (Fnormal radix bo s).
Hypothesis Nd: (Fnormal radix bo d).

Hypothesis Roundp : (EvenClosest bo radix precision (b*b)%R p).
Hypothesis Roundq : (EvenClosest bo radix precision (a*c)%R q).

Hypothesis Firstcase : (p+q <= 3*(Rabs (p-q)))%R ->
   (EvenClosest bo radix precision (p-q)%R d).

Hypothesis SRoundt : (3*(Rabs (p-q)) < p+q)%R -> (EvenClosest bo radix precision (p-q)%R t).
Hypothesis SdpEq : (3*(Rabs (p-q)) < p+q)%R -> (FtoRradix dp=b*b-p)%R.
Hypothesis SdqEq : (3*(Rabs (p-q)) < p+q)%R -> (FtoRradix dq=a*c-q)%R.
Hypothesis SRounds : (3*(Rabs (p-q)) < p+q)%R -> (EvenClosest bo radix precision (dp-dq)%R s).
Hypothesis SRoundd : (3*(Rabs (p-q)) < p+q)%R -> (EvenClosest bo radix precision (t+s)%R d).

Theorem discri: (delta <= 2*(Fulp bo radix precision d))%R.
case Np; intros.
assert (FtoRradix d=(Fopp q))%R.
apply sym_eq; unfold FtoRradix; apply RoundedModeProjectorIdemEq with bo precision (EvenClosest bo radix precision); auto with zarith.
apply EvenClosestRoundedModeP; auto.
apply oppBounded; auto.
replace (FtoR radix (Fopp q)) with (p-q)%R; [apply Firstcase|rewrite Fopp_correct; fold FtoRradix; rewrite H; ring].
rewrite H; ring_simplify (0-q)%R; ring_simplify (0+q)%R.
rewrite Rabs_Ropp; apply Rle_trans with (1:=(RRle_abs q)).
apply Rle_trans with (1*(Rabs q))%R; auto with real.
apply Rmult_le_compat_r; auto with real; apply Rabs_pos.
unfold delta; rewrite H0.
unfold FtoRradix; rewrite Fopp_correct; fold FtoRradix.
replace (-q-(b*b-a*c))%R with ((a*c-q)+(-(b*b)))%R;[idtac|ring].
apply Rle_trans with (1:=(Rabs_triang (a * c - q)%R (-(b*b))%R)).
apply Rle_trans with (/2*(Fulp bo radix precision q)+/2*(Fulp bo radix precision p))%R.
apply Rplus_le_compat; apply Rmult_le_reg_l with (2%nat)%R; auto with real zarith.
apply Rle_trans with (Fulp bo radix precision q);[idtac|simpl; right; field; auto with real].
unfold FtoRradix; apply ClosestUlp; auto.
elim Roundq; auto.
apply Rle_trans with (Fulp bo radix precision p);[idtac|simpl; right; field; auto with real].
replace (-(b*b))%R with (-(b*b-p))%R;[rewrite Rabs_Ropp|rewrite H;ring].
unfold FtoRradix; apply ClosestUlp; auto.
elim Roundp; auto.
apply Rle_trans with (/ 2 * Fulp bo radix precision q + / 2 * Fulp bo radix precision q)%R;[apply Rplus_le_compat; auto with real|idtac].
apply Rmult_le_compat_l; auto with real.
rewrite Fulp_zero.
unfold Fulp; apply Rle_powerRZ;auto with real zarith.
apply FnormalizeBounded; auto with zarith.
apply is_Fzero_rep2 with radix; auto with zarith.
apply Rle_trans with (1 * Fulp bo radix precision q)%R;[right; field; auto with real|idtac].
apply Rmult_le_compat; auto with real zarith.
unfold Fulp; auto with real zarith.
apply powerRZ_le, Rlt_IZR; auto with zarith.
right; apply trans_eq with (Fulp bo radix precision (Fopp q)).
unfold Fulp; rewrite Fnormalize_Fopp; auto with real zarith.
apply FulpComp; auto with zarith.
apply oppBounded; auto.
case Nq; intros.
assert (FtoRradix d=p)%R.
apply sym_eq; unfold FtoRradix; apply RoundedModeProjectorIdemEq with bo precision (EvenClosest bo radix precision); auto with zarith.
apply EvenClosestRoundedModeP; auto.
replace (FtoR radix p) with (p-q)%R; [apply Firstcase|fold FtoRradix; rewrite H0; ring].
rewrite H0; ring_simplify (p+0)%R; ring_simplify (p-0)%R.
apply Rle_trans with (1:=(RRle_abs p)).
apply Rle_trans with (1*(Rabs p))%R; auto with real.
apply Rmult_le_compat_r; auto with real; apply Rabs_pos.
unfold delta; rewrite H1.
replace (p-(b*b-a*c))%R with (a*c+(-(b*b-p)))%R;[idtac|ring].
apply Rle_trans with (1:=(Rabs_triang (a * c )%R (-(b*b-p))%R)).
apply Rle_trans with (/2*(Fulp bo radix precision q)+/2*(Fulp bo radix precision p))%R.
apply Rplus_le_compat; apply Rmult_le_reg_l with (2%nat)%R; auto with real zarith.
apply Rle_trans with (Fulp bo radix precision q);[idtac|simpl; right; field; auto with real].
replace (a*c)%R with (a*c-q)%R;[idtac|rewrite H0;ring].
unfold FtoRradix; apply ClosestUlp; auto.
elim Roundq; auto.
apply Rle_trans with (Fulp bo radix precision p);[idtac|simpl; right; field; auto with real].
rewrite Rabs_Ropp; unfold FtoRradix; apply ClosestUlp; auto.
elim Roundp; auto.
apply Rle_trans with (/ 2 * Fulp bo radix precision p + / 2 * Fulp bo radix precision p)%R;[apply Rplus_le_compat; auto with real|idtac].
apply Rmult_le_compat_l; auto with real.
rewrite Fulp_zero.
unfold Fulp; apply Rle_powerRZ;auto with real zarith.
apply FnormalizeBounded; auto with zarith.
apply is_Fzero_rep2 with radix; auto with zarith.
apply Rle_trans with (1 * Fulp bo radix precision p)%R;[right; field; auto with real|idtac].
apply Rmult_le_compat; auto with real zarith.
unfold Fulp; auto with real zarith.
apply powerRZ_le, Rlt_IZR; auto with zarith.
right; apply FulpComp; auto with zarith.
case (Rle_or_lt (p + q)%R (3 * Rabs (p - q))%R); intros.
unfold delta;apply discri1 with p q; auto.
apply Rle_trans with (Rsqr (FtoR 2 b)) ; auto with real.
case (Ns H1); intros.
unfold delta; apply discri3 with p q t dp dq s; auto with real zarith.
apply Rle_trans with (Rsqr (FtoR 2 b)) ; auto with real.
2:unfold delta; apply discri9 with p q t dp dq s; auto.
exists (Float 0%Z 0%Z).
split.
unfold Fbounded; split; auto with zarith.
simpl; case (dExp bo); auto with zarith.
apply trans_eq with 0%R.
unfold FtoR; simpl; ring.
fold radix; fold FtoRradix.
rewrite <- H2.
unfold FtoRradix; rewrite plusExactR0 with bo radix precision dp (Fopp dq) s; auto.
rewrite Fopp_correct; ring.
apply oppBounded; auto.
replace (FtoR radix dp + FtoR radix (Fopp dq))%R with (dp - dq)%R; auto.
elim SRounds; auto.
unfold FtoRradix; rewrite Fopp_correct; auto with real.
intros; left; auto.
Qed.
End Discriminant5B.


Section Fast.
Variable b : Fbound.
Variable precision : nat.

Let radix := 2%Z.

Local Coercion FtoRradix := FtoR radix.
Hypothesis precisionGreaterThanOne : 1 < precision.
Hypothesis pGivesBound : Zpos (vNum b) = Zpower_nat radix precision.
Variable Iplus : float -> float -> float.
Hypothesis
  IplusCorrect :
    forall p q : float,
    Fbounded b p -> Fbounded b q -> Closest b radix (p + q) (Iplus p q).
Hypothesis IplusSym : forall p q : float, Iplus p q = Iplus q p.
Hypothesis
  IplusOp : forall p q : float, Fopp (Iplus p q) = Iplus (Fopp p) (Fopp q).
Variable Iminus : float -> float -> float.
Hypothesis IminusPlus : forall p q : float, Iminus p q = Iplus p (Fopp q).

Let radixMoreThanOne : (1 < radix)%Z.
unfold radix in |- *; red in |- *; simpl in |- *; auto.
Qed.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ radixMoreThanOne).

Theorem IminusCorrect :
 forall p q : float,
 Fbounded b p -> Fbounded b q -> Closest b radix (p - q) (Iminus p q).
intros p q H' H'0.
rewrite IminusPlus.
unfold Rminus in |- *.
unfold FtoRradix in |- *; rewrite <- Fopp_correct.
apply IplusCorrect; auto.
apply oppBounded; auto.
Qed.

Theorem ErrorBoundedIplus :
 forall p q : float,
 Fbounded b p ->
 Fbounded b q ->
 exists error : float, error = (p + q - Iplus p q)%R :>R /\ Fbounded b error.
intros p q H' H'0.
case
 (errorBoundedPlus b radix precision)
  with (p := p) (q := q) (pq := Iplus p q); auto.
intros x H'1; elim H'1; intros H'2 H'3; elim H'3; intros H'4 H'5; exists x;
 auto.
Qed.

Theorem IplusOr :
 forall p q : float,
 Fbounded b p -> Fbounded b q -> q = 0%R :>R -> Iplus p q = p :>R.
intros p q H' H'0 H'1.
cut (Closest b radix (p + q) (Iplus p q)); auto.
rewrite H'1.
rewrite Rplus_0_r.
intros H'2; apply sym_eq.
unfold FtoRradix in |- *; apply (ClosestIdem b radix); auto.
Qed.

Theorem IminusId :
 forall p q : float,
 Fbounded b p -> Fbounded b q -> p = q :>R -> Iminus p q = 0%R :>R.
intros p q H' H'0 H'1.
cut (Closest b radix (p - q) (Iminus p q)); auto.
replace (p - q)%R with 0%R; [ idtac | rewrite <- H'1; ring ].
replace 0%R with (FtoRradix (Fzero (- dExp b))).
intros H'2; apply sym_eq.
unfold FtoRradix in |- *; apply (ClosestIdem b radix); auto.
apply FboundedFzero; auto.
unfold Fzero, FtoRradix, FtoR in |- *; simpl in |- *; ring.
apply IminusCorrect; easy.
Qed.

Theorem IplusBounded :
 forall p q : float, Fbounded b p -> Fbounded b q -> Fbounded b (Iplus p q).
intros p q H' H'0; cut (Closest b radix (p + q) (Iplus p q)); auto.
intros H1; case H1; auto.
Qed.

Theorem IminusBounded :
 forall p q : float, Fbounded b p -> Fbounded b q -> Fbounded b (Iminus p q).
intros p q H' H'0; cut (Closest b radix (p - q) (Iminus p q)); auto.
intros H1; case H1; auto.
apply IminusCorrect; easy.
Qed.

Theorem MDekkerAux1 :
 forall p q : float,
 Iminus (Iplus p q) p = (Iplus p q - p)%R :>R ->
 Fbounded b p ->
 Fbounded b q -> Iminus q (Iminus (Iplus p q) p) = (p + q - Iplus p q)%R :>R.
intros p q H' H'0 H'1.
elim (ErrorBoundedIplus p q);
 [ intros error E; elim E; intros H'7 H'8; clear E | idtac | idtac ];
 auto.
cut
 (Closest b radix (q - Iminus (Iplus p q) p)
    (Iminus q (Iminus (Iplus p q) p))); auto.
rewrite H'.
replace (q - (Iplus p q - p))%R with (p + q - Iplus p q)%R; [ idtac | ring ].
rewrite <- H'7.
intros H'2.
apply sym_eq; apply (ClosestIdem b radix); auto.
apply IminusCorrect; auto.
repeat (try apply IplusBounded; try apply IminusBounded; auto); auto.
Qed.

Theorem MDekkerAux2 :
 forall p q : float,
 Iplus p q = (p + q)%R :>R ->
 Fbounded b p -> Fbounded b q -> Iminus (Iplus p q) p = (Iplus p q - p)%R :>R.
intros p q H' H'0 H'1.
cut (Closest b radix (Iplus p q - p) (Iminus (Iplus p q) p)); auto.
repeat rewrite H'.
replace (p + q - p)%R with (FtoRradix q); [ idtac | ring ].
intros H'2.
apply sym_eq; apply (ClosestIdem b radix); auto.
apply IminusCorrect; auto.
apply IplusBounded; auto.
Qed.

Theorem MDekkerAux3 :
 forall p q : float,
 Fbounded b (Fplus radix p q) ->
 Fbounded b p -> Fbounded b q -> Iminus (Iplus p q) p = (Iplus p q - p)%R :>R.
intros p q H' H'0 H'1.
apply MDekkerAux2; auto.
unfold FtoRradix in |- *; rewrite <- Fplus_correct; auto.
apply sym_eq; apply (ClosestIdem b radix); auto.
rewrite Fplus_correct; auto.
Qed.

Theorem MDekkerAux4 :
 forall p q : float,
 Fbounded b (Fminus radix (Iplus p q) p) ->
 Fbounded b p -> Fbounded b q -> Iminus (Iplus p q) p = (Iplus p q - p)%R :>R.
intros p q H' H'0 H'1.
unfold FtoRradix in |- *; rewrite <- Fminus_correct; auto.
apply sym_eq; apply (ClosestIdem b radix); auto.
rewrite Fminus_correct; auto.
apply IminusCorrect; auto.
repeat (try apply IplusBounded; try apply IminusBounded; auto); auto.
Qed.

Theorem Dekker1_FTS :
 forall p q : float,
 (0 <= q)%R ->
 (q <= p)%R ->
 Fbounded b p -> Fbounded b q -> Iminus (Iplus p q) p = (Iplus p q - p)%R :>R.
intros p q H' H'0 H'1 H'2.
apply MDekkerAux4; auto.
apply oppBoundedInv; auto.
rewrite Fopp_Fminus; auto.
apply Sterbenz; auto.
apply IplusBounded; auto.
apply Rmult_le_reg_l with (r := INR 2); auto with real arith.
rewrite <- Rmult_assoc; rewrite Rinv_r; auto with real arith;
 rewrite Rmult_1_l.
apply (RoundedModeMult b radix) with (P := Closest b radix) (r := (p + q)%R);
 auto.
apply ClosestRoundedModeP with (precision := precision); auto.
replace (radix * FtoR radix p)%R with (p + p)%R;
 [ idtac | simpl in |- *; fold FtoRradix; unfold radix; ring ]; auto with real.
case H'; clear H'; intros H'.
apply Rmult_le_reg_l with (r := (/ radix)%R); auto with real.
rewrite <- Rmult_assoc; rewrite Rinv_l; auto with real; rewrite Rmult_1_l.
apply (FmultRadixInv b radix precision) with (y := (p + q)%R); auto.
apply Rmult_lt_reg_l with (r := INR 2); auto with real.
rewrite <- Rmult_assoc; rewrite Rinv_r; auto with real; rewrite Rmult_1_l.
apply Rle_lt_trans with (radix * p)%R.
apply Rledouble; auto.
apply Rlt_le; apply Rlt_le_trans with (FtoRradix q); auto.
apply Rmult_lt_reg_l with (r := (/ radix)%R); auto with real.
repeat rewrite <- Rmult_assoc; repeat rewrite Rinv_l; auto with real zarith;
 repeat rewrite Rmult_1_l.
pattern (FtoRradix p) at 1 in |- *; replace (FtoRradix p) with (p + 0)%R;
 [ idtac | ring ]; auto with real.
rewrite IplusOr; auto.
apply Rledouble; auto.
rewrite H'; auto.
Qed.

Theorem Dekker2_FTS :
 forall p q : float,
 (0 <= p)%R ->
 (- q <= p)%R ->
 (p <= radix * - q)%R ->
 Fbounded b p -> Fbounded b q -> Iminus (Iplus p q) p = (Iplus p q - p)%R :>R.
intros p q H' H'0 H'1 H'2 H'3.
apply MDekkerAux3; auto.
rewrite <- (Fopp_Fopp q).
change (Fbounded b (Fminus radix p (Fopp q))) in |- *.
apply Sterbenz; auto.
apply oppBounded; auto.
apply Rmult_le_reg_l with (r := INR 2); auto with real arith.
rewrite <- Rmult_assoc; rewrite Rinv_r; auto with real arith;
 rewrite Rmult_1_l; auto.
rewrite Fopp_correct; auto.
apply Rle_trans with (FtoRradix p); auto.
apply Rledouble; auto.
rewrite Fopp_correct; auto.
Qed.

Theorem Dekker3 :
 forall p q : float,
 (q <= 0)%R ->
 (radix * - q < p)%R ->
 Fbounded b p -> Fbounded b q -> Iminus (Iplus p q) p = (Iplus p q - p)%R :>R.
intros p q H' H'0 H'1 H'2.
apply MDekkerAux4; auto.
apply Sterbenz; auto.
apply IplusBounded; auto.
apply (FmultRadixInv b radix precision) with (y := (p + q)%R); auto.
apply Rmult_lt_reg_l with (r := INR 2); auto with real arith.
rewrite <- Rmult_assoc; rewrite Rinv_r; auto with real arith;
 rewrite Rmult_1_l.
replace (2%nat * (p + q))%R with (2%nat * p + 2%nat * q)%R;
 [ idtac | simpl in |- *; ring ].
apply Rplus_lt_reg_l with (r := (- FtoR radix p)%R).
replace (- FtoR radix p + FtoR radix p)%R with 0%R;
 [ idtac | simpl in |- *; ring ].
replace (- FtoR radix p + (2%nat * p + 2%nat * q))%R with (p + 2%nat * q)%R;
 [ idtac | simpl in |- *; unfold FtoRradix in |- *; ring ].
apply Rplus_lt_reg_l with (r := (2%nat * - q)%R).
replace (2%nat * - q + 0)%R with (2%nat * - q)%R;
 [ idtac | simpl in |- *; ring ].
replace (2%nat * - q + (p + 2%nat * q))%R with (FtoRradix p);
 [ idtac | simpl in |- *; ring ]; auto.
apply (RoundedModeMult b radix) with (P := Closest b radix) (r := (p + q)%R);
 auto.
apply ClosestRoundedModeP with (precision := precision); auto.
apply Rle_trans with (FtoR radix p); auto.
replace (FtoR radix p) with (p + 0)%R; [ idtac | fold FtoRradix; ring ].
apply Rplus_le_compat_l; auto.
apply Rledouble; auto.
apply Rlt_le; auto.
apply Rle_lt_trans with (2 := H'0).
apply Rle_trans with (- q)%R; auto.
replace 0%R with (-0)%R; auto with real.
apply Rledouble; auto.
replace 0%R with (-0)%R; auto with real.
Qed.

Theorem MDekkerAux5 :
 forall p q : float,
 Fbounded b p ->
 Fbounded b q ->
 Iminus (Iplus (Fopp p) (Fopp q)) (Fopp p) =
 (Iplus (Fopp p) (Fopp q) - Fopp p)%R :>R ->
 Iminus (Iplus p q) p = (Iplus p q - p)%R :>R.
intros p q H' H'0 H'1.
rewrite <- (Fopp_Fopp (Iminus (Iplus p q) p)); auto.
repeat rewrite IminusPlus; auto.
rewrite IplusOp; auto.
rewrite <- IminusPlus.
rewrite IplusOp; auto.
unfold FtoRradix in |- *; rewrite Fopp_correct.
unfold FtoRradix in H'1; rewrite H'1.
rewrite <- IplusOp; auto.
repeat rewrite Fopp_correct.
ring.
Qed.

Theorem MDekker :
 forall p q : float,
 Fbounded b p ->
 Fbounded b q ->
 (Rabs q <= Rabs p)%R -> Iminus (Iplus p q) p = (Iplus p q - p)%R :>R.
intros p q H' H'0 H'1.
case (Rle_or_lt 0 p); intros Rl1.
rewrite (Rabs_right p) in H'1; auto with real.
case (Rle_or_lt 0 q); intros Rl2.
rewrite (Rabs_right q) in H'1; auto with real.
apply Dekker1_FTS; auto.
rewrite (Rabs_left1 q) in H'1; auto with real.
case (Rle_or_lt p (radix * - q)); intros Rl3.
apply Dekker2_FTS; auto.
apply Dekker3; auto.
apply Rlt_le; auto.
rewrite (Rabs_left1 p) in H'1; auto with real.
apply MDekkerAux5; auto.
case (Rle_or_lt 0 q); intros Rl2.
rewrite (Rabs_right q) in H'1; auto with real.
case (Rle_or_lt (- p) (radix * q)); intros Rl3.
apply Dekker2_FTS; auto with zarith.
unfold FtoRradix in |- *; rewrite Fopp_correct; auto.
replace 0%R with (-0)%R; auto with real.
unfold FtoRradix in |- *; repeat rewrite Fopp_correct; auto.
rewrite Ropp_involutive; auto.
unfold FtoRradix in |- *; repeat rewrite Fopp_correct; auto with real.
rewrite Ropp_involutive; auto.
apply oppBounded; auto.
apply oppBounded; auto.
apply Dekker3; auto with zarith.
unfold FtoRradix in |- *; rewrite Fopp_correct; auto.
replace 0%R with (-0)%R; auto with real.
unfold FtoRradix in |- *; repeat rewrite Fopp_correct; auto.
rewrite Ropp_involutive; auto.
apply oppBounded; auto.
apply oppBounded; auto.
rewrite (Rabs_left1 q) in H'1; auto with real.
apply Dekker1_FTS; auto with zarith.
unfold FtoRradix in |- *; rewrite Fopp_correct; auto.
replace 0%R with (-0)%R; auto with real.
unfold FtoRradix in |- *; repeat rewrite Fopp_correct; auto.
apply oppBounded; auto.
apply oppBounded; auto.
Qed.

Theorem Dekker_FTS :
 forall p q : float,
 Fbounded b p ->
 Fbounded b q ->
 (Rabs q <= Rabs p)%R ->
 Iminus q (Iminus (Iplus p q) p) = (p + q - Iplus p q)%R :>R.
intros p q H' H'0 H'1.
apply MDekkerAux1; auto.
apply MDekker; auto.
Qed.
End Fast.

Section GenericA.
Variable bo : Fbound.
Variable radix : Z.
Variable p : nat.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.
Hypothesis radixMoreThanOne : (1 < radix)%Z.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ radixMoreThanOne).

Hypothesis pGivesBound : Zpos (vNum bo) = Zpower_nat radix p.
Hypothesis precisionGreaterThanOne : 3 <= p.

Hypothesis Evenradix: (Even radix).

Variable a b e x y:float.

Hypothesis eLea: (Rabs e <= /2*Fulp bo radix p a)%R.
Hypothesis eLeb: (Rabs e <= /2*Fulp bo radix p b)%R.
Hypothesis xDef: Closest bo radix (a+b)%R x.
Hypothesis yDef: Closest bo radix (a+b+e)%R y.
Hypothesis Nx: Fnormal radix bo x.
Hypothesis Ny: Fnormal radix bo y.
Hypothesis Cb: Fcanonic radix bo b.
Hypothesis Ca: Fcanonic radix bo a.

Hypothesis Fexpb: (- dExp bo < Fexp b)%Z.

Let Unmoins := (1 - (powerRZ radix (Z.succ (-p)))/2)%R.
Let Unplus := (1 + (powerRZ radix (Z.succ (-p)))/2)%R.

Lemma UnMoinsPos: (0 < Unmoins)%R.
unfold Unmoins.
assert (powerRZ radix (Z.succ (-p)) / 2 < 1)%R; auto with real.
apply Rmult_lt_reg_l with 2%R; auto with real.
apply Rle_lt_trans with (powerRZ radix (Z.succ (-p)));
  [right; field; auto with real| ring_simplify (2*1)%R].
apply Rle_lt_trans with (powerRZ radix (Z.succ (-1))); unfold Z.pred; auto with real zarith.
apply Rle_powerRZ;try apply Rle_IZR; auto with zarith.
simpl; auto with real.
apply Rplus_lt_reg_l with (powerRZ radix (Z.succ (- p)) / 2)%R.
ring_simplify; assumption.
Qed.

Lemma ClosestRoundeLeNormal: forall (z : R) (f : float),
       Closest bo radix z f ->
       Fnormal radix bo f ->
       (Rabs f <= (Rabs z) / Unmoins)%R.
intros.
generalize UnMoinsPos; intros U1.
apply Rmult_le_reg_l with Unmoins; auto with real.
apply Rle_trans with (Rabs z);[idtac|right; field; auto with real].
apply Rplus_le_reg_l with (-Rabs z+(1-Unmoins)*Rabs f)%R.
apply Rle_trans with (Rabs f-Rabs z)%R;[right; ring|idtac].
apply Rle_trans with (Rabs (f-z));[apply Rabs_triang_inv|idtac].
replace (f-z)%R with (-(z-f))%R;[rewrite Rabs_Ropp|ring].
apply Rle_trans with (Rabs (f) * (/ S 1 * powerRZ radix (Z.succ (- p))))%R.
unfold FtoRradix; apply ClosestErrorBoundNormal with bo; auto with zarith.
rewrite FcanonicFnormalizeEq; auto with zarith; left; auto.
unfold Unmoins, Rdiv; right; simpl; field.
Qed.

Lemma ClosestRoundeGeNormal: forall (z : R) (f : float),
       Closest bo radix z f ->
       Fnormal radix bo f ->
       (Rabs z <= (Rabs f) * Unplus)%R.
intros.
apply Rplus_le_reg_l with (-(Rabs f))%R.
apply Rle_trans with (Rabs z-Rabs f)%R;[right; ring|idtac].
apply Rle_trans with (Rabs (z-f));[apply Rabs_triang_inv|idtac].
apply Rle_trans with (Rabs (f) * (/ S 1 * powerRZ radix (Z.succ (- p))))%R.
unfold FtoRradix; apply ClosestErrorBoundNormal with bo; auto with zarith.
rewrite FcanonicFnormalizeEq; auto with zarith; left; auto.
unfold Unplus; simpl; unfold Rdiv; right; field.
Qed.

Lemma abeLeab: (Rabs b <= Rabs a)%R -> (2*powerRZ radix (Fexp b) <= Rabs (a+b))%R
               -> (Rabs (a+b) <= Rabs (a+b+e) *4/3)%R.
intros.
assert (0 <3)%R; [apply Rlt_trans with 2%R; auto with real|idtac].
assert (0 <4)%R; [apply Rlt_trans with 3%R; auto with real|idtac].
apply Rmult_le_reg_l with (3/4)%R; auto with real.
unfold Rdiv; apply Rmult_lt_0_compat; auto with real.
apply Rle_trans with (Rabs (a + b + e));[idtac|right; field; auto with real].
apply Rplus_le_reg_l with (/4*Rabs (a + b))%R.
apply Rle_trans with (Rabs (a+b));[right; field; auto with real|idtac].
pattern (a+b)%R at 1; replace (a+b)%R with ((a+b+e)+(-e))%R;[idtac|ring].
apply Rle_trans with (Rabs (a+b+e)+ Rabs (-e))%R;[apply Rabs_triang|idtac].
rewrite Rplus_comm; apply Rplus_le_compat_r.
rewrite Rabs_Ropp; apply Rle_trans with (1:=eLeb).
apply Rmult_le_reg_l with 4%R; auto with real.
apply Rle_trans with (Rabs (a+b));[idtac|right; field; auto with real].
apply Rle_trans with (2:=H0).
unfold Fulp; rewrite FcanonicFnormalizeEq; auto with zarith.
right; field; auto with real.
Qed.

Lemma xLe2y_aux1: (Rabs b <= Rabs a)%R -> (powerRZ radix (Fexp b) = Rabs (a+b))%R
              -> (Rabs x <= 2*Rabs y)%R.
intros.
apply Rle_trans with (Rabs (a+b));[right|idtac].
rewrite <- H0; unfold FtoRradix; rewrite <- Fabs_correct; auto.
apply trans_eq with (FtoRradix (Float 1 (Fexp b)));
   [apply sym_eq |unfold FtoRradix, FtoR; simpl; ring].
apply RoundedModeProjectorIdemEq with (P:=(Closest bo radix)) (3:=pGivesBound);
   auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
split; simpl.
apply vNumbMoreThanOne with radix p; auto with zarith.
assert (Fbounded bo b); auto with zarith.
apply FcanonicBound with radix; auto.
replace (FtoR radix (Float 1 (Fexp b))) with (Rabs (a+b)).
apply ClosestFabs with p; auto with zarith.
rewrite <- H0; unfold FtoR; simpl; ring.
rewrite <- H0; apply Rmult_le_reg_l with (/2)%R; auto with real.
apply Rle_trans with (Rabs y);[idtac|right; field; auto with real].
elim Evenradix; intros n Hn.
apply Rle_trans with (FtoRradix (Float n ((Fexp b)-1))).
right; unfold FtoRradix, FtoR; simpl.
unfold Zminus; rewrite powerRZ_add; try apply IZR_neq; auto with real zarith; simpl.
ring_simplify (radix*1)%R; rewrite Hn; rewrite mult_IZR; simpl; field; auto with real zarith.
apply IZR_neq; auto with zarith.
unfold FtoRradix; apply RoundAbsMonotonel with bo p (Closest bo radix) (a+b+e)%R;
   auto with real zarith.
apply ClosestRoundedModeP with p; auto with zarith.
split; simpl; auto with zarith.
rewrite pGivesBound; apply Z.lt_trans with (Zpower_nat radix 1); auto with zarith.
unfold Zpower_nat; simpl; rewrite Z.abs_eq; auto with zarith.
apply Zpower_nat_monotone_lt; auto with zarith.
replace (a+b+e)%R with ((a+b)-(-e))%R;[idtac|ring].
apply Rle_trans with (Rabs (a+b)-Rabs (-e))%R;[idtac|apply Rabs_triang_inv].
apply Rle_trans with (powerRZ radix (Fexp b)-/2*(powerRZ radix (Fexp b)))%R.
right; unfold FtoRradix, FtoR; simpl.
unfold Zminus; rewrite powerRZ_add; try apply IZR_neq;auto with real zarith; simpl.
ring_simplify (radix*1)%R; rewrite Hn; repeat rewrite mult_IZR; simpl; field; auto with real zarith.
apply IZR_neq; auto with zarith.
unfold Rminus; apply Rplus_le_compat; auto with real.
apply Ropp_le_contravar; rewrite Rabs_Ropp.
apply Rle_trans with (1:=eLeb); right.
unfold Fulp; rewrite FcanonicFnormalizeEq; auto with zarith real.
Qed.

Lemma xLe2y_aux2 : (Rabs b <= Rabs a)%R -> (Rabs x <= 2*Rabs y)%R.
intros.
assert ((a+b=0)%R \/ (powerRZ radix (Fexp b) = Rabs (a+b))%R
  \/ (2*powerRZ radix (Fexp b) <= Rabs (a+b))%R).
unfold FtoRradix; rewrite <- Fplus_correct; auto.
rewrite <- Fabs_correct; auto.
unfold FtoR.
replace (Fexp (Fabs (Fplus radix a b))) with (Fexp (Fplus radix a b)); auto with zarith.
replace (Fexp (Fplus radix a b)) with (Fexp b).
case (Zle_lt_or_eq 0 (Z.abs (Fnum (Fplus radix a b)))); auto with zarith.
intros; case (Zle_lt_or_eq 1 (Z.abs (Fnum (Fplus radix a b)))); auto with zarith.
right; right.
apply Rmult_le_compat_r; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
apply Rle_IZR; generalize H1; unfold Fabs; simpl; auto with zarith.
intros; right; left.
 generalize H1; unfold Fabs; simpl; auto with zarith real.
intros H2; rewrite <- H2; simpl; ring.
left; replace (Fnum (Fplus radix a b)) with 0%Z;[simpl; ring|auto with zarith].
apply sym_eq; unfold Fplus; simpl; apply Zmin_le2.
apply Fcanonic_Rle_Zle with radix bo p; auto with zarith.
case H0;intros H1;[idtac|case H1; clear H1; intros H1]; clear H0.
replace (FtoRradix x) with 0%R.
rewrite Rabs_R0; apply Rmult_le_pos; auto with real.
apply Rabs_pos.
apply trans_eq with (FtoRradix (Float 0 (-(dExp bo))));
   [unfold FtoRradix, FtoR; simpl; ring|idtac].
apply RoundedModeProjectorIdemEq with (P:=(Closest bo radix)) (3:=pGivesBound);
  auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
split; simpl; auto with zarith.
replace (FtoR radix (Float 0 (- dExp bo))) with (a+b)%R; auto.
rewrite H1; unfold FtoR; simpl; ring.
apply xLe2y_aux1; auto.
generalize UnMoinsPos; intros U1.
apply Rle_trans with ((Rabs (a+b))/Unmoins)%R.
apply ClosestRoundeLeNormal; auto.
apply Rmult_le_reg_l with Unmoins; auto with real.
apply Rle_trans with (Rabs (a+b));[right; field; auto with real|idtac].
apply Rle_trans with (Rabs (a + b + e) * 4 / 3)%R.
apply abeLeab; auto.
assert (0 <3)%R; [apply Rlt_trans with 2%R; auto with real|idtac].
assert (0 <4)%R; [apply Rlt_trans with 3%R; auto with real|idtac].
apply Rmult_le_reg_l with (3/4)%R; [unfold Rdiv; apply Rmult_lt_0_compat; auto with real|idtac].
apply Rle_trans with (Rabs (a+b+e));[right; field; auto with real|idtac].
apply Rle_trans with ((Rabs y)*Unplus)%R.
apply ClosestRoundeGeNormal; auto.
apply Rle_trans with (Rabs y *(3/4*2*Unmoins))%R;[idtac|right; ring].
apply Rmult_le_compat_l; auto with real.
apply Rabs_pos.
apply Rmult_le_reg_l with 2%R; auto with real.
apply Rle_trans with (3*Unmoins)%R;[idtac|right; field; auto with real].
unfold Unplus, Unmoins.
apply Rplus_le_reg_l with (-2+3* (powerRZ radix (Z.succ (- p)) / 2))%R.
ring_simplify.
apply Rle_trans with (powerRZ radix 0);[idtac|simpl; auto with real].
apply Rle_trans with (powerRZ radix (3-p));
  [idtac|apply Rle_powerRZ; try apply Rle_IZR; auto with zarith real].
apply Rle_trans with ((5/2)*(powerRZ radix (Z.succ (-p))))%R;[right; field; auto with real|idtac].
apply Rle_trans with ((radix*radix)*(powerRZ radix (Z.succ (- p))))%R.
apply Rmult_le_compat_r; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
apply Rle_trans with 4%R; auto with real zarith.
apply Rmult_le_reg_l with 2%R; auto with real.
apply Rle_trans with 5%R;[simpl; right; field; auto with real|idtac].
apply Rle_trans with 8%R; [auto with real zarith|simpl; right; ring].
rewrite <- mult_IZR; apply Rle_IZR.
apply Z.le_trans with (2*2)%Z; auto with zarith.
apply Zmult_le_compat; auto with zarith.
unfold Z.succ, Zminus; repeat rewrite powerRZ_add; try apply IZR_neq; auto with real zarith.
simpl; right; ring.
Qed.

Lemma yLe2x_aux : (Rabs b <= Rabs a)%R -> ~(FtoRradix x=0)%R
              -> (Rabs y <= 2*Rabs x)%R.
intros.
assert ((a+b=0)%R \/ (powerRZ radix (Fexp b) <= Rabs (a+b))%R).
unfold FtoRradix; rewrite <- Fplus_correct; auto.
rewrite <- Fabs_correct; auto.
unfold FtoR.
replace (Fexp (Fabs (Fplus radix a b))) with (Fexp (Fplus radix a b)); auto with zarith.
replace (Fexp (Fplus radix a b)) with (Fexp b).
case (Zle_lt_or_eq 0 (Z.abs (Fnum (Fplus radix a b)))); auto with zarith.
intros; right.
apply Rle_trans with (1%Z*(powerRZ radix (Fexp b)))%R;[right; simpl; ring|idtac].
apply Rmult_le_compat_r; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
apply Rle_IZR; generalize H1; simpl; auto with zarith.
intros; left; replace (Fnum (Fplus radix a b)) with 0%Z;[simpl; ring|auto with zarith].
apply sym_eq; unfold Fplus; simpl; apply Zmin_le2.
apply Fcanonic_Rle_Zle with radix bo p; auto with zarith.
case H1;intros H2; clear H1.
absurd (FtoRradix x=0)%R; auto with real.
apply sym_eq; apply trans_eq with (FtoRradix (Float 0 (-(dExp bo))));
   [unfold FtoRradix, FtoR; simpl; ring|idtac].
apply RoundedModeProjectorIdemEq with (P:=(Closest bo radix)) (3:=pGivesBound);
  auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
split; simpl; auto with zarith.
replace (FtoR radix (Float 0 (- dExp bo))) with (a+b)%R; auto.
rewrite H2; unfold FtoR; simpl; ring.
generalize UnMoinsPos; intros U1.
apply Rle_trans with ((Rabs (a+b+e))/Unmoins)%R.
apply ClosestRoundeLeNormal; auto.
apply Rmult_le_reg_l with Unmoins; auto with real.
apply Rle_trans with (Rabs (a+b+e));[right; field; auto with real|idtac].
apply Rle_trans with (Rabs (a + b) * 3 / 2)%R.
apply Rle_trans with (Rabs (a+b)+Rabs e)%R;[apply Rabs_triang|idtac].
apply Rle_trans with (Rabs (a + b) + Rabs (a + b) /2)%R;
   [apply Rplus_le_compat_l|right; field; auto with real].
apply Rle_trans with (1:=eLeb); unfold Rdiv; rewrite Rmult_comm.
apply Rmult_le_compat_r; auto with real.
unfold Fulp; rewrite FcanonicFnormalizeEq; auto with zarith real.
assert (0 <3)%R; [apply Rlt_trans with 2%R; auto with real|idtac].
apply Rmult_le_reg_l with (2/3)%R; [unfold Rdiv; apply Rmult_lt_0_compat; auto with real|idtac].
apply Rle_trans with (Rabs (a+b));[right; field; apply prod_neq_R0; auto with real|idtac].
apply Rle_trans with ((Rabs x)*Unplus)%R.
apply ClosestRoundeGeNormal; auto.
apply Rle_trans with (Rabs x *(2/3*2*Unmoins))%R;[idtac|right; ring].
apply Rmult_le_compat_l; auto with real.
apply Rabs_pos.
apply Rmult_le_reg_l with 3%R; auto with real.
apply Rle_trans with (4*Unmoins)%R;[idtac|right; field; auto with real].
unfold Unplus, Unmoins.
apply Rplus_le_reg_l with (-3+4* (powerRZ radix (Z.succ (- p)) / 2))%R.
ring_simplify.
apply Rle_trans with (powerRZ radix 0);[idtac|simpl; auto with real].
apply Rle_trans with (powerRZ radix (3-p));
  [idtac|apply Rle_powerRZ; auto with zarith real].
apply Rle_trans with ((7/2)*(powerRZ radix (Z.succ (-p))))%R;[right; field; auto with real|idtac].
apply Rle_trans with ((radix*radix)*(powerRZ radix (Z.succ (- p))))%R.
apply Rmult_le_compat_r; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
apply Rle_trans with 4%R; auto with real zarith.
apply Rmult_le_reg_l with 2%R; auto with real.
apply Rle_trans with 7%R;[simpl; right; field; auto with real|idtac].
apply Rle_trans with 8%R; [auto with real zarith|simpl; right; ring].
rewrite <- mult_IZR; apply Rle_IZR.
apply Z.le_trans with (2*2)%Z; auto with zarith.
apply Zmult_le_compat; auto with zarith.
unfold Z.succ, Zminus; repeat rewrite powerRZ_add; auto with real zarith.
simpl; right; field.
apply IZR_neq; omega.
apply IZR_neq; omega.
apply Rle_IZR; omega.
Qed.

End GenericA.

Section GenericB.
Variable bo : Fbound.
Variable radix : Z.
Variable p : nat.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.
Hypothesis radixMoreThanOne : (1 < radix)%Z.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ radixMoreThanOne).

Hypothesis pGivesBound : Zpos (vNum bo) = Zpower_nat radix p.
Hypothesis precisionGreaterThanOne : 3 <= p.

Hypothesis Evenradix: (Even radix).

Variable a b e x y:float.

Hypothesis eLea: (Rabs e <= /2*Fulp bo radix p a)%R.
Hypothesis eLeb: (Rabs e <= /2*Fulp bo radix p b)%R.
Hypothesis xDef: Closest bo radix (a+b)%R x.
Hypothesis yDef: Closest bo radix (a+b+e)%R y.
Hypothesis Nx: Fnormal radix bo x.
Hypothesis Ny: Fnormal radix bo y.
Hypothesis Cb: Fcanonic radix bo b.
Hypothesis Ca: Fcanonic radix bo a.

Hypothesis Fexpb: (- dExp bo < Fexp b)%Z.
Hypothesis Fexpa: (- dExp bo < Fexp a)%Z.

Hypothesis dsd: ((0<= y)%R -> (0<= x)%R) /\ ((y <= 0)%R -> (x <= 0)%R).

Lemma xLe2y : (Rabs x <= 2*Rabs y)%R.
case (Rle_or_lt (Rabs b) (Rabs a)); intros.
unfold FtoRradix; apply xLe2y_aux2 with bo p a b e; auto.
unfold FtoRradix; apply xLe2y_aux2 with bo p b a e; auto with real; fold FtoRradix.
rewrite Rplus_comm; auto.
replace (b+a+e)%R with (a+b+e)%R; auto; ring.
Qed.

Lemma yLe2x: ~(FtoRradix x=0)%R -> (Rabs y <= 2*Rabs x)%R.
case (Rle_or_lt (Rabs b) (Rabs a)); intros.
unfold FtoRradix; apply yLe2x_aux with bo p a b e; auto.
unfold FtoRradix; apply yLe2x_aux with bo p b a e; auto with real; fold FtoRradix.
rewrite Rplus_comm; auto.
replace (b+a+e)%R with (a+b+e)%R; auto; ring.
Qed.

Lemma Subexact: exists v:float, (FtoRradix v=x-y)%R /\ (Fbounded bo v) /\
        (Fexp v=Z.min (Fexp x) (Fexp y))%Z.
case (Req_dec (FtoRradix x) 0); intros.
absurd (FtoRradix x =0)%R; auto with real.
assert (~ is_Fzero x).
apply FnormalNotZero with radix bo; auto.
unfold is_Fzero in H0.
unfold FtoRradix, FtoR; apply prod_neq_R0; auto with real zarith.
apply IZR_neq; auto.
apply Rgt_not_eq; apply powerRZ_lt, Rlt_IZR; omega.
case (Rle_or_lt 0 y); intros S.
exists (Fminus radix x y); split.
unfold FtoRradix; rewrite Fminus_correct; auto with real zarith.
split;[idtac|simpl; auto with zarith].
apply Sterbenz; auto with zarith; fold FtoRradix.
elim Nx; auto.
elim Ny; auto.
apply Rmult_le_reg_l with 2%R; auto with real.
apply Rle_trans with (FtoRradix y);[simpl; right; field; auto with real|idtac].
rewrite <- (Rabs_right y);[idtac|apply Rle_ge; auto].
rewrite <- (Rabs_right x).
apply yLe2x; auto.
elim dsd; intros I1 I2; apply Rle_ge; apply I1; auto.
rewrite <- (Rabs_right y);[idtac|apply Rle_ge; auto].
rewrite <- (Rabs_right x).
simpl; apply xLe2y; auto.
elim dsd; intros I1 I2; apply Rle_ge; apply I1; auto.
exists (Fopp (Fminus radix (Fopp x) (Fopp y))); split.
unfold FtoRradix; rewrite Fopp_correct; rewrite Fminus_correct; auto with real zarith.
rewrite Fopp_correct;rewrite Fopp_correct; ring.
split;[idtac|simpl; auto with zarith].
apply oppBounded.
apply Sterbenz; auto with zarith; fold FtoRradix.
apply oppBounded; elim Nx; auto.
apply oppBounded; elim Ny; auto.
apply Rmult_le_reg_l with 2%R; auto with real.
apply Rle_trans with (FtoRradix (Fopp y));[simpl; right; field; auto with real|idtac].
unfold FtoRradix; repeat rewrite Fopp_correct; fold FtoRradix.
rewrite <- (Rabs_left1 y); auto with real.
rewrite <- (Rabs_left1 x).
apply yLe2x; auto.
elim dsd; intros I1 I2; apply I2; auto with real.
unfold FtoRradix; repeat rewrite Fopp_correct; fold FtoRradix.
rewrite <- (Rabs_left1 y); auto with real.
rewrite <- (Rabs_left1 x).
simpl; apply xLe2y; auto.
elim dsd; intros I1 I2; apply I2; auto with real.
Qed.

End GenericB.

Section GenericC.
Variable bo : Fbound.
Variable radix : Z.
Variable p : nat.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.
Hypothesis radixMoreThanOne : (1 < radix)%Z.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ radixMoreThanOne).

Hypothesis pGivesBound : Zpos (vNum bo) = Zpower_nat radix p.
Hypothesis precisionGreaterThanOne : 1 < p.
Hypothesis Evenradix: (Even radix).

Lemma LSB_Pred: forall x y:float,
    (Rabs x < Rabs y)%R -> (LSB radix x <= LSB radix y)%Z
       -> (Rabs x <= Rabs y - powerRZ radix (LSB radix x))%R.
intros.
assert (exists nx:Z, (Rabs x=nx*powerRZ radix (LSB radix x))%R).
unfold FtoRradix; rewrite <- Fabs_correct; auto.
elim LSB_rep_min with radix (Fabs x); auto.
intros nx H1; exists nx; rewrite H1.
rewrite <- LSB_abs; auto.
assert (exists ny:Z, (Rabs y=ny*powerRZ radix (LSB radix x))%R).
unfold FtoRradix; rewrite <- Fabs_correct; auto.
elim LSB_rep_min with radix (Fabs y); auto.
intros ny1 H2; exists (ny1*Zpower_nat radix (Z.abs_nat (LSB radix y - LSB radix x)))%Z;
   rewrite H2.
rewrite <- LSB_abs; auto; rewrite mult_IZR; unfold FtoR; simpl;
  unfold FtoR; simpl.
rewrite Zpower_nat_Z_powerRZ.
rewrite Rmult_assoc; rewrite <- powerRZ_add; auto with real zarith.
replace (Z.abs_nat (LSB radix y - LSB radix x)+LSB radix x)%Z with (LSB radix y); auto with real.
rewrite <- Zabs_absolu; rewrite Z.abs_eq; auto with zarith.
apply IZR_neq; omega.
apply Rplus_le_reg_l with (powerRZ radix (LSB radix x)-Rabs x)%R.
ring_simplify.
elim H1; intros nx H1'; elim H2; intros ny H2'; rewrite H1'; rewrite H2'.
apply Rle_trans with ((IZR 1)*powerRZ radix (LSB radix x))%R;[simpl; right; ring|idtac].
apply Rle_trans with ((ny-nx)*powerRZ radix (LSB radix x))%R;[idtac|simpl; right; ring].
apply Rmult_le_compat_r; auto with real zarith.
assert (ny-nx=(ny-nx)%Z)%R.
unfold Zminus; rewrite plus_IZR; rewrite Ropp_Ropp_IZR; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
rewrite <- minus_IZR; apply Rle_IZR.
assert (0 < ny - nx)%Z; auto with real zarith.
apply Zlt_Rlt; rewrite minus_IZR.
apply Rplus_lt_reg_r with nx.
ring_simplify.
apply Rmult_lt_reg_l with ( powerRZ radix (LSB radix x)); auto with real zarith.
apply powerRZ_lt, Rlt_IZR; omega.
rewrite Rmult_comm; rewrite <- H1'; rewrite Rmult_comm; rewrite <- H2'; auto.
Qed.

Variables x1 x2 y f:float.
Hypothesis x1Def: Closest bo radix (x1+x2)%R x1.
Hypothesis fDef : Closest bo radix (x1+x2+y)%R f.
Hypothesis yLe: (MSB radix y < LSB radix x2)%Z.
Hypothesis Nx1: Fnormal radix bo x1.
Hypothesis x1Pos: (0 < x1)%R.
Hypothesis x2NonZero: ~(FtoRradix x2 =0)%R.
Hypothesis x1Exp: (-dExp bo < Fexp x1)%Z.

Lemma Midpoint_aux_aux:
    (FtoRradix x1= f) \/ (exists v:float, (FtoRradix v=x2)%R /\ (Fexp x1 -2 <= Fexp v)%Z).
case (Z.eq_dec (Fnum x1) (nNormMin radix p)); intros H1.
case (Rle_or_lt 0 x2); intros G.
assert (Rabs x2 <= powerRZ radix (Fexp x1)/2)%R.
apply Rmult_le_reg_l with (INR 2); auto with real zarith.
apply Rle_trans with (powerRZ radix (Fexp x1));[idtac|right; simpl; field; auto with real].
replace (FtoRradix x2) with ((x1+x2) -x1)%R;[idtac|ring].
apply Rle_trans with (Fulp bo radix p x1).
unfold FtoRradix; apply ClosestUlp; auto with zarith.
unfold Fulp; rewrite FcanonicFnormalizeEq; auto with zarith real; left; auto.
case H; clear H; intros H.
assert (Rabs (x2 + y) < powerRZ radix (Fexp x1) / 2)%R.
apply Rle_lt_trans with (Rabs x2+Rabs y)%R;[apply Rabs_triang|idtac].
apply Rplus_lt_reg_l with (-Rabs y)%R.
ring_simplify (- Rabs y + (Rabs x2 + Rabs y))%R.
elim Evenradix; intros n Hn.
assert (powerRZ radix (Fexp x1) / 2 = Float n (Fexp x1 -1))%R.
unfold FtoRradix, FtoR; simpl; unfold Zminus;
  rewrite powerRZ_add; try apply IZR_neq; auto with real zarith; simpl.
repeat (rewrite Hn;rewrite mult_IZR); simpl; field; auto with real zarith.
apply IZR_neq; auto with zarith.
apply Rle_lt_trans with (Rabs (Float n (Fexp x1 - 1)) - powerRZ radix (LSB radix x2))%R.
apply LSB_Pred; auto.
rewrite (Rabs_right (Float n (Fexp x1 - 1))); auto with real.
rewrite <- H0; auto.
apply Rle_ge; unfold FtoRradix; apply LeFnumZERO; simpl; auto with real zarith.
apply Z.le_trans with (MSB radix x2).
apply LSB_le_MSB; auto with zarith.
contradict x2NonZero.
unfold FtoRradix; apply is_Fzero_rep1; auto.
apply Z.le_trans with (Fexp x1-1)%Z.
2: apply Z.le_trans with (Fexp ((Float n (Fexp x1 - 1)))); auto with zarith.
2: apply Fexp_le_LSB.
assert (MSB radix x2 < Fexp x1)%Z; auto with zarith.
apply Zlt_powerRZ with radix; auto with real.
apply Rle_IZR; omega.
apply Rle_lt_trans with (FtoR radix (Float (S 0) (MSB radix x2)));
       [right; unfold FtoR; simpl; ring|idtac].
apply Rle_lt_trans with (FtoR radix (Fabs x2)).
apply MSB_le_abs; auto.
contradict x2NonZero.
unfold FtoRradix; apply is_Fzero_rep1; auto.
rewrite Fabs_correct; auto; fold FtoRradix.
apply Rlt_le_trans with (1:=H).
unfold Rdiv; apply Rle_trans with ( (powerRZ radix (Fexp x1)*1))%R; auto with real zarith.
apply Rmult_le_compat_l; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
apply Rle_trans with (/1)%R; auto with real.
rewrite (Rabs_right (Float n (Fexp x1 - 1))); auto with real.
2: apply Rle_ge; unfold FtoRradix; apply LeFnumZERO; simpl; auto with real zarith.
rewrite <- H0; unfold Rminus; rewrite Rplus_comm.
apply Rplus_lt_compat_r; apply Ropp_lt_contravar.
unfold FtoRradix; rewrite <- Fabs_correct; auto.
apply Rlt_le_trans with (FtoR radix (Float (S 0) (Z.succ (MSB radix y))))%R.
apply abs_lt_MSB; auto.
unfold FtoR; simpl; ring_simplify (1 * powerRZ radix (Z.succ (MSB radix y)))%R.
apply Rle_powerRZ; auto with real zarith.
apply Rle_IZR; omega.
assert (powerRZ radix (Fexp x1 + p - 1) <= x1)%R.
apply Rle_trans with (((nNormMin radix p))*(powerRZ radix (Fexp x1)))%R.
unfold nNormMin; rewrite Zpower_nat_Z_powerRZ.
replace (Fexp x1 + p - 1)%Z with (Fexp x1 + pred p)%Z.
rewrite powerRZ_add; try apply IZR_neq; auto with real zarith; right;ring.
rewrite inj_pred; unfold Z.pred; auto with zarith.
unfold FtoRradix, FtoR; apply Rmult_le_compat_r; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
apply Rle_IZR; auto with zarith.
left; unfold FtoRradix.
apply ImplyClosestStrict with bo p (x1+x2+y)%R (Fexp x1); auto with zarith.
elim x1Def; auto.
left; auto.
replace (x1+x2+y)%R with (x1+(x2+y))%R;[idtac|ring].
apply Rle_trans with ((powerRZ radix (Fexp x1 + p - 1) + 0))%R;[right; ring|idtac].
apply Rplus_le_compat; auto.
apply Rplus_le_reg_l with (-y)%R.
ring_simplify.
apply Rle_trans with (Rabs (-y));[apply RRle_abs|rewrite Rabs_Ropp].
rewrite <- (Rabs_right x2); auto with real.
case (Req_dec y 0); intros.
rewrite H3; rewrite Rabs_R0; auto with real.
apply Rabs_pos.
case (Rle_or_lt (Rabs y) (Rabs x2)); auto.
intros I; absurd (MSB radix y < LSB radix x2)%Z; auto.
apply Zle_not_lt.
apply Z.le_trans with (MSB radix x2).
apply LSB_le_MSB; auto with zarith.
contradict x2NonZero.
unfold FtoRradix; apply is_Fzero_rep1; auto.
apply MSB_monotone; auto with zarith real.
contradict x2NonZero.
unfold FtoRradix; apply is_Fzero_rep1; auto.
contradict H3.
unfold FtoRradix; apply is_Fzero_rep1; auto.
repeat rewrite Fabs_correct; auto with real zarith.
fold FtoRradix; replace (x1+x2+y-x1)%R with (x2+y)%R; auto with real; ring.
right; elim Evenradix; intros n Hn.
exists (Float n (Fexp x1-1)); split;[idtac|simpl; auto with zarith].
apply trans_eq with (Rabs x2);[rewrite H|idtac].
unfold FtoRradix, FtoR; simpl; unfold Zminus; rewrite powerRZ_add; try apply IZR_neq; auto with real zarith.
simpl; ring_simplify (radix*1)%R; rewrite Hn; rewrite mult_IZR; simpl; field; auto with real zarith.
apply IZR_neq; auto with zarith.
rewrite Rabs_right; auto with real.
assert (Rabs x2 <= powerRZ radix (Fexp x1-1)/2)%R.
apply Rmult_le_reg_l with 2%R; auto with real.
apply Rle_trans with (powerRZ radix (Fexp x1-1));[idtac|right; simpl; field; auto with real].
assert (FPred bo radix p x1 - (x1 + x2) = -(powerRZ radix (Fexp x1-1)-Rabs x2))%R.
apply trans_eq with (-(Fminus radix x1 (FPred bo radix p x1) - Rabs x2))%R.
rewrite (Rabs_left x2); auto with real.
unfold FtoRradix; rewrite Fminus_correct; auto with zarith; ring.
unfold FtoRradix; rewrite FPredDiff3; auto with zarith.
unfold FtoR, Z.pred,Zminus; simpl;ring.
apply Rplus_le_reg_l with (-(Rabs x2))%R.
ring_simplify (- Rabs x2 + 2 * Rabs x2)%R.
apply Rle_trans with (Rabs ((FPred bo radix p x1)-(x1+x2)))%R.
pattern (FtoRradix x2) at 1; replace (FtoRradix x2) with (-(x1-(x1+x2)))%R;[rewrite Rabs_Ropp|ring].
elim x1Def; intros Y1 Y2; unfold FtoRradix; apply Y2; auto.
apply FBoundedPred; auto with zarith.
rewrite H; rewrite Rabs_Ropp.
rewrite Rabs_right.
right; ring.
apply Rle_ge; apply Rplus_le_reg_l with (Rabs x2).
ring_simplify.
case (Rle_or_lt (Rabs x2) (powerRZ radix (Fexp x1 - 1))); auto; intros.
absurd (Rabs x2 <= Rabs (FPred bo radix p x1 - (x1 + x2)))%R.
rewrite H; rewrite Rabs_Ropp; apply Rlt_not_le.
rewrite Rabs_left; auto with real.
apply Rplus_lt_reg_r with (-Rabs x2+powerRZ radix (Fexp x1-1))%R.
ring_simplify; auto with real zarith.
apply powerRZ_lt, Rlt_IZR; omega.
pattern (FtoRradix x2) at 1; replace (FtoRradix x2) with (-(x1-(x1+x2)))%R;[rewrite Rabs_Ropp|ring].
elim x1Def; intros Y1 Y2; unfold FtoRradix; apply Y2; auto.
apply FBoundedPred; auto with zarith.
case H; clear H; intros H.
assert (Rabs (x2 + y) < powerRZ radix (Fexp x1-1) / 2)%R.
apply Rle_lt_trans with (Rabs x2+Rabs y)%R;[apply Rabs_triang|idtac].
apply Rplus_lt_reg_l with (-Rabs y)%R.
ring_simplify (- Rabs y + (Rabs x2 + Rabs y))%R.
elim Evenradix; intros n Hn.
assert (powerRZ radix (Fexp x1-1) / 2 = Float n (Fexp x1 -2))%R.
unfold FtoRradix, FtoR; simpl; unfold Zminus;
  repeat rewrite powerRZ_add; try apply IZR_neq; auto with real zarith; simpl.
repeat (rewrite Hn; rewrite mult_IZR);simpl; field; auto with real zarith.
apply IZR_neq; auto with zarith.
apply Rle_lt_trans with (Rabs (Float n (Fexp x1 - 2)) - powerRZ radix (LSB radix x2))%R.
apply LSB_Pred; auto.
rewrite (Rabs_right (Float n (Fexp x1 - 2))); auto with real.
rewrite <- H0; auto.
apply Rle_ge; unfold FtoRradix; apply LeFnumZERO; simpl; auto with real zarith.
apply Z.le_trans with (MSB radix x2).
apply LSB_le_MSB; auto with zarith.
contradict x2NonZero.
unfold FtoRradix; apply is_Fzero_rep1; auto.
apply Z.le_trans with (Fexp x1-2)%Z.
2: apply Z.le_trans with (Fexp ((Float n (Fexp x1 - 2)))); auto with zarith.
2: apply Fexp_le_LSB.
assert (MSB radix x2 < Fexp x1-1)%Z; auto with zarith.
apply Zlt_powerRZ with radix; auto with real.
apply Rle_IZR; omega.
apply Rle_lt_trans with (FtoR radix (Float (S 0) (MSB radix x2)));
       [right; unfold FtoR; simpl; ring|idtac].
apply Rle_lt_trans with (FtoR radix (Fabs x2)).
apply MSB_le_abs; auto.
contradict x2NonZero.
unfold FtoRradix; apply is_Fzero_rep1; auto.
rewrite Fabs_correct; auto; fold FtoRradix.
apply Rlt_le_trans with (1:=H).
unfold Rdiv; apply Rle_trans with ( (powerRZ radix (Fexp x1-1)*1))%R; auto with real zarith.
apply Rmult_le_compat_l; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
apply Rle_trans with (/1)%R; auto with real.
rewrite (Rabs_right (Float n (Fexp x1 - 2))); auto with real.
2: apply Rle_ge; unfold FtoRradix; apply LeFnumZERO; simpl; auto with real zarith.
rewrite <- H0; unfold Rminus; rewrite Rplus_comm.
apply Rplus_lt_compat_r; apply Ropp_lt_contravar.
unfold FtoRradix; rewrite <- Fabs_correct; auto.
apply Rlt_le_trans with (FtoR radix (Float (S 0) (Z.succ (MSB radix y))))%R.
apply abs_lt_MSB; auto.
unfold FtoR; simpl; ring_simplify (1 * powerRZ radix (Z.succ (MSB radix y)))%R.
apply Rle_powerRZ; auto with real zarith.
apply Rle_IZR; omega.
assert (powerRZ radix (Fexp x1 + p - 1) = x1)%R.
unfold FtoRradix, FtoR; replace (Fexp x1+p-1)%Z with (Fexp x1+pred p)%Z.
rewrite powerRZ_add; auto with real zarith.
rewrite H1; unfold nNormMin; rewrite Zpower_nat_Z_powerRZ; auto with real zarith.
apply IZR_neq; omega.
rewrite inj_pred; auto with zarith; unfold Z.pred; ring.
left; unfold FtoRradix.
apply ImplyClosestStrict with bo p (x1+x2+y)%R (Fexp x1-1)%Z; auto with zarith.
elim x1Def; auto.
left; auto.
replace (x1+x2+y)%R with (x1+(x2+y))%R;[idtac|ring].
apply Rle_trans with ((powerRZ radix (Fexp x1 + p - 1) +
    -((powerRZ radix (Fexp x1-1))/2)))%R.
apply Rplus_le_reg_l with ((powerRZ radix (Fexp x1-1))/2)%R.
ring_simplify.
apply Rle_trans with (powerRZ radix (Fexp x1 - 1 + p - 1)+powerRZ radix (Fexp x1 - 1))%R.
rewrite Rplus_comm; apply Rplus_le_compat_l.
apply Rle_trans with (powerRZ radix (Fexp x1 - 1) *1)%R;
   [unfold Rdiv; apply Rmult_le_compat_l; auto with real zarith|right; ring].
apply powerRZ_le, Rlt_IZR; omega.
apply Rle_trans with (/1)%R; auto with real.
apply Rle_trans with (powerRZ radix ((Fexp x1 - 1 + p - 1)+1)).
apply powerRZSumRle; auto with zarith.
apply Rle_powerRZ; auto with real zarith.
apply Rle_IZR; omega.
apply Rplus_le_compat; auto with real.
apply Rle_trans with (-(-(x2+y)))%R;[apply Ropp_le_contravar|right; ring].
apply Rle_trans with (Rabs (-(x2+y)));[apply RRle_abs|rewrite Rabs_Ropp; auto with real].
fold FtoRradix; rewrite <- H2; apply Rle_powerRZ; auto with zarith real.
apply Rle_IZR; omega.
fold FtoRradix; replace (x1+x2+y-x1)%R with (x2+y)%R; auto with real; ring.
right; elim Evenradix; intros n Hn.
exists (Float (-n) (Fexp x1-2)); split;[idtac|simpl; auto with zarith].
apply trans_eq with (-(Rabs x2))%R;[rewrite H|idtac].
unfold FtoRradix, FtoR; simpl; unfold Zminus; rewrite powerRZ_add; auto with real zarith.
 rewrite powerRZ_add; try apply IZR_neq; auto with real zarith;rewrite Ropp_Ropp_IZR.
simpl; ring_simplify (radix*1)%R; rewrite Hn; rewrite mult_IZR; simpl; field; auto with real.
apply IZR_neq; auto with zarith.
apply IZR_neq; auto with zarith.
rewrite Rabs_left; auto with real.
assert (Rabs x2 <= powerRZ radix (Fexp x1)/2)%R.
apply Rmult_le_reg_l with (INR 2); auto with real zarith.
apply Rle_trans with (powerRZ radix (Fexp x1));[idtac|right; simpl; field; auto with real].
replace (FtoRradix x2) with ((x1+x2) -x1)%R;[idtac|ring].
apply Rle_trans with (Fulp bo radix p x1).
unfold FtoRradix; apply ClosestUlp; auto with zarith.
unfold Fulp; rewrite FcanonicFnormalizeEq; auto with zarith real; left; auto.
case H; clear H; intros H.
assert (Rabs (x2 + y) < powerRZ radix (Fexp x1) / 2)%R.
apply Rle_lt_trans with (Rabs x2+Rabs y)%R;[apply Rabs_triang|idtac].
apply Rplus_lt_reg_l with (-Rabs y)%R.
ring_simplify (- Rabs y + (Rabs x2 + Rabs y))%R.
elim Evenradix; intros n Hn.
assert (powerRZ radix (Fexp x1) / 2 = Float n (Fexp x1 -1))%R.
unfold FtoRradix, FtoR; simpl; unfold Zminus;
  rewrite powerRZ_add; auto with real zarith; simpl.
repeat (rewrite Hn; rewrite mult_IZR); simpl; field; auto with real zarith.
apply IZR_neq; auto with zarith.
apply IZR_neq; auto with zarith.
apply Rle_lt_trans with (Rabs (Float n (Fexp x1 - 1)) - powerRZ radix (LSB radix x2))%R.
apply LSB_Pred; auto.
rewrite (Rabs_right (Float n (Fexp x1 - 1))); auto with real.
rewrite <- H0; auto.
apply Rle_ge; unfold FtoRradix; apply LeFnumZERO; simpl; auto with real zarith.
apply Z.le_trans with (MSB radix x2).
apply LSB_le_MSB; auto with zarith.
contradict x2NonZero.
unfold FtoRradix; apply is_Fzero_rep1; auto.
apply Z.le_trans with (Fexp x1-1)%Z.
2: apply Z.le_trans with (Fexp ((Float n (Fexp x1 - 1)))); auto with zarith.
2: apply Fexp_le_LSB.
assert (MSB radix x2 < Fexp x1)%Z; auto with zarith.
apply Zlt_powerRZ with radix; auto with real.
apply Rle_IZR; omega.
apply Rle_lt_trans with (FtoR radix (Float (S 0) (MSB radix x2)));
       [right; unfold FtoR; simpl; ring|idtac].
apply Rle_lt_trans with (FtoR radix (Fabs x2)).
apply MSB_le_abs; auto.
contradict x2NonZero.
unfold FtoRradix; apply is_Fzero_rep1; auto.
rewrite Fabs_correct; auto; fold FtoRradix.
apply Rlt_le_trans with (1:=H).
unfold Rdiv; apply Rle_trans with ( (powerRZ radix (Fexp x1)*1))%R; auto with real zarith.
apply Rmult_le_compat_l; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
apply Rle_trans with (/1)%R; auto with real.
rewrite (Rabs_right (Float n (Fexp x1 - 1))); auto with real.
2: apply Rle_ge; unfold FtoRradix; apply LeFnumZERO; simpl; auto with real zarith.
rewrite <- H0; unfold Rminus; rewrite Rplus_comm.
apply Rplus_lt_compat_r; apply Ropp_lt_contravar.
unfold FtoRradix; rewrite <- Fabs_correct; auto.
apply Rlt_le_trans with (FtoR radix (Float (S 0) (Z.succ (MSB radix y))))%R.
apply abs_lt_MSB; auto.
unfold FtoR; simpl; ring_simplify (1 * powerRZ radix (Z.succ (MSB radix y)))%R.
apply Rle_powerRZ; auto with real zarith.
apply Rle_IZR; omega.
assert (powerRZ radix (Fexp x1 + p - 1) +powerRZ radix (Fexp x1) <= x1)%R.
apply Rle_trans with (((nNormMin radix p)+1)*(powerRZ radix (Fexp x1)))%R.
unfold nNormMin; rewrite Zpower_nat_Z_powerRZ.
replace (Fexp x1 + p - 1)%Z with (Fexp x1 + pred p)%Z.
rewrite powerRZ_add; try apply IZR_neq; auto with real zarith; right;ring.
rewrite inj_pred; unfold Z.pred; auto with zarith.
unfold FtoRradix, FtoR; apply Rmult_le_compat_r; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
replace 1%R with (IZR 1); auto with zarith; rewrite <- plus_IZR.
assert (nNormMin radix p + 1 <= Fnum x1)%Z; auto with real zarith.
assert (nNormMin radix p <= Fnum x1)%Z; auto with real zarith.
elim Nx1; intros.
apply Zmult_le_reg_r with radix; auto with zarith.
rewrite Zmult_comm; rewrite <- PosNormMin with radix bo p; auto with zarith.
apply Z.le_trans with (1:=H3).
rewrite Z.abs_eq; auto with zarith.
apply Z.mul_nonneg_nonneg; auto with zarith.
apply LeR0Fnum with radix; auto with real zarith.
apply Rle_IZR; auto.
left; unfold FtoRradix.
apply ImplyClosestStrict with bo p (x1+x2+y)%R (Fexp x1); auto with zarith.
elim x1Def; auto.
left; auto.
replace (x1+x2+y)%R with (x1+(x2+y))%R;[idtac|ring].
apply Rle_trans with ((powerRZ radix (Fexp x1 + p - 1) + powerRZ radix (Fexp x1))
  +-powerRZ radix (Fexp x1))%R;[right; ring|idtac].
apply Rplus_le_compat; auto.
rewrite <- (Ropp_involutive (x2+y)); apply Ropp_le_contravar.
apply Rle_trans with (Rabs (-(x2+y))); [apply RRle_abs|rewrite Rabs_Ropp].
apply Rle_trans with (powerRZ radix (Fexp x1) / 2)%R; auto with real.
apply Rle_trans with (powerRZ radix (Fexp x1) *1)%R; auto with real.
unfold Rdiv; apply Rmult_le_compat_l; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
apply Rle_trans with (/1)%R; auto with real.
fold FtoRradix; apply Rle_trans with (2:=H2); auto with real zarith.
apply Rle_trans with (powerRZ radix (Fexp x1 + p - 1)+0)%R; auto with real zarith.
apply Rplus_le_compat_l.
apply powerRZ_le, Rlt_IZR; omega.
fold FtoRradix; replace (x1+x2+y-x1)%R with (x2+y)%R; auto with real; ring.
right; elim Evenradix; intros n Hn.
case (Rle_or_lt 0%R x2); intros.
exists (Float n (Fexp x1-1)); split;[idtac|simpl; auto with zarith].
apply trans_eq with (Rabs x2);[rewrite H|idtac].
unfold FtoRradix, FtoR; simpl; unfold Zminus; rewrite powerRZ_add; auto with real zarith.
simpl; ring_simplify (radix*1)%R; rewrite Hn; rewrite mult_IZR; simpl; field; auto with real zarith.
apply IZR_neq; auto with zarith.
apply IZR_neq; auto with zarith.
rewrite Rabs_right; auto with real.
exists (Float (-n) (Fexp x1-1)); split;[idtac|simpl; auto with zarith].
apply trans_eq with (-(Rabs x2))%R;[rewrite H|idtac].
unfold FtoRradix, FtoR; simpl; unfold Zminus; rewrite powerRZ_add; auto with real zarith.
rewrite Ropp_Ropp_IZR; simpl; ring_simplify (radix*1)%R; rewrite Hn; rewrite mult_IZR; simpl; field; auto with real zarith.
apply IZR_neq; auto with zarith.
apply IZR_neq; auto with zarith.
rewrite Rabs_left; auto with real.
Qed.

End GenericC.

Section GenericD.
Variable bo : Fbound.
Variable radix : Z.
Variable p : nat.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.
Hypothesis radixMoreThanOne : (1 < radix)%Z.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ radixMoreThanOne).

Hypothesis pGivesBound : Zpos (vNum bo) = Zpower_nat radix p.
Hypothesis precisionGreaterThanOne : 1 < p.
Hypothesis Evenradix: (Even radix).

Variables x1 x2 y f:float.
Hypothesis x1Def: Closest bo radix (x1+x2)%R x1.
Hypothesis fDef : Closest bo radix (x1+x2+y)%R f.
Hypothesis yLe: (MSB radix y < LSB radix x2)%Z.
Hypothesis Nx1: Fnormal radix bo x1.
Hypothesis x2NonZero: ~(FtoRradix x2 =0)%R.
Hypothesis x1Exp: (-dExp bo < Fexp x1)%Z.

Lemma Midpoint_aux:
    (FtoRradix x1= f) \/ (exists v:float, (FtoRradix v=x2)%R /\ (Fexp x1 -2 <= Fexp v)%Z).
case (Rle_or_lt 0 x1); intros H.
case H; clear H; intros H.
unfold FtoRradix; apply Midpoint_aux_aux with bo p y; auto.
absurd (FtoRradix x1 =0)%R; auto with real.
assert (~ is_Fzero x1).
apply FnormalNotZero with radix bo; auto.
unfold is_Fzero in H0.
unfold FtoRradix, FtoR; apply prod_neq_R0; auto with real zarith.
apply IZR_neq; auto.
apply Rgt_not_eq; apply powerRZ_lt, Rlt_IZR; omega.
elim Midpoint_aux_aux with bo radix p (Fopp x1) (Fopp x2) (Fopp y) (Fopp f); auto.
repeat rewrite Fopp_correct; fold FtoRradix; intros; left; auto with real.
apply Rmult_eq_reg_l with (-1)%R; auto with real.
apply trans_eq with (-x1)%R; [ring| apply trans_eq with (1:=H0);ring].
intros T; elim T; intros v T'; elim T'; intros; clear T T'.
right; exists (Fopp v); split.
unfold FtoRradix; rewrite Fopp_correct; rewrite H0; rewrite Fopp_correct; ring.
simpl; apply Z.le_trans with (2:=H1); simpl; auto with zarith.
replace (FtoR radix (Fopp x1) + FtoR radix (Fopp x2))%R with (-(x1+x2))%R.
apply ClosestOpp; auto.
repeat rewrite Fopp_correct; unfold FtoRradix; ring.
replace (FtoR radix (Fopp x1) + FtoR radix (Fopp x2)+ FtoR radix (Fopp y))%R with (-(x1+x2+y))%R.
apply ClosestOpp; auto.
repeat rewrite Fopp_correct; unfold FtoRradix; ring.
rewrite <- MSB_opp; rewrite <- LSB_opp; auto.
apply FnormalFop; auto.
rewrite Fopp_correct; auto with real.
rewrite Fopp_correct; auto with real.
Qed.

End GenericD.

Section Be2Zero.
Variable bo : Fbound.
Variable radix : Z.
Variable p : nat.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.
Hypothesis radixMoreThanOne : (1 < radix)%Z.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ radixMoreThanOne).

Hypothesis pGivesBound : Zpos (vNum bo) = Zpower_nat radix p.
Hypothesis precisionGreaterThanOne : 3 <= p.
Hypothesis Evenradix: (Even radix).

Theorem TwoSumProp: forall (a b x y:float),
  (Fbounded bo a) ->
  (Closest bo radix (a+b)%R x) -> (FtoRradix y=a+b-x)%R
  -> (Rabs y <= Rabs b)%R.
intros.
elim H0; fold FtoRradix; intros.
rewrite <- (Rabs_Ropp y); rewrite <- (Rabs_Ropp b).
replace (-y)%R with (x-(a+b))%R;[idtac|rewrite H1; ring].
replace (-b)%R with (a-(a+b))%R;[idtac|ring].
apply H3; auto.
Qed.

Variable a x y r1 u1 u2 al1 al2 be1 be2 gat ga :float.

Hypothesis Fa : Fbounded bo a.
Hypothesis Fx : Fbounded bo x.
Hypothesis Fy : Fbounded bo y.

Hypothesis Nbe1: Fnormal radix bo be1.
Hypothesis Nr1 : Fnormal radix bo r1.
Hypothesis Cal1: Fcanonic radix bo al1.
Hypothesis Cu1 : Fcanonic radix bo u1.
Hypothesis Exp1: (- dExp bo < Fexp al1)%Z.
Hypothesis Exp2: (- dExp bo < Fexp u1)%Z.

Hypothesis r1Def: (Closest bo radix (a*x+y)%R r1).
Hypothesis u1Def: (Closest bo radix (a*x)%R u1).
Hypothesis u2Def: (FtoRradix u2=a*x-u1)%R.
Hypothesis al1Def:(Closest bo radix (y+u2)%R al1).
Hypothesis al2Def:(FtoRradix al2=y+u2-al1)%R.
Hypothesis be1Def:(Closest bo radix (u1+al1)%R be1).
Hypothesis be2Def:(FtoRradix be2=u1+al1-be1)%R.
Hypothesis gatDef:(Closest bo radix (be1-r1)%R gat).
Hypothesis gaDef: (Closest bo radix (gat+be2)%R ga).

Lemma gatCorrect: exists v:float, (FtoRradix v=be1-r1)%R /\ (Fbounded bo v)
                     /\ (Fexp v=Z.min (Fexp be1) (Fexp r1))%Z.
unfold FtoRradix; apply Subexact with p u1 al1 al2; auto.
apply Rle_trans with (Rabs u2).
apply TwoSumProp with y al1; auto.
apply Rmult_le_reg_l with (INR 2); auto with real zarith.
apply Rle_trans with (Fulp bo radix p u1);[idtac|simpl; right; field; auto with real].
rewrite u2Def; unfold FtoRradix; apply ClosestUlp; auto with zarith.
apply Rmult_le_reg_l with (INR 2); auto with real zarith.
apply Rle_trans with (Fulp bo radix p al1);[idtac|simpl; right; field; auto with real].
fold FtoRradix; rewrite al2Def; unfold FtoRradix; apply ClosestUlp; auto with zarith.
fold FtoRradix; replace (u1+al1+al2)%R with (a*x+y)%R; auto.
rewrite al2Def; rewrite u2Def; ring.
case (Rle_or_lt 0 (a*x+y))%R; intros I1.
fold FtoRradix; split; intros I2.
unfold FtoRradix; apply RleRoundedR0 with bo p (Closest bo radix) (u1+al1)%R; auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
apply Rplus_le_reg_l with (-u1)%R.
ring_simplify.
unfold FtoRradix; rewrite <- Fopp_correct.
apply RleBoundRoundl with bo p (Closest bo radix) (y + u2)%R; auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
apply oppBounded; elim u1Def; auto.
rewrite Fopp_correct; fold FtoRradix; apply Rplus_le_reg_l with (FtoRradix u1).
ring_simplify (u1+-u1)%R; apply Rle_trans with (1:=I1).
right; rewrite u2Def; ring.
absurd (FtoRradix r1 =0)%R.
assert (~ is_Fzero r1).
apply FnormalNotZero with radix bo; auto.
unfold is_Fzero in H.
unfold FtoRradix, FtoR; apply prod_neq_R0; auto with real zarith.
apply IZR_neq; auto.
apply Rgt_not_eq; apply powerRZ_lt, Rlt_IZR; omega.
assert (I3: (0<= r1)%R); auto with real.
unfold FtoRradix; apply RleRoundedR0 with bo p (Closest bo radix) (a*x+y)%R; auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
fold FtoRradix; split; intros I2.
absurd (FtoRradix r1 =0)%R.
assert (~ is_Fzero r1).
apply FnormalNotZero with radix bo; auto.
unfold is_Fzero in H.
unfold FtoRradix, FtoR; apply prod_neq_R0; auto with real zarith.
apply IZR_neq; auto.
apply Rgt_not_eq; apply powerRZ_lt, Rlt_IZR; omega.
assert (I3: (r1 <= 0)%R); auto with real.
unfold FtoRradix; apply RleRoundedLessR0 with bo p (Closest bo radix) (a*x+y)%R;
   auto with zarith real.
apply ClosestRoundedModeP with p; auto with zarith.
unfold FtoRradix; apply RleRoundedLessR0 with bo p (Closest bo radix) (u1+al1)%R;
   auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
apply Rplus_le_reg_l with (-u1)%R.
ring_simplify.
unfold FtoRradix; rewrite <- Fopp_correct.
apply RleBoundRoundr with bo p (Closest bo radix) (y + u2)%R; auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
apply oppBounded; elim u1Def; auto.
rewrite Fopp_correct; fold FtoRradix; apply Rplus_le_reg_l with (FtoRradix u1).
ring_simplify (u1+-u1)%R; apply Rle_trans with (a*x+y)%R; auto with real.
right; rewrite u2Def; ring.
Qed.

Hypothesis Be2Zero: (FtoRradix be2=0)%R.

Theorem FmaErr_aux1: (a*x+y=r1+ga+al2)%R.
generalize gatCorrect; intros H.
replace (FtoRradix ga) with (FtoRradix gat).
replace (FtoRradix gat) with (be1-r1)%R.
rewrite al2Def; rewrite u2Def.
apply trans_eq with (a * x + y-0)%R;[ring|rewrite <- Be2Zero].
rewrite be2Def; ring.
elim H; intros v H'; elim H'; intros H1 H''; elim H''; intros H2 H3; rewrite <- H1.
unfold FtoRradix.
apply RoundedModeProjectorIdemEq with (P:=(Closest bo radix)) (3:=pGivesBound);
  auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
fold FtoRradix; rewrite H1; auto.
unfold FtoRradix.
apply RoundedModeProjectorIdemEq with (P:=(Closest bo radix)) (3:=pGivesBound);
  auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
elim gatDef; auto.
fold FtoRradix; replace (FtoRradix gat) with (gat+be2)%R; auto with real.
rewrite Be2Zero; ring.
Qed.

End Be2Zero.

Section Be2NonZero.
Variable bo : Fbound.
Variable radix : Z.
Variable p : nat.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.
Hypothesis radixMoreThanOne : (1 < radix)%Z.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ radixMoreThanOne).

Hypothesis pGivesBound : Zpos (vNum bo) = Zpower_nat radix p.
Hypothesis precisionGreaterThanOne : 3 <= p.
Hypothesis Evenradix: (Even radix).

Variable P: R -> float -> Prop.
Hypothesis P1: forall (r:R) (f:float), (P r f) -> (Closest bo radix r f).
Hypothesis P2: forall (r1 r2:R) (f1 f2:float),
     (P r1 f1) -> (P r2 f2) -> (r1=r2)%R -> (FtoRradix f1=f2)%R.

Variable a x y r1 u1 u2 al1 al2 be1 be2 gat ga :float.

Hypothesis Fa : Fbounded bo a.
Hypothesis Fx : Fbounded bo x.
Hypothesis Fy : Fbounded bo y.

Hypothesis Nbe1: Fnormal radix bo be1.
Hypothesis Nr1 : Fnormal radix bo r1.
Hypothesis Cal1: Fcanonic radix bo al1.
Hypothesis Cu1 : Fcanonic radix bo u1.
Hypothesis Exp1: (- dExp bo < Fexp al1)%Z.
Hypothesis Exp2: (- dExp bo < Fexp u1)%Z.
Hypothesis Exp3: (- dExp bo+1 < Fexp be1)%Z.

Hypothesis r1Def: (Closest bo radix (a*x+y)%R r1).
Hypothesis u1Def: (Closest bo radix (a*x)%R u1).
Hypothesis u2Def: (FtoRradix u2=a*x-u1)%R.
Hypothesis al1Def:(Closest bo radix (y+u2)%R al1).
Hypothesis al2Def:(FtoRradix al2=y+u2-al1)%R.
Hypothesis be1Def:(Closest bo radix (u1+al1)%R be1).
Hypothesis be2Def:(FtoRradix be2=u1+al1-be1)%R.
Hypothesis gatDef:(Closest bo radix (be1-r1)%R gat).
Hypothesis gaDef: (Closest bo radix (gat+be2)%R ga).
Hypothesis be2Bounded: Fbounded bo be2.

Hypothesis r1DefE: (P (a*x+y)%R r1).
Hypothesis be1DefE:(P (u1+al1)%R be1).

Lemma Expr1 : (Fexp r1 <= Fexp be1+1)%Z.
assert (radix*be1=(Float (Fnum be1) (Fexp be1+1)))%R.
unfold FtoRradix, FtoR; simpl; rewrite powerRZ_add; try apply IZR_neq;
   auto with real zarith; simpl; ring.
apply Z.le_trans with (Fexp (Float (Fnum be1) (Fexp be1+1))); auto with zarith.
apply Fcanonic_Rle_Zle with radix bo p; auto with zarith.
left; auto.
elim Nbe1; intros J1 J2; elim J1; intros J3 J4.
left; split;[split|idtac]; simpl; auto with zarith.
fold FtoRradix; rewrite <- H.
rewrite Rabs_mult; rewrite (Rabs_right radix).
2: apply Rle_ge, Rle_IZR; auto with real zarith.
apply Rle_trans with (2*Rabs be1)%R.
2: apply Rmult_le_compat_r; auto with real.
2: apply Rabs_pos.
2: apply Rle_IZR; auto with real zarith.
unfold FtoRradix; apply yLe2x with bo p u1 al1 al2; auto with real.
apply Rle_trans with (Rabs u2).
unfold FtoRradix; apply TwoSumProp with bo y al1; auto.
apply Rmult_le_reg_l with (INR 2); auto with real zarith.
apply Rle_trans with (Fulp bo radix p u1);[idtac|simpl; right; field; auto with real].
rewrite u2Def; unfold FtoRradix; apply ClosestUlp; auto with zarith.
apply Rmult_le_reg_l with (INR 2); auto with real zarith.
apply Rle_trans with (Fulp bo radix p al1);[idtac|simpl; right; field; auto with real].
fold FtoRradix; rewrite al2Def; unfold FtoRradix; apply ClosestUlp; auto with zarith.
fold FtoRradix; replace (u1+al1+al2)%R with (a*x+y)%R; auto.
rewrite al2Def; rewrite u2Def; ring.
assert (~ is_Fzero be1).
apply FnormalNotZero with radix bo; auto.
unfold is_Fzero in H0.
unfold FtoR; apply prod_neq_R0; auto with real zarith.
apply IZR_neq; auto.
apply Rgt_not_eq, powerRZ_lt, Rlt_IZR; omega.
Qed.

Lemma Expbe1: (Fexp be1 <= Fexp r1+1)%Z.
assert (radix*r1=(Float (Fnum r1) (Fexp r1+1)))%R.
unfold FtoRradix, FtoR; simpl; rewrite powerRZ_add; try apply IZR_neq;
   auto with real zarith; simpl; ring.
apply Z.le_trans with (Fexp (Float (Fnum r1) (Fexp r1+1))); auto with zarith.
apply Fcanonic_Rle_Zle with radix bo p; auto with zarith.
left; auto.
elim Nr1; intros J1 J2; elim J1; intros J3 J4.
left; split;[split|idtac]; simpl; auto with zarith.
fold FtoRradix; rewrite <- H.
rewrite Rabs_mult; rewrite (Rabs_right radix).
2: apply Rle_ge, Rle_IZR; auto with real zarith.
apply Rle_trans with (2*Rabs r1)%R.
2: apply Rmult_le_compat_r; auto with real.
2: apply Rabs_pos.
2: apply Rle_IZR; auto with real zarith.
unfold FtoRradix; apply xLe2y with bo p u1 al1 al2; auto with real.
apply Rle_trans with (Rabs u2).
unfold FtoRradix; apply TwoSumProp with bo y al1; auto.
apply Rmult_le_reg_l with (INR 2); auto with real zarith.
apply Rle_trans with (Fulp bo radix p u1);[idtac|simpl; right; field; auto with real].
rewrite u2Def; unfold FtoRradix; apply ClosestUlp; auto with zarith.
apply Rmult_le_reg_l with (INR 2); auto with real zarith.
apply Rle_trans with (Fulp bo radix p al1);[idtac|simpl; right; field; auto with real].
fold FtoRradix; rewrite al2Def; unfold FtoRradix; apply ClosestUlp; auto with zarith.
fold FtoRradix; replace (u1+al1+al2)%R with (a*x+y)%R; auto.
rewrite al2Def; rewrite u2Def; ring.
Qed.

Lemma Zmin_Zlt : forall z1 z2 z3 : Z,
       (z1 < z2)%Z -> (z1 < z3)%Z -> (z1 < Z.min z2 z3)%Z.
intros; unfold Z.min.
case (z2 ?= z3)%Z; auto.
Qed.

Hypothesis Be2NonZero: ~(FtoRradix be2=0)%R.

Lemma be2MuchSmaller:
  ~(FtoRradix al2=0)%R -> ~(FtoRradix u2=0)%R ->
  (MSB radix al2 < LSB radix be2)%Z.
intros.
assert (FtoRradix be2 = (Fminus radix (Fplus radix u1 al1) be1))%R.
rewrite be2Def; unfold FtoRradix; rewrite Fminus_correct; auto;
   rewrite Fplus_correct; auto; ring.
rewrite LSB_comp with radix be2 (Fminus radix (Fplus radix u1 al1) be1); auto with zarith.
2: contradict Be2NonZero.
2: unfold FtoRradix; apply is_Fzero_rep1; auto.
apply Z.lt_le_trans with (Z.min (LSB radix (Fplus radix u1 al1)) (LSB radix be1)).
2: apply LSBMinus; auto.
2: contradict Be2NonZero.
2: rewrite H1; unfold FtoRradix; apply is_Fzero_rep1; auto.
apply Zmin_Zlt.
apply Z.lt_le_trans with (Z.min (LSB radix u1 ) (LSB radix al1)).
2: apply LSBPlus; auto.
apply Zmin_Zlt.
apply Z.le_lt_trans with (MSB radix u2).
apply MSB_monotone; auto.
contradict H; unfold FtoRradix; apply is_Fzero_rep1; auto.
contradict H0; unfold FtoRradix; apply is_Fzero_rep1; auto.
repeat rewrite Fabs_correct; auto.
apply TwoSumProp with bo y al1; auto.
rewrite MSB_comp with radix u2 (Fminus radix (Fmult a x) u1); auto with zarith.
apply MSBroundLSB with bo p (Closest bo radix); auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
rewrite Fmult_correct; auto with real.
contradict H0.
apply trans_eq with (FtoRradix (Fminus radix (Fmult a x) u1)).
rewrite u2Def; unfold FtoRradix; rewrite Fminus_correct; auto; rewrite Fmult_correct; auto; ring.
unfold FtoRradix; apply is_Fzero_rep1; auto.
contradict H0; unfold FtoRradix; apply is_Fzero_rep1; auto.
fold FtoRradix; rewrite u2Def; unfold FtoRradix; rewrite Fminus_correct; auto;
  rewrite Fmult_correct; auto; ring.
rewrite MSB_comp with radix al2 (Fminus radix (Fplus radix y u2) al1); auto with zarith.
apply MSBroundLSB with bo p (Closest bo radix); auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
rewrite Fplus_correct; auto with real.
contradict H.
apply trans_eq with (FtoRradix (Fminus radix (Fplus radix y u2) al1)).
rewrite al2Def; unfold FtoRradix; rewrite Fminus_correct; auto; rewrite Fplus_correct; auto; ring.
unfold FtoRradix; apply is_Fzero_rep1; auto.
contradict H; unfold FtoRradix; apply is_Fzero_rep1; auto.
fold FtoRradix; rewrite al2Def; unfold FtoRradix; rewrite Fminus_correct; auto;
  rewrite Fplus_correct; auto; ring.
cut (~(u1+al1=0)%R).
intros I; contradict I.
unfold FtoRradix; rewrite <- Fplus_correct; auto; apply is_Fzero_rep1; auto.
contradict Be2NonZero.
rewrite be2Def; rewrite Be2NonZero.
assert (FtoRradix be1=0)%R; auto with real.
rewrite <- FzeroisReallyZero with radix (-(dExp bo))%Z.
unfold FtoRradix; apply sym_eq.
apply RoundedModeProjectorIdemEq with bo p (Closest bo radix); auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
apply FboundedFzero.
rewrite FzeroisReallyZero; rewrite <- Be2NonZero; auto.
apply Zlt_powerRZ with radix; auto with real zarith.
apply Rle_IZR; omega.
apply Rle_lt_trans with (FtoR radix (Float (S 0) (MSB radix al2)));
   [unfold FtoR; simpl; right; ring|idtac].
apply Rle_lt_trans with (FtoR radix (Fabs al2));[apply MSB_le_abs; auto with zarith|idtac].
contradict H; unfold FtoRradix; apply is_Fzero_rep1; auto.
rewrite Fabs_correct; auto; fold FtoRradix.
apply Rlt_le_trans with (powerRZ radix (Fexp be1)).
2: apply Rle_powerRZ; auto with real zarith.
2: apply Rle_IZR; omega.
2: apply Fexp_le_LSB; auto.
apply Rlt_le_trans with (powerRZ radix (Z.min (Fexp u1) (Fexp al1))).
cut (Rabs al2 < powerRZ radix (Fexp u1))%R;[intros I1|idtac].
cut (Rabs al2 < powerRZ radix (Fexp al1))%R;[intros I2|idtac].
unfold Z.min; case (Fexp u1 ?= Fexp al1)%Z; auto with real zarith.
rewrite al2Def; apply Rlt_le_trans with (Fulp bo radix p al1).
unfold FtoRradix; apply RoundedModeUlp with (Closest bo radix); auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
unfold Fulp; rewrite FcanonicFnormalizeEq; auto with zarith real; left; auto.
apply Rle_lt_trans with (Rabs u2).
unfold FtoRradix; apply TwoSumProp with bo y al1; auto.
rewrite u2Def; apply Rlt_le_trans with (Fulp bo radix p u1).
unfold FtoRradix; apply RoundedModeUlp with (Closest bo radix); auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
unfold Fulp; rewrite FcanonicFnormalizeEq; auto with zarith real; left; auto.
apply Rle_powerRZ; try apply Rle_IZR; auto with real zarith.
apply Zlt_le_weak.
case (Zle_or_lt (Fexp be1) (Z.min (Fexp u1) (Fexp al1))); auto.
intros; contradict Be2NonZero.
rewrite be2Def.
assert (FtoRradix be1=u1+al1)%R; auto with real.
unfold FtoRradix; apply plusExact1 with bo p; auto with zarith.
elim u1Def; auto.
elim al1Def; auto.
Qed.

Lemma gaCorrect: exists v:float, (FtoRradix v=be1-r1+be2)%R /\ (Fbounded bo v).
elim gatCorrect with bo radix p a x y r1 u1 u2 al1 al2 be1; auto.
intros v T; elim T; intros H1 T'; elim T'; intros H2 H3; clear T T'.
case (Req_dec al2 0); intros Z1.
exists be2; split; auto.
cut (FtoRradix be1=r1)%R.
intros T; rewrite T; ring.
apply P2 with (u1+al1)%R (a*x+y)%R; auto.
apply trans_eq with (u1+al1+al2)%R;[rewrite Z1; ring|rewrite al2Def; rewrite u2Def; ring].
case (Req_dec u2 0); intros Z2.
contradict Z1.
rewrite al2Def; rewrite Z2.
cut (FtoRradix y=al1);[intros I; rewrite I; ring|idtac].
apply RoundedModeProjectorIdemEq with (P:=(Closest bo radix)) (3:=pGivesBound); auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
replace (FtoR radix y) with (y+u2)%R; auto with real.
rewrite Z2; unfold FtoRradix; ring.
elim Midpoint_aux with bo radix p be1 be2 al2 r1; auto with zarith.
3: fold FtoRradix; replace (be1+be2)%R with (u1+al1)%R; auto.
3: rewrite be2Def; ring.
3: fold FtoRradix; replace (be1+be2+al2)%R with (a*x+y)%R; auto.
3: rewrite al2Def; rewrite u2Def; rewrite be2Def; ring.
3: apply be2MuchSmaller; auto.
fold FtoRradix; intros; exists be2; split; auto.
rewrite H; ring.
intros T'; elim T'; intros v' T; elim T; intros; clear T T'.
elim BoundedL with bo radix p (be1 - r1 + be2)%R (Fplus radix v v') (Fexp be1 - 2)%Z; auto with zarith.
intros v'' T; elim T; intros H4 T'; elim T'; intros; clear T T'.
exists v''; split; auto.
simpl; apply Zmin_Zle; auto.
rewrite H3; apply Zmin_Zle; auto with zarith.
generalize Expbe1; auto with zarith.
auto with zarith.
unfold FtoRradix; rewrite Fplus_correct; auto; rewrite H; rewrite H1; ring.
replace (be1-r1+be2)%R with (((a*x+y)-r1)+-al2)%R.
2: rewrite al2Def; rewrite u2Def; rewrite be2Def; ring.
apply Rle_lt_trans with (Rabs (a * x + y - r1) + Rabs(- al2))%R;
   [apply Rabs_triang|rewrite Rabs_Ropp].
apply Rle_lt_trans with ((powerRZ radix (Fexp be1+1))/2+(powerRZ radix (Fexp be1))/2)%R;
  [apply Rplus_le_compat|idtac].
apply Rmult_le_reg_l with (INR 2); auto with real zarith.
apply Rle_trans with (Fulp bo radix p r1).
unfold FtoRradix; apply ClosestUlp; auto with zarith.
unfold Fulp; rewrite FcanonicFnormalizeEq; auto with zarith;[idtac|left; auto].
apply Rle_trans with (powerRZ radix (Fexp be1 + 1));
  [apply Rle_powerRZ; try apply Rle_IZR; auto with real zarith|right; simpl; field; auto with real].
apply Expr1.
apply Rle_trans with (Rabs be2).
assert (MSB radix al2 < LSB radix be2)%Z.
apply be2MuchSmaller; auto.
unfold FtoRradix; repeat rewrite <- Fabs_correct; auto.
apply Rle_trans with (FtoR radix (Float (S 0) (Z.succ (MSB radix al2))))%R.
apply Rlt_le; apply abs_lt_MSB; auto.
apply Rle_trans with (FtoR radix (Float (S 0) (LSB radix be2))).
unfold FtoR; simpl; apply Rmult_le_compat_l; auto with real.
apply Rle_powerRZ; try apply Rle_IZR; auto with real zarith.
apply LSB_le_abs; auto.
contradict Be2NonZero; unfold FtoRradix; apply is_Fzero_rep1; auto.
apply Rmult_le_reg_l with (INR 2); auto with real zarith.
apply Rle_trans with (Fulp bo radix p be1).
rewrite be2Def; unfold FtoRradix; apply ClosestUlp; auto with zarith.
unfold Fulp; rewrite FcanonicFnormalizeEq; auto with zarith;[idtac|left; auto].
right; simpl; field; auto with real.
apply Rlt_le_trans with (powerRZ radix (Fexp be1 + 1));
  [idtac|apply Rle_powerRZ; auto with real zarith].
apply Rlt_le_trans with (powerRZ radix (Fexp be1 + 1) / 2 + powerRZ radix (Fexp be1+1) / 2)%R;
  [apply Rplus_lt_compat_l|right; field; auto with real].
unfold Rdiv; apply Rmult_lt_compat_r; auto with real.
apply Rlt_powerRZ; try apply Rlt_IZR; auto with real zarith.
apply Rle_IZR; auto with zarith.
Qed.

Theorem FmaErr_aux2: (a*x+y=r1+ga+al2)%R.
elim gatCorrect with bo radix p a x y r1 u1 u2 al1 al2 be1; auto.
intros v1 T; elim T; intros H1 T'; elim T'; intros H2 H3; clear T T'.
elim gaCorrect; intros v2 T; elim T; intros H4 H5; clear T.
replace (FtoRradix ga) with (FtoRradix v2).
rewrite H4; rewrite be2Def; rewrite al2Def; rewrite u2Def; ring.
unfold FtoRradix; apply RoundedModeProjectorIdemEq with
   (P:=(Closest bo radix)) (3:=pGivesBound); auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
replace (FtoR radix v2) with (gat+be2)%R; auto.
fold FtoRradix; rewrite H4; assert (FtoRradix gat=be1-r1)%R; auto with real.
unfold FtoRradix; rewrite <- H1.
apply sym_eq; apply RoundedModeProjectorIdemEq with
   (P:=(Closest bo radix)) (3:=pGivesBound); auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
rewrite H1; auto with real.
Qed.

End Be2NonZero.

Section Final.

Variable bo : Fbound.
Variable radix : Z.
Variable p : nat.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.
Hypothesis radixMoreThanOne : (1 < radix)%Z.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ radixMoreThanOne).

Hypothesis pGivesBound : Zpos (vNum bo) = Zpower_nat radix p.
Hypothesis precisionGreaterThanOne : 3 <= p.
Hypothesis Evenradix: (Even radix).

Variable P: R -> float -> Prop.
Hypothesis P1: forall (r:R) (f:float), (P r f) -> (Closest bo radix r f).
Hypothesis P2: forall (r1 r2:R) (f1 f2:float),
     (P r1 f1) -> (P r2 f2) -> (r1=r2)%R -> (FtoRradix f1=f2)%R.

Variable a x y r1 u1 u2 al1 al2 be1 be2 gat ga :float.

Hypothesis Fa : Fbounded bo a.
Hypothesis Fx : Fbounded bo x.
Hypothesis Fy : Fbounded bo y.

Hypothesis Nbe1: Fnormal radix bo be1.
Hypothesis Nr1 : Fnormal radix bo r1.
Hypothesis Cal1: Fcanonic radix bo al1.
Hypothesis Cu1 : Fcanonic radix bo u1.
Hypothesis Exp1: (- dExp bo < Fexp al1)%Z.
Hypothesis Exp2: (- dExp bo < Fexp u1)%Z.
Hypothesis Exp3: (- dExp bo+1 < Fexp be1)%Z.

Hypothesis r1Def: (Closest bo radix (a*x+y)%R r1).
Hypothesis u1Def: (Closest bo radix (a*x)%R u1).
Hypothesis u2Def: (FtoRradix u2=a*x-u1)%R.
Hypothesis al1Def:(Closest bo radix (y+u2)%R al1).
Hypothesis al2Def:(FtoRradix al2=y+u2-al1)%R.
Hypothesis be1Def:(Closest bo radix (u1+al1)%R be1).
Hypothesis be2Def:(FtoRradix be2=u1+al1-be1)%R.
Hypothesis gatDef:(Closest bo radix (be1-r1)%R gat).
Hypothesis gaDef: (Closest bo radix (gat+be2)%R ga).
Hypothesis be2Bounded: Fbounded bo be2.

Hypothesis r1DefE: (P (a*x+y)%R r1).
Hypothesis be1DefE:(P (u1+al1)%R be1).

Theorem FmaErr_aux: (a*x+y=r1+ga+al2)%R.
case (Req_dec be2 0); intros.
unfold FtoRradix; apply FmaErr_aux1 with bo p u1 u2 al1 be1 be2 gat; auto.
unfold FtoRradix; apply FmaErr_aux2 with bo p P u1 u2 al1 be1 be2 gat; auto.
Qed.
End Final.

Section Final2.

Variable bo : Fbound.
Variable radix : Z.
Variable p : nat.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.
Hypothesis radixMoreThanOne : (1 < radix)%Z.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ radixMoreThanOne).

Hypothesis pGivesBound : Zpos (vNum bo) = Zpower_nat radix p.
Hypothesis precisionGreaterThanOne : 3 <= p.
Hypothesis Evenradix: (Even radix).

Lemma ClosestZero1: forall (r:R) (f g:float), (Closest bo radix r f) -> (FtoRradix f=0)%R ->
   (r=g)%R -> (-dExp bo <= Fexp g)%Z -> (r=0)%R.
intros.
case (Req_dec r 0); auto; intros.
absurd (powerRZ radix (-(dExp bo)) <= Rabs f)%R; auto.
apply Rlt_not_le; rewrite H0; rewrite Rabs_R0; auto with real zarith.
apply powerRZ_lt, Rlt_IZR; auto with real zarith.
apply Rle_trans with (FtoRradix (Float 1 (-(dExp bo)))).
right; unfold FtoRradix, FtoR; simpl; ring.
unfold FtoRradix; apply RoundAbsMonotonel with bo p (Closest bo radix) r; auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
split; simpl; auto with zarith.
rewrite pGivesBound; apply Z.le_lt_trans with (Zpower_nat radix 0); auto with zarith.
apply Zpower_nat_monotone_lt; auto with zarith.
rewrite H1; unfold FtoRradix; rewrite <- Fabs_correct; auto.
unfold FtoR; simpl; apply Rmult_le_compat; auto with real zarith.
apply powerRZ_le, Rlt_IZR; auto with real zarith.
apply Rle_IZR.
case (Zle_lt_or_eq 0 (Z.abs (Fnum g))); auto with zarith; intros.
absurd (Rabs r=0)%R.
apply Rabs_no_R0; auto.
rewrite H1; unfold FtoRradix; rewrite <- Fabs_correct; auto; unfold FtoR; simpl.
rewrite <- H4; simpl; ring.
apply Rle_powerRZ; try apply Rle_IZR; auto with real zarith.
Qed.

Lemma ClosestZero2: forall (r:R) (x:float),
  (Closest bo radix r x) -> (r=0)%R -> (FtoRradix x=0)%R.
intros.
cut (0 <= FtoRradix x)%R;[intros |idtac].
cut (FtoRradix x <= 0)%R;[intros; auto with real |idtac].
unfold FtoRradix; apply RleRoundedLessR0 with bo p (Closest bo radix) r; auto with real.
apply ClosestRoundedModeP with p; auto with zarith.
unfold FtoRradix; apply RleRoundedR0 with bo p (Closest bo radix) r; auto with real.
apply ClosestRoundedModeP with p; auto with zarith.
Qed.

Lemma LeExpRound: forall f g:float, Closest bo radix f g
          -> exists g':float, Fbounded bo g' /\ FtoRradix g'=g /\ (Fexp f <= Fexp g')%Z.
intros.
case (Zle_or_lt (Fexp f) (Fexp g)); intros.
exists g; split; auto.
elim H; auto.
elim RoundedModeRep with bo radix p (Closest bo radix) f g; auto with zarith.
2: apply ClosestRoundedModeP with p; auto with zarith.
fold FtoRradix; intros n H1.
exists (Float n (Fexp f)).
elim H; intros H2 T; clear T.
split; split; simpl; auto with zarith real .
apply Z.le_lt_trans with (Z.abs (Fnum g)); auto with zarith; try apply H2.
apply Zle_Rle.
apply Rmult_le_reg_l with (powerRZ radix (Fexp g)); auto with real zarith.
apply powerRZ_lt, Rlt_IZR; auto with real zarith.
apply Rle_trans with (Rabs g);[idtac| unfold FtoRradix; rewrite <- Fabs_correct; auto;
     unfold FtoR; simpl; right; ring].
rewrite H1; unfold FtoRradix; rewrite <- Fabs_correct; auto; unfold FtoR; simpl.
rewrite Rmult_comm; apply Rmult_le_compat_l; auto with real zarith.
apply Rle_IZR; auto with zarith.
apply Rle_powerRZ; try apply Rle_IZR; auto with zarith.
apply Z.le_trans with (Fexp g); auto with zarith; apply H2.
Qed.

Lemma LeExpRound2: forall (n:Z) (f g:float), Closest bo radix f g
          -> (n <= Fexp f)%Z
          -> exists g':float, Fbounded bo g' /\ FtoRradix g'=g /\ (n <= Fexp g')%Z.
intros.
elim LeExpRound with f g; auto.
intros g' T; elim T; intros T1 T2; elim T2; intros T3 T4; clear T T2.
exists g'; split; auto; split; auto with zarith.
Qed.

Variable P: R -> float -> Prop.
Hypothesis P1: forall (r:R) (f:float), (P r f) -> (Closest bo radix r f).
Hypothesis P2: forall (r1 r2:R) (f1 f2:float),
     (P r1 f1) -> (P r2 f2) -> (r1=r2)%R -> (FtoRradix f1=f2)%R.

Variable a x y r1 u1 u2 al1 al2 be1 be2 gat ga :float.

Hypothesis Fa : Fbounded bo a.
Hypothesis Fx : Fbounded bo x.
Hypothesis Fy : Fbounded bo y.

Hypothesis Nbe1: Fcanonic radix bo be1.
Hypothesis Nr1 : Fcanonic radix bo r1.
Hypothesis Cal1: Fcanonic radix bo al1.
Hypothesis Cu1 : Fcanonic radix bo u1.
Hypothesis Exp1: (- dExp bo < Fexp al1)%Z \/ (FtoRradix al1=0)%R.
Hypothesis Exp2: (- dExp bo < Fexp u1)%Z \/ (FtoRradix u1=0)%R.
Hypothesis Exp3: (- dExp bo+1 < Fexp be1)%Z\/ (FtoRradix be1=0)%R.
Hypothesis Exp4: (Fnormal radix bo r1) \/ (FtoRradix r1=0)%R.
Hypothesis Exp5: (-dExp bo <= Fexp a+Fexp x)%Z.

Hypothesis u1Def: (Closest bo radix (a*x)%R u1).
Hypothesis u2Def: (FtoRradix u2=a*x-u1)%R.
Hypothesis al1Def:(Closest bo radix (y+u2)%R al1).
Hypothesis al2Def:(FtoRradix al2=y+u2-al1)%R.
Hypothesis be2Def:(FtoRradix be2=u1+al1-be1)%R.
Hypothesis gatDef:(Closest bo radix (be1-r1)%R gat).
Hypothesis gaDef: (Closest bo radix (gat+be2)%R ga).
Hypothesis r1DefE: (P (a*x+y)%R r1).
Hypothesis be1DefE:(P (u1+al1)%R be1).

Theorem FmaErr: (a*x+y=r1+ga+al2)%R.
case Exp1; intros I1.
case Exp2; intros I2.
case Exp3; intros I3.
case Exp4; intros I4.
elim errorBoundedPlus with bo radix p u1 al1 be1; auto with zarith.
2: elim u1Def; auto.
2: elim al1Def; auto.
fold FtoRradix; intros be2' T; elim T; intros J1 T'; elim T'; intros J2 J3; clear T T'.
unfold FtoRradix; apply FmaErr_aux with bo p P u1 u2 al1 be1 be2' gat; auto.
elim Nbe1; intros; auto.
elim H; intros T1 T2; elim T2; intros T3 T4; contradict T3; auto with zarith.
fold FtoRradix; rewrite J1; rewrite <- be2Def; auto.
assert (a*x+y=0)%R.
apply ClosestZero1 with r1 (Fplus radix (Fmult a x) y); auto.
unfold FtoRradix; rewrite Fplus_correct; auto; rewrite Fmult_correct; auto with real zarith.
simpl; apply Zmin_Zle; auto with zarith.
elim Fy; auto.
rewrite H.
assert (FtoRradix r1=0)%R.
apply ClosestZero2 with (a*x+y)%R; auto; elim r1DefE; auto.
assert (FtoRradix u1= (Fopp y))%R.
unfold FtoRradix; apply sym_eq;
   apply RoundedModeProjectorIdemEq with bo p (Closest bo radix); auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
apply oppBounded; auto.
replace (FtoR radix (Fopp y)) with (a*x)%R; auto.
rewrite Fopp_correct; fold FtoRradix; auto with real.
apply Rplus_eq_reg_l with y; rewrite Rplus_comm; rewrite H; ring.
assert (FtoRradix u2=0)%R.
rewrite u2Def; rewrite H1; unfold FtoRradix; rewrite Fopp_correct; fold FtoRradix.
rewrite <- H; ring.
assert (FtoRradix al1= y)%R.
unfold FtoRradix; apply sym_eq;
   apply RoundedModeProjectorIdemEq with bo p (Closest bo radix); auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
replace (FtoR radix y) with (y+u2)%R; auto with real.
rewrite H2; auto with real.
assert (FtoRradix al2=0)%R.
rewrite al2Def; rewrite H2; rewrite H3; ring.
assert (FtoRradix be1=0)%R.
apply ClosestZero2 with (u1+al1)%R; auto.
rewrite H3; rewrite H1; unfold FtoRradix; rewrite Fopp_correct; ring.
assert (FtoRradix be2=0)%R.
rewrite be2Def; rewrite H1; rewrite H3; rewrite H5; unfold FtoRradix; rewrite Fopp_correct; ring.
assert (FtoRradix gat=0)%R.
apply ClosestZero2 with (be1-r1)%R; auto.
rewrite H5; rewrite H0; ring.
assert (FtoRradix ga=0)%R.
apply ClosestZero2 with (gat+be2)%R; auto.
rewrite H7; rewrite H6; ring.
rewrite H0; rewrite H8; rewrite H4; ring.
assert (u1 + al1=0)%R.
apply ClosestZero1 with be1 (Fplus radix u1 al1); auto.
unfold FtoRradix; rewrite Fplus_correct; auto.
simpl; apply Zmin_Zle.
elim u1Def; intros T1 T2; elim T1; auto.
elim al1Def; intros T1 T2; elim T1; auto.
assert (FtoRradix be2=0)%R.
rewrite be2Def; rewrite I3; rewrite H; ring.
assert (FtoRradix gat= (Fopp r1))%R.
unfold FtoRradix; apply sym_eq;
   apply RoundedModeProjectorIdemEq with bo p (Closest bo radix); auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
apply oppBounded; auto.
assert (T:(Closest bo radix (a*x+y)%R r1)); auto; elim T; auto.
replace (FtoR radix (Fopp r1)) with (be1-r1)%R; auto.
rewrite I3; rewrite Fopp_correct; fold FtoRradix; auto with real.
assert (FtoRradix ga= gat)%R.
unfold FtoRradix; apply sym_eq;
   apply RoundedModeProjectorIdemEq with bo p (Closest bo radix); auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
elim gatDef; auto.
replace (FtoR radix gat) with (gat+be2)%R; auto.
rewrite H0; auto with real.
apply sym_eq; apply trans_eq with (FtoRradix al2).
rewrite H2; rewrite H1; unfold FtoRradix; rewrite Fopp_correct; ring.
rewrite al2Def; rewrite u2Def.
apply trans_eq with (y+a*x-(u1+al1))%R;[idtac|rewrite H]; ring.
assert (a*x=0)%R.
apply ClosestZero1 with u1 (Fmult a x); auto.
unfold FtoRradix; rewrite Fmult_correct; auto.
assert (FtoRradix r1=y).
unfold FtoRradix; apply sym_eq;
   apply RoundedModeProjectorIdemEq with bo p (Closest bo radix); auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
replace (FtoR radix y) with (a*x+y)%R;[idtac|rewrite H]; auto with real.
assert (FtoRradix u2=0)%R.
rewrite u2Def; rewrite I2; rewrite H; ring.
assert (FtoRradix al1=y).
unfold FtoRradix; apply sym_eq;
   apply RoundedModeProjectorIdemEq with bo p (Closest bo radix); auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
replace (FtoR radix y) with (y+u2)%R;[idtac|rewrite H1]; auto with real.
assert (FtoRradix al2=0)%R.
rewrite al2Def; rewrite H1; rewrite H2; ring.
assert (FtoRradix be1=y).
unfold FtoRradix; apply sym_eq;
   apply RoundedModeProjectorIdemEq with bo p (Closest bo radix); auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
replace (FtoR radix y) with (u1+al1)%R;[idtac|rewrite I2; rewrite H2]; auto with real.
assert (FtoRradix gat=0).
apply ClosestZero2 with (be1-r1)%R; auto.
rewrite H4; rewrite H0; ring.
assert (FtoRradix ga=0).
apply ClosestZero2 with (gat+be2)%R; auto.
rewrite be2Def;rewrite H5; rewrite I2; rewrite H2; rewrite H4; simpl; ring.
rewrite H; rewrite H0; rewrite H6; rewrite H3;simpl; ring.
assert (y + u2=0)%R.
elim errorBoundedMult with bo radix p (Closest bo radix) a x u1; auto with zarith.
fold FtoRradix; intros u2' T; elim T; intros H1 T'; elim T'; intros H2 H3; clear T T'.
apply ClosestZero1 with al1 (Fplus radix y u2'); auto.
unfold FtoRradix; rewrite Fplus_correct; auto; fold FtoRradix; rewrite H1; rewrite u2Def; ring.
simpl; apply Zmin_Zle; auto with zarith; apply Fy.
apply ClosestRoundedModeP with p; auto with zarith.
assert (FtoRradix be1=u1).
unfold FtoRradix; apply sym_eq;
   apply RoundedModeProjectorIdemEq with bo p (Closest bo radix); auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
elim u1Def; auto.
replace (FtoR radix u1) with (u1+al1)%R;[idtac|rewrite I1]; auto with real.
assert (FtoRradix r1=u1).
unfold FtoRradix; apply sym_eq;
   apply RoundedModeProjectorIdemEq with bo p (Closest bo radix); auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
elim u1Def; auto.
replace (FtoR radix u1) with (a*x+y)%R;[idtac|idtac]; auto with real.
fold FtoRradix; apply trans_eq with (u1+(y+u2))%R;[rewrite u2Def| rewrite H]; ring.
assert (FtoRradix gat=0)%R.
apply ClosestZero2 with (be1-r1)%R; auto; rewrite H0; rewrite H1; auto with real.
assert (FtoRradix ga=0)%R.
apply ClosestZero2 with (gat+be2)%R; auto.
rewrite H2; rewrite be2Def; rewrite I1; rewrite H0; ring.
rewrite H1; rewrite H3; rewrite al2Def; rewrite u2Def; rewrite I1; ring.
Qed.

Theorem Fma_FTS: (exists ga_e:float, exists al2_e:float,
                 (FtoRradix ga_e=ga)%R /\ (FtoRradix al2_e=al2)%R
                  /\ (Fbounded bo ga_e) /\ (Fbounded bo al2_e) /\
                   (Fexp al2_e <= Fexp ga_e)%Z).
elim errorBoundedMult with bo radix p (Closest bo radix) a x u1; auto with zarith.
2:apply ClosestRoundedModeP with p; auto with zarith.
fold FtoRradix; intros u2' T; elim T; intros E1 T'; elim T'; intros E2 E3; clear T T'.
elim errorBoundedPlus with bo radix p y u2' al1; auto with zarith.
2: fold FtoRradix; rewrite E1; rewrite <- u2Def; auto.
rewrite E3; intros al2' T; elim T; intros F1 T'; elim T'; intros F2 F3; clear T T'.
elim LeExpRound2 with (Z.min (Fexp y) (Fexp a+Fexp x)) (Fplus radix y u2') al1; auto.
2: unfold FtoRradix; rewrite Fplus_correct; auto; fold FtoRradix;
     rewrite E1; rewrite <- u2Def; auto.
2:simpl; rewrite E3; auto with zarith.
intros al1' T; elim T; intros U1 T'; elim T'; intros U2 U3; clear T T'.
elim LeExpRound2 with (Z.min (Fexp y) (Fexp a+Fexp x)) (Fmult a x) u1; auto with zarith.
2: unfold FtoRradix; rewrite Fmult_correct; auto.
intros u1' T; elim T; intros U4 T'; elim T'; intros U5 U6; clear T T'.
elim LeExpRound2 with (Z.min (Fexp y) (Fexp a+Fexp x)) (Fplus radix (Fmult a x) y) r1; auto with zarith.
2: unfold FtoRradix; rewrite Fplus_correct; auto; rewrite Fmult_correct; auto.
2: simpl; rewrite Zmin_sym; auto with zarith.
intros r1' T; elim T; intros U7 T'; elim T'; intros U8 U9; clear T T'.
elim LeExpRound2 with (Z.min (Fexp y) (Fexp a+Fexp x)) (Fplus radix u1' al1') be1; auto with zarith.
2: unfold FtoRradix; rewrite Fplus_correct; auto; fold FtoRradix; rewrite U5; rewrite U2; auto.
2: simpl; apply Zmin_Zle; auto with zarith.
intros be1' T; elim T; intros V1 T'; elim T'; intros V2 V3; clear T T'.
elim errorBoundedPlus with bo radix p u1' al1' be1; auto with zarith.
2: fold FtoRradix; rewrite U5; rewrite U2; auto.
fold FtoRradix; intros be2' T; elim T; intros V4 T'; elim T'; intros V5 V6; clear T T'.
elim LeExpRound2 with (Z.min (Fexp y) (Fexp a+Fexp x)) (Fminus radix be1' r1') gat; auto with zarith.
2: unfold FtoRradix; rewrite Fminus_correct; auto; fold FtoRradix; rewrite V2; rewrite U8; auto.
2: simpl; apply Zmin_Zle; auto with zarith.
intros gat' T; elim T; intros V7 T'; elim T'; intros V8 V9; clear T T'.
elim LeExpRound2 with (Z.min (Fexp y) (Fexp a+Fexp x)) (Fplus radix gat' be2') ga; auto with zarith.
2: unfold FtoRradix; rewrite Fplus_correct; auto; fold FtoRradix; rewrite V8;
   rewrite V4; rewrite U5; rewrite U2; rewrite <- be2Def; auto.
2: simpl; apply Zmin_Zle; auto with zarith.
intros ga' T; elim T; intros W1 T'; elim T'; intros W2 W3; clear T T'.
exists ga'; exists al2'; split; auto; split.
unfold FtoRradix; rewrite F1; fold FtoRradix.
rewrite E1; rewrite al2Def; rewrite u2Def; ring.
split; auto; split; auto with zarith.
Qed.

End Final2.

Section tBounded.
Variable bo : Fbound.
Variable radix : Z.
Variable p : nat.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.
Hypothesis radixMoreThanOne : (1 < radix)%Z.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ radixMoreThanOne).

Hypothesis pGivesBound : Zpos (vNum bo) = Zpower_nat radix p.
Hypothesis precisionGreaterThanOne : 3 <= p.

Variables a x b ph pl uh z:float.

Hypothesis Fb: Fbounded bo b.
Hypothesis Fa: Fbounded bo a.
Hypothesis Fx: Fbounded bo x.
Hypothesis Cb: Fcanonic radix bo b.
Hypothesis Nph: Fnormal radix bo ph.
Hypothesis Nz: Fnormal radix bo z.
Hypothesis Nuh: Fnormal radix bo uh.

Hypothesis Exp1: (- dExp bo <= Fexp a+Fexp x)%Z.

Hypothesis zDef : Closest bo radix (a*x+b)%R z.
Hypothesis phDef: Closest bo radix (a*x)%R ph.
Hypothesis plDef: (FtoRradix pl=a*x-ph)%R.
Hypothesis uhDef: Closest bo radix (ph+b)%R uh.
Hypothesis Posit: (0 <= a*x+b)%R.

Lemma zPos: (0 <= z)%R.
unfold FtoRradix; apply RleRoundedR0 with bo p (Closest bo radix) (a*x+b)%R; auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
Qed.

Lemma uhPos: (0 <= uh)%R.
unfold FtoRradix; apply RleRoundedR0 with bo p (Closest bo radix) (ph+b)%R; auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
apply Rplus_le_reg_l with (-b)%R.
ring_simplify.
unfold FtoRradix; rewrite <- Fopp_correct; auto.
apply RleBoundRoundl with bo p (Closest bo radix) (a*x)%R; auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
apply oppBounded; auto.
rewrite Fopp_correct; auto; fold FtoRradix; apply Rplus_le_reg_l with b.
apply Rle_trans with 0%R; auto with real; rewrite Rplus_comm; auto.
Qed.

Theorem tBounded_aux: exists v:float, Fbounded bo v /\ (FtoRradix v=uh-z)%R.
case (Req_dec (ph+b)%R 0);intros H1.
exists (Fopp z); split.
apply oppBounded; elim zDef; auto.
unfold FtoRradix; rewrite Fopp_correct; replace (FtoR radix uh) with 0%R;[ring|idtac].
apply trans_eq with (FtoR radix (Fzero (-(dExp bo)))).
rewrite FzeroisReallyZero; auto.
apply RoundedModeProjectorIdemEq with bo p (Closest bo radix); auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
apply FboundedFzero.
rewrite FzeroisReallyZero; rewrite <- H1; auto.
case (Rle_or_lt (Rabs pl) (/4*Rabs (ph+b))); intros H2.
exists (Fminus radix uh z).
split;[idtac|unfold FtoRradix; apply Fminus_correct; auto].
apply Sterbenz; auto.
elim uhDef; auto.
elim zDef; auto.
apply Rmult_le_reg_l with 2%R; auto with real.
apply Rle_trans with (FtoR radix z);[right; simpl; field; auto with real|idtac].
rewrite <- (Rabs_right (FtoR radix z)).
2: apply Rle_ge; generalize zPos; auto with real.
apply Rle_trans with (Rabs (a*x+b) / (1 - powerRZ radix (Z.succ (- p)) / 2))%R.
apply ClosestRoundeLeNormal with bo; auto with zarith.
assert (0 < 1 - powerRZ radix (Z.succ (- p)) / 2)%R.
apply UnMoinsPos; auto with zarith.
apply Rmult_le_reg_l with (1 - powerRZ radix (Z.succ (- p)) / 2)%R; auto with real.
apply Rle_trans with (Rabs (a*x+b));[right; field; auto with real|idtac].
assert (0 < 2 - powerRZ radix (Z.succ (- p)))%R; auto with real.
replace (2 - powerRZ radix (Z.succ (- p)))%R with
   (2*(1 - powerRZ radix (Z.succ (- p)) / 2))%R; auto with real.
replace 0%R with (2*0)%R; auto with real.
field; auto with real.
replace (a*x+b)%R with ((ph+b)+pl)%R;[idtac|rewrite plDef; ring].
apply Rle_trans with (Rabs (ph+b)+Rabs pl)%R;[apply Rabs_triang|idtac].
apply Rle_trans with (Rabs (ph+b)+/ 4 * Rabs (ph + b))%R;auto with real.
assert (0 < 4)%R;[apply Rmult_lt_0_compat; auto with real|idtac].
apply Rle_trans with (5/4*(Rabs (ph+b)))%R;[right; field; auto with real|idtac].
apply Rle_trans with (5/4*(Rabs (FtoR radix uh) * (1 + powerRZ radix (Z.succ (- p)) / 2)))%R.
apply Rmult_le_compat_l.
apply Rlt_le; unfold Rdiv;apply Rmult_lt_0_compat; auto with real.
apply ClosestRoundeGeNormal with bo; auto with zarith.
fold FtoRradix; apply Rle_trans with ((5/4*(1 + powerRZ radix (Z.succ (- p)) / 2))*Rabs uh)%R;
  [right; ring|idtac].
rewrite <- Rmult_assoc with (r3:=uh).
pattern (FtoRradix uh) at 2; rewrite <- (Rabs_right uh).
2: apply Rle_ge; generalize uhPos; auto with real.
apply Rmult_le_compat_r; auto with real.
apply Rabs_pos.
apply Rmult_le_reg_l with 8%R; [apply Rmult_lt_0_compat; auto with real|idtac].
apply Rle_trans with (10+5*powerRZ radix (Z.succ (- p)))%R;[right; field; auto with real|idtac].
apply Rle_trans with (16-8*powerRZ radix (Z.succ (- p)))%R;[idtac|right; field; auto with real].
apply Rplus_le_reg_l with (-10+8* powerRZ radix (Z.succ (- p)))%R.
ring_simplify.
apply Rle_trans with (13* /4)%R.
apply Rmult_le_compat_l; auto with real.
apply Rle_trans with (powerRZ radix (Z.succ (- 3))); unfold Z.succ; auto with zarith real.
apply Rle_powerRZ; auto with real zarith.
apply Rle_IZR; omega.
simpl; ring_simplify (radix*1)%R.
apply Rle_Rinv; auto with real.
rewrite <- mult_IZR; apply Rle_IZR.
apply Z.le_trans with (2*2)%Z; auto with zarith.
apply Zmult_le_compat; auto with zarith.
apply Rmult_le_reg_l with 4%R; auto with real.
apply Rle_trans with 13%R;[right; field; auto with real|idtac].
apply Rle_trans with 24%R; auto with real zarith.
right; ring.
simpl; fold FtoRradix.
rewrite <- (Rabs_right z);[idtac|apply Rle_ge; generalize zPos; auto with real].
rewrite <- (Rabs_right uh);[idtac|apply Rle_ge; generalize uhPos; auto with real].
assert (0 < 1 - powerRZ radix (Z.succ (- p)) / 2)%R.
apply UnMoinsPos; auto with zarith.
assert (0 < 4)%R;[apply Rmult_lt_0_compat; auto with real|idtac].
assert (0 < 3)%R;[apply Rlt_trans with 2%R; auto with real|idtac].
apply Rle_trans with (Rabs (ph+b) / (1 - powerRZ radix (Z.succ (- p)) / 2))%R.
unfold FtoRradix; apply ClosestRoundeLeNormal with bo; auto with zarith.
apply Rmult_le_reg_l with (1 - powerRZ radix (Z.succ (- p)) / 2)%R; auto with real.
apply Rle_trans with (Rabs (ph+b));[right; field; auto with real|idtac].
replace (2 - powerRZ radix (Z.succ (- p)))%R with
  (2* (1 - powerRZ radix (Z.succ (- p)) / 2))%R;[idtac|field].
apply prod_neq_R0; auto with real.
apply Rle_trans with ((4/3)*(Rabs (a*x+b)))%R.
apply Rmult_le_reg_l with (3/4)%R.
unfold Rdiv; apply Rmult_lt_0_compat; auto with real.
apply Rplus_le_reg_l with (/4*Rabs (ph + b))%R.
apply Rle_trans with (Rabs (ph+b));[right; field; auto with real|idtac].
apply Rle_trans with (/4*Rabs (ph+b)+Rabs (a*x+b))%R;[idtac|right; field; auto with real].
pattern (ph+b)%R at 1; replace (ph+b)%R with (-(pl)+(a*x+b))%R;[idtac|rewrite plDef; ring].
apply Rle_trans with (Rabs (-pl)+Rabs (a*x+b))%R;
   [apply Rabs_triang|rewrite Rabs_Ropp; auto with real].
apply Rle_trans with (4/3*(Rabs z * (1 + powerRZ radix (Z.succ (- p)) / 2)))%R.
apply Rmult_le_compat_l; auto with real.
apply Rlt_le; unfold Rdiv; apply Rmult_lt_0_compat; auto with real.
unfold FtoRradix; apply ClosestRoundeGeNormal with bo; auto with zarith.
apply Rle_trans with ((4 / 3 * (1 + powerRZ radix (Z.succ (- p)) / 2))*(Rabs z))%R;
  [right; ring|idtac].
rewrite <- Rmult_assoc with (r3:=Rabs z).
apply Rmult_le_compat_r; auto with real.
apply Rabs_pos.
apply Rmult_le_reg_l with 6%R; [apply Rmult_lt_0_compat; auto with real|idtac].
apply Rle_trans with (8+4*powerRZ radix (Z.succ (- p)))%R;[right; field; auto with real|idtac].
apply Rle_trans with (12-6*powerRZ radix (Z.succ (- p)))%R;[idtac|right; field; auto with real].
apply Rplus_le_reg_l with (-8+6* powerRZ radix (Z.succ (- p)))%R.
ring_simplify.
apply Rle_trans with (10* /4)%R.
apply Rmult_le_compat_l; auto with real.
apply Rle_trans with (powerRZ radix (Z.succ (- 3))); unfold Z.succ; auto with zarith real.
apply Rle_powerRZ; auto with real zarith.
apply Rle_IZR; omega.
simpl; ring_simplify (radix*1)%R; auto with real zarith.
apply Rle_Rinv; auto with real.
rewrite <- mult_IZR; apply Rle_IZR.
apply Z.le_trans with (2*2)%Z; auto with zarith.
apply Zmult_le_compat; auto with zarith.
apply Rmult_le_reg_l with 4%R; auto with real.
apply Rle_trans with 10%R;[right; field; auto with real|idtac].
apply Rle_trans with 16%R; auto with real zarith.
right; ring.
assert (K:(Fexp a+Fexp x < Fexp ph)%Z).
elim errorBoundedMult with bo radix p (Closest bo radix) a x ph; auto with zarith.
2: apply ClosestRoundedModeP with p; auto with zarith.
fold FtoRradix; intros pl' T; elim T; intros H3 T'; elim T'; intros H4 H5; clear T T'.
rewrite <- H5.
apply ClosestErrorExpStrict with bo radix p (a*x)%R; auto with zarith.
elim phDef; auto.
fold FtoRradix; contradict H2.
apply Rle_not_lt; rewrite plDef; rewrite <- H3; rewrite H2.
rewrite Rabs_R0; apply Rle_trans with (/4*0)%R; auto with real.
apply Rmult_le_compat_l; auto with real.
apply Rabs_pos.
assert (K':(-dExp bo < Fexp ph)%Z); auto with zarith.
assert (Rabs (ph+b) < 2* powerRZ radix (Fexp ph))%R.
apply Rmult_lt_reg_l with (/4)%R; auto with real.
assert (0 < 4)%R; auto with real.
apply Rlt_le_trans with (1:=H2).
apply Rmult_le_reg_l with (INR 2); auto with real zarith.
apply Rle_trans with (Fulp bo radix p ph).
rewrite plDef; unfold FtoRradix; apply ClosestUlp; auto with zarith.
rewrite CanonicFulp; auto with zarith.
unfold FtoR; right; simpl; field; auto with real.
left; auto.
assert (Fexp ph -1 <= Fexp b)%Z.
apply Z.le_trans with (Fexp (Float (Zpos (vNum bo)-2*radix) (Fexp ph -1)));
   [simpl; auto with zarith|idtac].
assert (0 < Zpos (vNum bo)-2*radix )%Z.
assert (2*radix < Zpos (vNum bo))%Z; auto with zarith.
rewrite pGivesBound; apply Z.le_lt_trans with (Zpower_nat radix 2); auto with zarith.
unfold Zpower_nat; simpl (nat_rect _ _ _ _).
ring_simplify (radix*1)%Z; apply Zmult_le_compat_r; auto with zarith.
apply Zpower_nat_monotone_lt; omega.
apply Fcanonic_Rle_Zle with radix bo p; auto with zarith.
left; split;[split|idtac]; auto with zarith.
apply Z.le_lt_trans with (Z.abs (Zpos (vNum bo) - 2 * radix)); auto with zarith.
simpl; auto with zarith.
apply Z.le_trans with (Z.abs (radix * (Zpos (vNum bo) - 2 * radix))); auto with zarith.
rewrite Zabs_Zmult; repeat rewrite Z.abs_eq; auto with zarith.
apply Z.le_trans with (2* (Zpos (vNum bo) - 2 * radix))%Z; auto with zarith.
apply Zplus_le_reg_l with (4*radix)%Z.
apply Z.le_trans with (Zpos (vNum bo)+ Zpos (vNum bo))%Z; auto with zarith.
assert (4*radix <= Zpos (vNum bo))%Z; auto with zarith.
rewrite pGivesBound; apply Z.le_trans with (Zpower_nat radix 3); auto with zarith.
unfold Zpower_nat; simpl (nat_rect _ _ _ _).
ring_simplify (radix*1)%Z; rewrite Zmult_assoc; apply Zmult_le_compat_r; auto with zarith.
apply Z.le_trans with (2*radix)%Z; auto with zarith.
apply Zpower_nat_monotone_le; omega.
fold FtoRradix; apply Rplus_le_reg_l with (2*powerRZ radix (Fexp ph))%R.
apply Rle_trans with (Rabs (ph + b)+Rabs b)%R; auto with real.
apply Rle_trans with (Rabs ph).
2: pattern (FtoRradix ph) at 1; replace (FtoRradix ph) with ((ph+b)+-b)%R;[idtac|ring].
2:apply Rle_trans with (Rabs (ph + b) + Rabs (-b))%R;
  [apply Rabs_triang| rewrite Rabs_Ropp; auto with real].
unfold FtoRradix; repeat rewrite <- Fabs_correct; auto.
unfold FtoR, Fabs.
simpl (Fexp (Float (Zpos (vNum bo) - 2 * radix) (Fexp ph - 1))).
simpl (Fexp
         (Float
            (Z.abs (Fnum (Float (Zpos (vNum bo) - 2 * radix) (Fexp ph - 1))))
            (Fexp ph - 1))).
replace (Fnum
      (Float (Z.abs (Fnum (Float (Zpos (vNum bo) - 2 * radix) (Fexp ph - 1))))
         (Fexp ph - 1))) with (Z.abs (Zpos (vNum bo) - 2 * radix)); auto with zarith.
simpl (Fnum (Float (Z.abs (Fnum ph)) (Fexp ph))).
simpl (Fexp (Float (Z.abs (Fnum ph)) (Fexp ph))).
apply Rle_trans with ((2+ ( Z.abs (Zpos (vNum bo) - 2 * radix))/radix)*powerRZ radix (Fexp ph))%R.
unfold Zminus; rewrite powerRZ_add; try apply IZR_neq; auto with real zarith; simpl; ring_simplify (radix*1)%R.
right; field; auto with real zarith.
apply IZR_neq; auto with zarith.
apply Rmult_le_compat_r; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
apply Rmult_le_reg_l with radix; auto with real zarith.
apply Rlt_IZR; omega.
apply Rle_trans with (IZR (Zpos (vNum bo))).
rewrite Z.abs_eq; auto with zarith.
unfold Zminus; rewrite plus_IZR; rewrite Ropp_Ropp_IZR; rewrite mult_IZR.
right; simpl; field; try apply IZR_neq; auto with real zarith.
elim Nph; intros.
apply Rle_trans with (IZR (Z.abs (radix * Fnum ph))); auto with real zarith.
apply Rle_IZR; auto.
right; rewrite Zabs_Zmult; rewrite mult_IZR; rewrite Z.abs_eq; auto with zarith real.
assert (exists uh':float, (FtoRradix uh'=ph+b)%R /\ (Fbounded bo uh') /\ (Fexp uh'=Fexp ph-1)%Z).
unfold FtoRradix; apply BoundedL with p (Fplus radix ph b); auto with zarith.
simpl; apply Zmin_Zle; auto with zarith.
apply Fplus_correct; auto with zarith.
fold FtoRradix; apply Rlt_le_trans with (1:=H).
apply Rle_trans with (powerRZ radix (Fexp ph+1));
  [rewrite powerRZ_add; try apply IZR_neq; auto with real zarith; simpl|apply Rle_powerRZ; auto with real zarith].
2: apply Rle_IZR; omega.
ring_simplify (radix*1)%R; rewrite Rmult_comm; apply Rmult_le_compat_l; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
apply Rle_IZR; omega.
elim H3; clear H3; intros uh' T; elim T; intros H5 T'; elim T'; intros H6 H7; clear T T'.
assert (FtoRradix uh=uh').
unfold FtoRradix; apply sym_eq.
apply RoundedModeProjectorIdemEq with bo p (Closest bo radix); auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
fold FtoRradix; rewrite H5; auto.
elim errorBoundedMult with bo radix p (Closest bo radix) a x ph; auto with zarith.
2: apply ClosestRoundedModeP with p; auto with zarith.
fold FtoRradix; intros pl' T; elim T; intros H8 T'; elim T'; intros H9 H10; clear T T'.
elim LeExpRound2 with bo radix p (Fexp a+Fexp x)%Z (Fplus radix (Fplus radix ph pl') b) z;
  auto with zarith.
2: rewrite Fplus_correct; auto with zarith;rewrite Fplus_correct; auto with zarith.
2: fold FtoRradix; rewrite H8.
2: ring_simplify (ph + (a * x - ph) + b)%R; auto with real.
2: simpl; apply Zmin_Zle; auto with zarith.
fold FtoRradix; intros z' T; elim T; intros H15 T'; elim T'; intros H16 H17; clear T T'.
cut ( exists v : float, (FtoRradix v = uh - z)%R /\ Fbounded bo v /\ (Fexp v=Fexp a+Fexp x)%Z).
intros T; elim T; intros v T1; elim T1; intros T2 T3; elim T3; intros; exists v; split; auto.
unfold FtoRradix; apply BoundedL with p (Fminus radix uh' z'); auto with zarith.
simpl; apply Zmin_Zle; auto with zarith.
rewrite Fminus_correct; auto with zarith; fold FtoRradix;
  rewrite <- H3; rewrite H16; auto with real.
apply Rlt_le_trans with (powerRZ radix (Fexp ph));
  [idtac|apply Rle_powerRZ; auto with real zarith].
fold FtoRradix; rewrite H3; rewrite H5.
replace (ph+b-z)%R with ((a*x+b-z)+-pl)%R;[idtac|rewrite plDef; ring].
apply Rle_lt_trans with (Rabs (a*x+b-z)+Rabs (-pl))%R;[apply Rabs_triang|rewrite Rabs_Ropp].
apply Rmult_lt_reg_l with (INR 2); auto with real zarith.
apply Rle_lt_trans with (S 1 * Rabs (a * x + b - z) + S 1*Rabs pl)%R;[right; ring|idtac].
apply Rlt_le_trans with (powerRZ radix (Fexp ph)+powerRZ radix (Fexp ph))%R;
  [idtac|simpl;right; ring].
cut (S 1 * Rabs (a * x + b - z) < powerRZ radix (Fexp ph))%R;[intros I1|idtac].
cut (S 1 * Rabs (pl) <= powerRZ radix (Fexp ph))%R;[intros I2; auto with real|idtac].
apply Rle_trans with (Fulp bo radix p ph).
rewrite plDef; unfold FtoRradix; apply ClosestUlp; auto with zarith.
rewrite CanonicFulp; auto with zarith real.
right; unfold FtoR; simpl; ring.
left; auto.
apply Rle_lt_trans with (Fulp bo radix p z).
unfold FtoRradix; apply ClosestUlp; auto with zarith.
rewrite CanonicFulp; auto with zarith real.
2: left; auto.
2: apply Rle_IZR; omega.
unfold FtoR; simpl; ring_simplify (1*powerRZ radix (Fexp z))%R.
apply Rlt_powerRZ; auto with real zarith.
apply Rlt_IZR; omega.
apply Z.le_lt_trans with (Fexp ph-1)%Z; auto with zarith.
assert (Fbounded bo (Float (3*radix) (Fexp ph-1))).
split; [idtac|simpl; auto with zarith].
apply Z.le_lt_trans with (Z.abs (3*radix)); auto with zarith.
rewrite Zabs_Zmult; repeat rewrite Z.abs_eq; auto with zarith.
rewrite pGivesBound; apply Z.lt_le_trans with (Zpower_nat radix 3); auto with zarith.
unfold Zpower_nat; simpl (nat_rect _ _ _ _)%Z.
rewrite Zmult_comm; ring_simplify (radix*1)%Z; apply Zmult_lt_compat_l; auto with zarith.
apply Z.lt_le_trans with (2*2)%Z; auto with zarith.
apply Z.le_trans with (2*radix)%Z; auto with zarith.
apply Zpower_nat_monotone_le; omega.
apply Z.le_trans with (Fexp (Float (3*radix) (Fexp ph-1))); auto with zarith.
apply Z.le_trans with (Fexp (Fnormalize radix bo p (Float (3*radix) (Fexp ph-1)))).
apply Fcanonic_Rle_Zle with radix bo p; auto with zarith.
left; auto.
apply FnormalizeCanonic; auto with zarith.
rewrite FnormalizeCorrect; auto.
apply Rle_trans with (FtoR radix (Float (3 * radix) (Fexp ph - 1)));
  [idtac|rewrite <- Fabs_correct; auto; unfold FtoR].
2: simpl; rewrite Z.abs_eq; auto with zarith real.
2: apply Z.le_trans with (3*radix)%Z; auto with zarith.
apply RoundAbsMonotoner with bo p (Closest bo radix) (a*x+b)%R; auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
replace (a*x+b)%R with ((ph+b)+pl)%R;[idtac|rewrite plDef; ring].
apply Rle_trans with (Rabs (ph+b)+Rabs pl)%R;[apply Rabs_triang|idtac].
apply Rle_trans with (2 * powerRZ radix (Fexp ph)+/2*powerRZ radix (Fexp ph))%R;
  [apply Rplus_le_compat; auto with real|idtac].
apply Rmult_le_reg_l with (INR 2); auto with real zarith.
apply Rle_trans with (Fulp bo radix p ph).
rewrite plDef; unfold FtoRradix; apply ClosestUlp; auto with zarith.
rewrite CanonicFulp; auto with zarith.
unfold FtoR; right; simpl; field; auto with real.
left; auto.
apply Rle_trans with ((2+/2)*powerRZ radix (Fexp ph))%R;[right; ring|idtac].
apply Rle_trans with (3*powerRZ radix (Fexp ph))%R;
  [idtac|right; unfold FtoR; simpl; unfold Zminus; rewrite powerRZ_add; auto with real zarith].
2: simpl; ring_simplify (radix*1)%R.
2: replace (IZR (match radix with
  | Z0 => 0%Z
  | Zpos y' => Zpos ((y' + xO y'))
  | Zneg y' => Zneg ((y' + xO y')) end)) with (3*radix)%R;[field; auto with real zarith|idtac].
2: apply IZR_neq; omega.
2: apply trans_eq with (IZR (3*radix)); auto with real zarith; rewrite mult_IZR; simpl; ring.
2: apply IZR_neq; omega.
apply Rmult_le_compat_r; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
apply Rle_trans with (2+1)%R; auto with real.
apply Rplus_le_compat_l; apply Rle_trans with (/1)%R; auto with real.
right; ring.
apply FcanonicLeastExp with radix bo p; auto with zarith.
rewrite FnormalizeCorrect; auto with real zarith.
apply FnormalizeCanonic; auto with zarith.
assert (0 < pPred (vNum bo))%Z.
apply pPredMoreThanOne with radix p; auto with zarith.
assert (Fbounded bo (Float (pPred (vNum bo)) (Fexp a+Fexp x+p))).
split; simpl; auto with zarith.
rewrite Z.abs_eq;[unfold pPred|idtac]; auto with zarith.
apply Z.le_trans with (Fexp (Float (pPred (vNum bo)) (Fexp a+Fexp x+p))); auto with zarith.
apply Fcanonic_Rle_Zle with radix bo p; auto with zarith.
left; auto.
left; split; auto.
simpl; rewrite Zabs_Zmult; repeat rewrite Z.abs_eq; auto with zarith.
apply Z.le_trans with (2* pPred (vNum bo))%Z; auto with zarith.
unfold pPred, Z.pred; apply Zplus_le_reg_l with 2%Z.
apply Z.le_trans with (Zpos (vNum bo) + Zpos (vNum bo))%Z; auto with zarith.
assert (2 <= Zpos (vNum bo))%Z; auto with zarith.
rewrite pGivesBound; apply Z.le_trans with (Zpower_nat radix 1); auto with zarith.
unfold Zpower_nat; simpl; auto with zarith.
apply Zpower_nat_monotone_le; omega.
apply Rle_trans with (FtoR radix (Float (pPred (vNum bo)) (Fexp a + Fexp x + p))).
2: right; rewrite <- Fabs_correct; auto; unfold FtoR; simpl;
    rewrite Z.abs_eq; auto with real zarith.
apply RoundAbsMonotoner with bo p (Closest bo radix) (a*x)%R; auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
rewrite Rabs_mult; unfold FtoRradix; repeat rewrite <- Fabs_correct; auto.
unfold FtoR; simpl; repeat rewrite powerRZ_add; try apply IZR_neq; auto with real zarith.
apply Rle_trans with (Z.abs (Fnum a) *
    (powerRZ radix (Fexp a) * powerRZ radix (Fexp x) * Z.abs (Fnum x)))%R;[right; ring|idtac].
apply Rmult_le_compat; auto with real zarith.
apply Rle_IZR; auto with zarith.
apply Rle_trans with (0*Z.abs (Fnum x))%R; auto with real.
apply Rmult_le_compat_r; auto with real zarith.
apply Rle_IZR; auto with zarith.
apply Rlt_le; apply Rmult_lt_0_compat; auto with real zarith; apply powerRZ_lt, Rlt_IZR; omega.
elim Fa; intros; unfold pPred; apply IZR_le; auto with real zarith.
apply Rmult_le_compat_l; auto with real zarith.
apply Rlt_le; apply Rmult_lt_0_compat; auto with real zarith; apply powerRZ_lt, Rlt_IZR; omega.
apply Rle_trans with (IZR (Zpos (vNum bo))).
apply Rle_IZR; elim Fx; intros; auto with real zarith.
rewrite pGivesBound; rewrite Zpower_nat_Z_powerRZ; auto with real.
Qed.

End tBounded.

Section tBounded2.
Variable bo : Fbound.
Variable radix : Z.
Variable p : nat.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.
Hypothesis radixMoreThanOne : (1 < radix)%Z.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ radixMoreThanOne).

Hypothesis pGivesBound : Zpos (vNum bo) = Zpower_nat radix p.
Hypothesis precisionGreaterThanOne : 3 <= p.

Variables a x b ph pl uh z:float.

Hypothesis Fb: Fbounded bo b.
Hypothesis Fa: Fbounded bo a.
Hypothesis Fx: Fbounded bo x.
Hypothesis Cb: Fcanonic radix bo b.
Hypothesis Nph: Fnormal radix bo ph \/ (FtoRradix ph=0).
Hypothesis Nz: Fnormal radix bo z \/ (FtoRradix z =0).
Hypothesis Nuh: Fnormal radix bo uh \/ (FtoRradix uh=0).

Hypothesis Exp1: (- dExp bo <= Fexp a+Fexp x)%Z.

Hypothesis zDef : Closest bo radix (a*x+b)%R z.
Hypothesis phDef: Closest bo radix (a*x)%R ph.
Hypothesis plDef: (FtoRradix pl=a*x-ph)%R.
Hypothesis uhDef: Closest bo radix (ph+b)%R uh.

Theorem tBounded: exists v:float, Fbounded bo v /\ (FtoRradix v=uh-z)%R.
case Nuh; intros I1.
case Nz; intros I2.
case Nph; intros I3.
case (Rle_or_lt 0 (a*x+b)); intros H.
unfold FtoRradix; apply tBounded_aux with p a x b ph pl; auto.
elim tBounded_aux with bo radix p
   (Fopp a) x (Fopp b) (Fopp ph) (Fopp pl) (Fopp uh) (Fopp z); auto with zarith.
fold FtoRradix; intros v T; elim T; intros; clear T.
exists (Fopp v); split;[apply oppBounded; auto|idtac].
unfold FtoRradix; rewrite Fopp_correct; fold FtoRradix; rewrite H1;
   unfold FtoRradix; repeat rewrite Fopp_correct; ring.
apply oppBounded; auto.
apply oppBounded; auto.
apply FcanonicFopp; auto.
apply FnormalFop; auto.
apply FnormalFop; auto.
apply FnormalFop; auto.
replace (FtoR radix (Fopp a) * FtoR radix x + FtoR radix (Fopp b))%R with (-(a*x+b))%R;
  [apply ClosestOpp; auto|repeat rewrite Fopp_correct; fold FtoRradix;ring].
replace (FtoR radix (Fopp a) * FtoR radix x )%R with (-(a*x))%R;
  [apply ClosestOpp; auto|unfold FtoRradix; repeat rewrite Fopp_correct; ring].
repeat rewrite Fopp_correct; fold FtoRradix; rewrite plDef; ring.
replace (FtoR radix (Fopp ph) + FtoR radix (Fopp b) )%R with (-(ph+b))%R;
  [apply ClosestOpp; auto|unfold FtoRradix; repeat rewrite Fopp_correct; ring].
apply Rle_trans with (-(a*x+b))%R; auto with real.
right; unfold FtoRradix; repeat rewrite Fopp_correct; ring.
assert (a*x=0)%R.
apply ClosestZero1 with bo radix p ph (Fmult a x); auto with zarith.
rewrite Fmult_correct; fold FtoRradix; auto.
exists (Fzero (-(dExp bo))); split.
apply FboundedFzero; auto.
unfold FtoRradix; rewrite FzeroisReallyZero; fold FtoRradix.
replace (FtoRradix uh) with (FtoRradix b).
replace (FtoRradix z) with (FtoRradix b);[ring|idtac].
unfold FtoRradix; apply RoundedModeProjectorIdemEq with bo p (Closest bo radix); auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
replace (FtoR radix b) with (a*x+b)%R; auto; rewrite H; auto with real.
unfold FtoRradix; apply RoundedModeProjectorIdemEq with bo p (Closest bo radix); auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
replace (FtoR radix b) with (ph+b)%R; auto; rewrite I3; auto with real.
exists uh; split.
elim uhDef; auto.
rewrite I2; simpl; ring.
exists (Fopp z); split.
apply oppBounded; elim zDef; auto.
rewrite I1; unfold FtoRradix; rewrite Fopp_correct; fold FtoRradix;simpl ; ring.
Qed.

End tBounded2.

Section uhExact.
Variable bo : Fbound.
Variable radix : Z.
Variable p : nat.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.
Hypothesis radixMoreThanOne : (1 < radix)%Z.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ radixMoreThanOne).

Hypothesis pGivesBound : Zpos (vNum bo) = Zpower_nat radix p.
Hypothesis precisionGreaterThanOne : 3 <= p.

Variables a x b ph pl uh ul z t v w:float.

Hypothesis Fb: Fbounded bo b.
Hypothesis Fa: Fbounded bo a.
Hypothesis Fx: Fbounded bo x.
Hypothesis Cb: Fcanonic radix bo b.
Hypothesis Nph: Fnormal radix bo ph \/ (FtoRradix ph=0).
Hypothesis Nuh: Fnormal radix bo uh \/ (FtoRradix uh=0).

Hypothesis Nz: Fnormal radix bo z \/ (FtoRradix z=0).
Hypothesis Nw: Fnormal radix bo w \/ (FtoRradix w=0).
Hypothesis Fpl: Fbounded bo pl.

Hypothesis Exp1: (- dExp bo <= Fexp a+Fexp x)%Z.

Hypothesis zDef : Closest bo radix (a*x+b)%R z.
Hypothesis phDef: Closest bo radix (a*x)%R ph.
Hypothesis plDef: (FtoRradix pl=a*x-ph)%R.
Hypothesis uhDef: Closest bo radix (ph+b)%R uh.
Hypothesis ulDef: (FtoRradix ul=ph+b-uh)%R.
Hypothesis tDef : Closest bo radix (uh-z)%R t.
Hypothesis vDef : Closest bo radix (pl+ul)%R v.
Hypothesis wDef : Closest bo radix (t+v)%R w.

Hypothesis Case1:(FtoRradix ul=0)%R.

Theorem ErrFmaApprox_1_aux:
   Fnormal radix bo z -> Fnormal radix bo w ->
   (Rabs (z+w-(a*x+b)) <= (3*radix/2+/2)*powerRZ radix (2-2*p)*Rabs z)%R.
intros M1 M2.
apply Rle_trans with (3*powerRZ radix (2-2*p)*Rabs z)%R.
assert (FtoRradix t=uh-z)%R.
elim tBounded with bo radix p a x b ph pl uh z; auto.
fold FtoRradix; intros t' T; elim T; intros T1 T2; clear T.
rewrite <- T2; unfold FtoRradix; apply sym_eq.
apply RoundedModeProjectorIdemEq with bo p (Closest bo radix); auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
fold FtoRradix; rewrite T2; auto.
assert (Rabs (z + w - (a * x + b)) <= (Fulp bo radix p w))%R.
replace (z + w - (a * x + b))%R with (-((t+v)-w))%R.
2: rewrite H; replace (a*x)%R with (a*x-0)%R; auto with real; rewrite <- Case1.
2: rewrite ulDef; replace (FtoRradix v) with (FtoRradix pl).
2: rewrite plDef; ring.
2: unfold FtoRradix; apply RoundedModeProjectorIdemEq
    with bo p (Closest bo radix); auto with zarith.
2: apply ClosestRoundedModeP with p; auto with zarith.
2: replace (FtoR radix pl) with (pl+ul)%R; auto;fold FtoRradix; rewrite Case1; auto with real.
rewrite Rabs_Ropp; apply Rlt_le; unfold FtoRradix.
apply RoundedModeUlp with (Closest bo radix); auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
apply Rle_trans with (1:=H0).
apply Rle_trans with (Rabs (FtoR radix w) * powerRZ radix (Z.succ (- p)))%R.
apply FulpLe2; auto with zarith.
elim M2; auto.
rewrite FcanonicFnormalizeEq; auto with zarith.
left; auto.
fold FtoRradix; apply Rle_trans with
  ((3*(powerRZ radix (Z.succ (-p))*Rabs z))* powerRZ radix (Z.succ (- p)))%R.
2: unfold Z.succ; replace (2-2*p)%Z with (1+1+-p+-p)%Z; auto with zarith.
2: repeat rewrite powerRZ_add; try apply IZR_neq; auto with real zarith; right; ring.
apply Rmult_le_compat_r; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
apply Rle_trans with (3*Fulp bo radix p z)%R.
2: apply Rmult_le_compat_l; auto with real zarith.
2: rewrite Rmult_comm; unfold FtoRradix; apply FulpLe2; auto with zarith.
2: elim M1; auto.
2: rewrite FcanonicFnormalizeEq; auto with zarith; left; auto.
assert (0 < 1- powerRZ radix (Z.succ (- p)))%R.
assert (powerRZ radix (Z.succ (- p)) < 1)%R; auto with real.
apply Rlt_le_trans with (powerRZ radix (Z.succ (-1))); unfold Z.succ; auto with zarith real.
apply Rlt_powerRZ; try apply Rlt_IZR; omega.
now apply Rlt_Rminus.
apply Rmult_le_reg_l with (1- powerRZ radix (Z.succ (- p)))%R; auto with real.
apply Rle_trans with (Fulp bo radix p z).
apply Rplus_le_reg_l with (powerRZ radix (Z.succ (- p))*Rabs w)%R.
apply Rle_trans with (Rabs w);[right; ring|idtac].
pattern (FtoRradix w) at 1; replace (FtoRradix w) with ((z+w-(a*x+b))+((a*x+b)-z))%R;
  [idtac|ring].
apply Rle_trans with (Rabs (z+w-(a*x+b))+Rabs (((a*x+b)-z)))%R; [apply Rabs_triang|idtac].
apply Rplus_le_compat.
apply Rle_trans with (1:=H0).
rewrite Rmult_comm; unfold FtoRradix; apply FulpLe2; auto with zarith.
elim M2; auto.
rewrite FcanonicFnormalizeEq; auto with zarith; left; auto.
apply Rlt_le; unfold FtoRradix.
apply RoundedModeUlp with (Closest bo radix); auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
apply Rle_trans with (1*Fulp bo radix p z)%R; auto with real.
unfold Fulp; rewrite <- Rmult_assoc; apply Rmult_le_compat_r; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
apply Rplus_le_reg_l with (-1+3*powerRZ radix (Z.succ (- p)))%R.
apply Rle_trans with (3 * powerRZ radix (Z.succ (- p)))%R;[right; ring|idtac].
ring_simplify ( -1 + 3 * powerRZ radix (Z.succ (- p)) +
    (1 - powerRZ radix (Z.succ (- p))) * 3)%R.
apply Rle_trans with (powerRZ radix 0); auto with real zarith.
apply Rle_trans with (powerRZ radix (2+Z.succ (-p))); auto with real zarith.
2: unfold Z.succ; apply Rle_powerRZ; auto with real zarith.
2: apply Rle_IZR; omega.
rewrite powerRZ_add; try apply IZR_neq; auto with real zarith; simpl.
apply Rmult_le_compat_r; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
ring_simplify (radix*1)%R; apply Rle_trans with 4%R; auto with real zarith.
rewrite <- mult_IZR; apply Rle_IZR.
apply Z.le_trans with (2*2)%Z; auto with zarith.
apply Zmult_le_compat; auto with zarith.
simpl; auto with real.
apply Rmult_le_compat_r; auto with real.
apply Rabs_pos.
apply Rmult_le_compat_r; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
apply Rle_trans with (3 * radix / 2)%R; auto with real.
apply Rmult_le_reg_l with (IZR 2); auto with real zarith.
apply Rle_trans with (radix*3)%R;[apply Rmult_le_compat_r; auto with real zarith|
      right; simpl; field; auto with real].
apply Rle_IZR; omega.
apply Rle_trans with (3 * radix / 2+0)%R; auto with real.
Qed.

Theorem ErrFmaApprox_1: (Rabs (z+w-(a*x+b)) <= (3*radix/2+/2)*powerRZ radix (2-2*p)*Rabs z)%R.
case Nz; intros I1.
case Nw; intros I2.
apply ErrFmaApprox_1_aux; auto.
replace (z+w-(a*x+b))%R with 0%R.
rewrite Rabs_R0; apply Rle_trans with (0*Rabs z)%R;
   [right; ring|apply Rmult_le_compat_r; auto with real].
apply Rabs_pos.
apply Rlt_le; apply Rmult_lt_0_compat; auto with real zarith.
apply Rle_lt_trans with (0+0)%R; auto with real; apply Rplus_lt_compat; auto with real.
unfold Rdiv;repeat apply Rmult_lt_0_compat; auto with real zarith.
apply Rlt_IZR; omega.
apply powerRZ_lt, Rlt_IZR; omega.
apply sym_eq; apply trans_eq with (z+w-uh-(pl+ul))%R;[rewrite plDef; rewrite ulDef; ring|idtac].
rewrite I2; rewrite Case1; simpl.
ring_simplify.
assert (FtoRradix t=uh-z)%R.
elim tBounded with bo radix p a x b ph pl uh z; auto with zarith.
fold FtoRradix; intros t' T; elim T; intros; clear T.
rewrite <- H0; unfold FtoRradix; apply sym_eq.
apply RoundedModeProjectorIdemEq with bo p (Closest bo radix); auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
fold FtoRradix; rewrite H0; auto.
cut (FtoRradix pl=-t)%R;[intros K; rewrite K; rewrite H; ring|idtac].
apply trans_eq with (FtoRradix v).
unfold FtoRradix;apply RoundedModeProjectorIdemEq with bo p (Closest bo radix); auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
replace (FtoR radix pl) with (pl+ul)%R; auto; rewrite Case1; auto with real.
assert (t+v=0)%R; auto with real.
apply ClosestZero1 with bo radix p w (Fplus radix t v); auto with zarith.
rewrite Fplus_correct; fold FtoRradix; auto with real.
simpl; apply Zmin_Zle.
elim tDef; intros J1 J2; elim J1; auto.
elim vDef; intros J1 J2; elim J1; auto.
apply Rplus_eq_reg_l with t; rewrite H0; auto with real.
replace (z+w-(a*x+b))%R with 0%R.
rewrite Rabs_R0; apply Rle_trans with (0*Rabs z)%R;
   [right; ring|apply Rmult_le_compat_r; auto with real].
apply Rabs_pos.
apply Rlt_le; apply Rmult_lt_0_compat; auto with real zarith.
apply Rle_lt_trans with (0+0)%R; auto with real; apply Rplus_lt_compat; auto with real.
unfold Rdiv;repeat apply Rmult_lt_0_compat; auto with real zarith.
apply Rlt_IZR; omega.
apply powerRZ_lt, Rlt_IZR; omega.
assert (a*x+b=0)%R.
apply ClosestZero1 with bo radix p z (Fplus radix (Fmult a x) b); auto with zarith.
rewrite Fplus_correct; auto; rewrite Fmult_correct; fold FtoRradix; auto with real.
simpl; apply Zmin_Zle; auto with zarith; apply Fb.
rewrite I1; rewrite H; simpl; ring_simplify.
apply sym_eq; unfold FtoRradix; apply ClosestZero2 with bo p (t+v)%R; auto.
cut (FtoRradix uh=0);[intros M1|idtac].
replace (FtoRradix t) with 0%R.
replace (FtoRradix v) with 0%R;[ring|idtac].
apply sym_eq; unfold FtoRradix; apply ClosestZero2 with bo p (pl+ul)%R; auto.
rewrite plDef; rewrite ulDef.
apply trans_eq with ((a*x+b)-uh)%R;[ring|rewrite H; rewrite M1; simpl; ring].
apply sym_eq; unfold FtoRradix; apply ClosestZero2 with bo p (uh-z)%R; auto.
rewrite M1; rewrite I1; ring.
unfold FtoRradix; apply ClosestZero2 with bo p (ph+b)%R; auto.
replace (FtoRradix ph) with (FtoRradix (Fopp b));[unfold FtoRradix; rewrite Fopp_correct; ring|idtac].
unfold FtoRradix;apply RoundedModeProjectorIdemEq with bo p (Closest bo radix); auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
apply oppBounded; auto.
replace (FtoR radix (Fopp b)) with (a*x)%R; auto.
apply Rplus_eq_reg_l with (FtoRradix b); rewrite Rplus_comm; rewrite H;
  rewrite Fopp_correct;auto with real.
Qed.

End uhExact.

Section uhInexact.
Variable bo : Fbound.
Variable radix : Z.
Variable p : nat.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.
Hypothesis radixMoreThanOne : (1 < radix)%Z.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ radixMoreThanOne).

Hypothesis pGivesBound : Zpos (vNum bo) = Zpower_nat radix p.
Hypothesis precisionGreaterThanOne : 4 <= p.

Lemma RleRoundedAbs: forall (f:float) (r:R), (Closest bo radix r f) ->
  (Fnormal radix bo f) -> (-(dExp bo) < Fexp f)%Z ->
  ((powerRZ radix (p - 1) + - / (2 * radix)) * powerRZ radix (Fexp f) <= Rabs r)%R.
intros f r H H0 H'0.
assert (exists f':float, f'=(Float (nNormMin radix p) (Fexp f))).
exists (Float (nNormMin radix p) (Fexp f)); auto.
elim H1; clear H1; intros f' H1.
assert (Fbounded bo f').
rewrite H1; split; simpl; auto with zarith.
rewrite Z.abs_eq.
apply ZltNormMinVnum; auto with zarith.
apply Zlt_le_weak; apply nNormPos; auto with zarith.
assert (Fnormal radix bo f').
split; auto.
rewrite H1; simpl; rewrite <- PosNormMin with radix bo p; auto with zarith.
assert (f' <= Fabs f)%R.
rewrite H1; unfold FtoRradix, FtoR; simpl.
apply Rmult_le_compat_r; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
assert (nNormMin radix p <= Z.abs (Fnum f))%Z; try apply Rle_IZR; auto with real zarith.
apply pNormal_absolu_min with bo; auto with zarith.
apply Rle_trans with (f'-powerRZ radix (Z.pred (Fexp f))/2)%R.
right; rewrite H1; unfold FtoRradix, FtoR, Z.pred, Zminus; simpl.
rewrite powerRZ_add with (n:=Fexp f); try apply IZR_neq; auto with real zarith; simpl; ring_simplify (radix*1)%R.
replace (IZR (nNormMin radix p)) with (powerRZ radix (p + -1)).
field; try apply IZR_neq; auto with real zarith.
unfold nNormMin; rewrite Zpower_nat_Z_powerRZ; rewrite inj_pred; auto with real zarith.
case (Rle_or_lt (f' - powerRZ radix (Z.pred (Fexp f)) / 2) (Rabs r));
  auto; intros I.
absurd (Rabs (Fabs f - Rabs r) <= Rabs ((FPred bo radix p f') - Rabs r))%R.
2: assert (K: (Closest bo radix (Rabs r) (Fabs f))).
2: apply ClosestFabs with p; auto with zarith.
2: elim K; intros K1 K2; unfold FtoRradix; apply K2; auto.
2: apply FBoundedPred; auto with zarith.
apply Rlt_not_le.
rewrite Rabs_left1.
rewrite Rabs_right with (Fabs f - Rabs r)%R.
apply Rplus_lt_reg_r with (Rabs r+(FPred bo radix p f'))%R.
ring_simplify.
apply Rlt_le_trans with (2*(f' - powerRZ radix (Z.pred (Fexp f)) / 2))%R.
apply Rmult_lt_compat_l; auto with real.
apply Rle_trans with (f'+ FPred bo radix p f')%R; auto with real.
apply Rle_trans with (2*f'- powerRZ radix (Z.pred (Fexp f)))%R;
   [right; field; auto with real|idtac].
apply Rplus_le_reg_l with (-f'- FPred bo radix p f'+ powerRZ radix (Z.pred (Fexp f)))%R.
ring_simplify.
right; apply trans_eq with (FtoRradix (Fminus radix f' (FPred bo radix p f'))).
unfold FtoRradix; rewrite Fminus_correct; auto with real zarith.
unfold FtoRradix; rewrite FPredDiff3; auto with zarith.
rewrite H1; simpl; unfold FtoR; simpl; ring.
rewrite H1; auto.
rewrite H1; simpl; auto with zarith.
rewrite Rplus_comm; auto with real.
apply Rle_ge.
assert (Rabs r < Fabs f)%R; auto with real.
apply Rlt_le_trans with (1:=I).
apply Rle_trans with (Fabs f-0)%R; auto with real; unfold Rminus;
  apply Rplus_le_compat; auto with real zarith.
apply Ropp_le_contravar; unfold Rdiv; apply Rlt_le; apply Rmult_lt_0_compat; auto with real zarith.
apply powerRZ_lt, Rlt_IZR; omega.
left; apply Rlt_Rminus; auto.
assert ((FPred bo radix p f' <= Rabs r))%R; auto with real.
2: apply Rle_trans with (Rabs r-Rabs r)%R; unfold Rminus; auto with real.
case (Rle_or_lt (FPred bo radix p f') (Rabs r)); auto; intros I'.
absurd (FPred bo radix p f' < f')%R.
2: unfold FtoRradix; apply FPredLt; auto with zarith.
apply Rle_not_lt.
apply Rle_trans with (1:=H4).
unfold FtoRradix; rewrite Fabs_correct; auto.
apply RoundAbsMonotoner with bo p (Closest bo radix) r; auto with zarith real.
apply ClosestRoundedModeP with p; auto with zarith.
apply FBoundedPred; auto with zarith.
Qed.

Variables a x b ph pl uh ul z t v w:float.

Hypothesis Fb: Fbounded bo b.
Hypothesis Fa: Fbounded bo a.
Hypothesis Fx: Fbounded bo x.
Hypothesis Cb: Fcanonic radix bo b.
Hypothesis Nph: Fnormal radix bo ph.
Hypothesis Nuh: Fnormal radix bo uh.

Hypothesis Nz: Fnormal radix bo z.
Hypothesis Nw: Fnormal radix bo w.
Hypothesis Nv: Fnormal radix bo v.

Hypothesis Exp1: (- dExp bo <= Fexp a+Fexp x)%Z.

Hypothesis zDef : Closest bo radix (a*x+b)%R z.
Hypothesis phDef: Closest bo radix (a*x)%R ph.
Hypothesis plDef: (FtoRradix pl=a*x-ph)%R.
Hypothesis uhDef: Closest bo radix (ph+b)%R uh.
Hypothesis ulDef: (FtoRradix ul=ph+b-uh)%R.
Hypothesis tDef : Closest bo radix (uh-z)%R t.
Hypothesis vDef : Closest bo radix (pl+ul)%R v.
Hypothesis wDef : Closest bo radix (t+v)%R w.

Hypothesis Case2: ~(FtoRradix ul=0)%R.

Lemma LeExp1: (Fexp ph <= Fexp uh+1)%Z.
case (Zle_or_lt (Fexp ph) (Fexp uh+1)); auto with zarith; intros.
absurd (FtoRradix ul=0); auto with real.
rewrite ulDef; assert (FtoRradix uh=ph+b)%R; auto with real.
unfold FtoRradix; apply plusExact2 with bo p; auto with real zarith.
left; auto.
Qed.

Lemma LeExp2: (Fexp uh <= Fexp z+1)%Z.
assert (Rabs (uh-z) <= (1+radix)* Fulp bo radix p uh + Fulp bo radix p z)%R.
replace (uh-z)%R with (-(ph+b-uh)+-(a*x-ph)+(a*x+b-z))%R;[idtac|ring].
apply Rle_trans with (1:=Rabs_triang (-(ph+b-uh)+-(a*x-ph)) (a*x+b-z)).
apply Rplus_le_compat.
apply Rle_trans with (1:=Rabs_triang (-(ph+b-uh)) (-(a*x-ph))).
apply Rle_trans with (Fulp bo radix p uh+radix*Fulp bo radix p uh)%R;
   [apply Rplus_le_compat|right; ring].
rewrite Rabs_Ropp; apply Rlt_le; unfold FtoRradix.
apply RoundedModeUlp with (Closest bo radix); auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
apply Rle_trans with (Fulp bo radix p ph).
rewrite Rabs_Ropp; apply Rlt_le; unfold FtoRradix.
apply RoundedModeUlp with (Closest bo radix); auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
unfold Fulp; repeat rewrite FcanonicFnormalizeEq; auto with zarith.
2: left; auto.
2: left; auto.
apply Rle_trans with (powerRZ radix (Fexp uh+1));
  [apply Rle_powerRZ; try apply Rle_IZR; auto with real zarith; generalize LeExp1; auto with zarith|
      rewrite powerRZ_add; try apply IZR_neq; auto with real zarith; simpl; right;ring].
apply Rlt_le; unfold FtoRradix.
apply RoundedModeUlp with (Closest bo radix); auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
apply Z.le_trans with (Fexp (Float (Fnum z) (Fexp z+1))); auto with zarith.
apply Fcanonic_Rle_Zle with radix bo p; auto with zarith.
left; auto.
elim Nz; intros A1 A2; elim A1; intros.
left; split;[split|idtac];simpl; auto with zarith.
fold FtoRradix; apply Rle_trans with (radix*Rabs z)%R.
2: unfold FtoRradix; repeat rewrite <- Fabs_correct; auto; unfold FtoR; simpl.
2: rewrite powerRZ_add; try apply IZR_neq; auto with real zarith; simpl; right; ring.
apply Rmult_le_reg_l with (1-(1+radix)*powerRZ radix (Z.succ (-p)))%R.
assert ((1+radix) * powerRZ radix (Z.succ (-p)) < 1)%R; auto with real.
2: apply Rlt_Rminus; auto.
apply Rlt_le_trans with ((radix*radix)*powerRZ radix (Z.succ (-p)))%R;
   [apply Rmult_lt_compat_r; auto with real zarith|idtac].
apply powerRZ_lt, Rlt_IZR; omega.
apply Rlt_le_trans with (radix+radix)%R; auto with real.
apply Rplus_lt_compat_r, Rlt_IZR; auto with zarith.
apply Rle_trans with (2*radix)%R; [right; ring|apply Rmult_le_compat_r; auto with real zarith]; apply Rle_IZR; omega.
apply Rle_trans with (powerRZ radix (2+Z.succ (-p)));
  [right; rewrite powerRZ_add; try apply IZR_neq; auto with real zarith; simpl; ring|idtac].
apply Rle_trans with (powerRZ radix 0); auto with real zarith.
apply Rle_powerRZ; try apply Rle_IZR; auto with zarith.
apply Rle_trans with (Rabs uh-(1+radix)*(Rabs uh*powerRZ radix (Z.succ (-p))))%R;
    [right; ring|idtac].
apply Rle_trans with (Rabs uh-(1+radix)*Fulp bo radix p uh)%R;
  [unfold Rminus;apply Rplus_le_compat_l|idtac].
apply Ropp_le_contravar; apply Rmult_le_compat_l; auto with real.
apply Rle_trans with (1+0)%R; try apply Rplus_le_compat_l; auto with real.
apply Rle_trans with 1%R; auto with real.
apply Rle_IZR; omega.
unfold FtoRradix; apply FulpLe2; auto with zarith.
elim uhDef; auto.
rewrite FcanonicFnormalizeEq; auto with zarith; left; auto.
apply Rle_trans with (Rabs z+Fulp bo radix p z)%R.
apply Rplus_le_reg_l with (-Rabs z+ (1+radix) * Fulp bo radix p uh)%R.
ring_simplify.
apply Rle_trans with (Rabs uh - Rabs z)%R;[right; ring|idtac].
apply Rle_trans with (Rabs (uh-z));[apply Rabs_triang_inv|idtac].
apply Rle_trans with (1:=H); right; ring.
apply Rle_trans with (Rabs z+(Rabs z*powerRZ radix (Z.succ (-p))))%R;
  [apply Rplus_le_compat_l|idtac].
unfold FtoRradix; apply FulpLe2; auto with zarith.
elim zDef; auto.
rewrite FcanonicFnormalizeEq; auto with zarith; left; auto.
rewrite <- Rmult_assoc.
apply Rle_trans with ((1+ powerRZ radix (Z.succ (- p)))*Rabs z)%R;
  [right; ring|apply Rmult_le_compat_r; auto with real].
apply Rabs_pos.
apply Rplus_le_reg_l with (-1+(1+radix)*radix*powerRZ radix (Z.succ (- p)))%R.
apply Rle_trans with ((1+radix+radix*radix)*powerRZ radix (Z.succ (- p)))%R;[right; ring|idtac].
apply Rle_trans with (radix-1)%R;[idtac|right; ring].
apply Rle_trans with ((radix*radix*radix)*powerRZ radix (Z.succ (-p)))%R;
   [apply Rmult_le_compat_r; auto with real zarith|idtac].
apply powerRZ_le, Rlt_IZR; omega.
apply Rle_trans with (radix+radix+radix*radix)%R; auto with real.
apply Rplus_le_compat_r, Rplus_le_compat_r, Rle_IZR; omega.
apply Rle_trans with ((2+radix)*radix)%R;
   [right; ring|apply Rmult_le_compat_r; auto with real zarith].
apply Rle_IZR; omega.
apply Rle_trans with (radix+radix)%R; auto with real.
apply Rplus_le_compat_r, Rle_IZR; omega.
apply Rle_trans with (2*radix)%R;
   [right; ring|apply Rmult_le_compat_r; auto with real zarith]; apply Rle_IZR; omega.
apply Rle_trans with (powerRZ radix (3+Z.succ (-p)));
  [right; rewrite powerRZ_add; try apply IZR_neq; auto with real zarith; simpl; ring|idtac].
apply Rle_trans with (powerRZ radix 0); auto with real zarith.
unfold Z.succ; apply Rle_powerRZ; try apply Rle_IZR; auto with real zarith.
assert (2 <= radix)%R by (apply Rle_IZR; auto with zarith).
simpl; apply Rplus_le_reg_l with 1%R; auto with real.
ring_simplify (1 + (radix - 1))%R; auto with real.
Qed.

Lemma LeExp3: (Fexp ph = Fexp uh+1)%Z -> (Fexp uh = Fexp z+1)%Z -> False.
intros.
absurd (powerRZ radix (p+Fexp z) <= Rabs z)%R.
apply Rlt_not_le.
unfold FtoRradix; rewrite <- Fabs_correct; auto; unfold FtoR; simpl.
rewrite powerRZ_add; try apply IZR_neq; auto with real zarith; apply Rmult_lt_compat_r; auto with real zarith.
apply powerRZ_lt, Rlt_IZR; omega.
elim Nz; intros A1 A2; elim A1; intros.
apply Rlt_le_trans with (IZR (Zpos (vNum bo))); try apply Rlt_IZR; auto with zarith real.
right; rewrite pGivesBound; rewrite Zpower_nat_Z_powerRZ; auto with real.
apply Rle_trans with (FtoRradix (Float 1 (p+Fexp z))).
right; unfold FtoRradix, FtoR; simpl; ring.
unfold FtoRradix; apply RoundAbsMonotonel with bo p (Closest bo radix) (a*x+b)%R; auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
elim Nz; intros A1 A2; elim A1; intros; split; simpl; auto with zarith.
apply vNumbMoreThanOne with radix p; auto with zarith.
apply Rle_trans with (powerRZ radix (p+Fexp z));[right; unfold FtoRradix, FtoR; simpl; ring|idtac].
replace (a*x+b)%R with (a*x-(-b))%R;[idtac|ring].
apply Rle_trans with (Rabs (a*x)-Rabs (-b))%R;[rewrite Rabs_Ropp|apply Rabs_triang_inv].
apply Rle_trans with ((powerRZ radix (p-1) -/(2*radix))*powerRZ radix (Fexp ph)
     -(powerRZ radix p-1)*powerRZ radix (Fexp b))%R.
apply Rle_trans with ((powerRZ radix (p-1) -/(2*radix))*powerRZ radix (Fexp z+2)
     -(powerRZ radix p-1)*powerRZ radix (Fexp z))%R.
apply Rle_trans with ((powerRZ radix (p + 1) - radix/ 2 ) * powerRZ radix (Fexp z) -
    (powerRZ radix p-1) * powerRZ radix (Fexp z))%R;
   [idtac|unfold Rminus; apply Rplus_le_compat_r].
rewrite powerRZ_add; try apply IZR_neq; auto with real zarith.
apply Rle_trans with ((powerRZ radix (p + 1) - radix / 2-powerRZ radix p+1)*
   powerRZ radix (Fexp z))%R;[apply Rmult_le_compat_r; auto with real zarith|right; ring].
apply powerRZ_le, Rlt_IZR; omega.
case (Zle_lt_or_eq 2 radix); auto with zarith; intros I.
assert (3 <= radix)%Z; auto with zarith.
assert (3 <= radix)%R; try apply Rle_IZR; auto with real zarith.
apply Rplus_le_reg_l with (radix/2-1-powerRZ radix p)%R.
ring_simplify.
unfold Rminus; rewrite Rplus_comm.
apply Rle_trans with (0+radix/2)%R; auto with real.
apply Rle_trans with (0+radix*1)%R; unfold Rdiv; auto with real.
apply Rplus_le_compat_l; apply Rmult_le_compat_l; try apply Rle_IZR; auto with real zarith.
apply Rle_trans with (/1)%R; auto with real.
apply Rle_trans with (powerRZ radix 1);[simpl; right; ring|idtac].
apply Rle_trans with (powerRZ radix p);auto with real zarith.
apply Rle_powerRZ; try apply Rle_IZR; omega.
apply Rle_trans with (- 2*powerRZ radix p + powerRZ radix p * (3))%R;
  [right; ring|auto with real].
rewrite powerRZ_add; auto with real zarith.
apply Rplus_le_compat_l; apply Rmult_le_compat_l; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
simpl; ring_simplify (radix*1)%R; auto with real.
apply IZR_neq; omega.
right; rewrite <- I; rewrite powerRZ_add; auto with real zarith; simpl.
field; auto with real.
unfold Zminus; repeat rewrite powerRZ_add; try apply IZR_neq; auto with real zarith; simpl.
ring_simplify (radix*1)%R; right; field; auto with real.
apply IZR_neq; omega.
unfold Rminus; apply Rplus_le_compat.
replace (Fexp ph) with (Fexp z+2)%Z; auto with real zarith.
apply Ropp_le_contravar; apply Rmult_le_compat_l.
assert (1 <= powerRZ radix p)%R; auto with real.
apply Rle_trans with (powerRZ radix 0); auto with real zarith.
apply Rle_powerRZ; try apply Rle_IZR; omega.
apply Rplus_le_reg_l with 1%R.
ring_simplify (1+0)%R; apply Rle_trans with (1:=H1); right; ring.
apply Rle_powerRZ; try apply Rle_IZR; auto with real zarith.
apply Z.le_trans with (Fexp uh -1)%Z; auto with zarith.
case (Zle_or_lt (Fexp b) (Fexp uh-1)); auto with zarith; intros.
absurd (FtoRradix ul=0); auto with real.
elim errorBoundedPlus with bo radix p ph b uh; auto with zarith.
2: elim phDef; auto.
fold FtoRradix; intros ul' T; elim T; intros H2 T'; elim T'; intros; clear T T'.
rewrite ulDef; rewrite <- H2.
absurd (Fexp uh < Fexp uh)%Z; auto with zarith.
apply Z.le_lt_trans with (Fexp ul'); auto with zarith.
apply ClosestErrorExpStrict with bo radix p (ph+b)%R; auto with zarith.
elim uhDef; auto.
fold FtoRradix; rewrite H2; rewrite <- ulDef; auto with real.
unfold Rminus; apply Rplus_le_compat.
apply RleRoundedAbs; auto.
assert (-dExp bo <= Fexp uh)%Z; auto with zarith.
elim uhDef; intros I1 I2; elim I1; auto.
apply Ropp_le_contravar; unfold FtoRradix; rewrite <- Fabs_correct; auto.
unfold FtoR; simpl; apply Rmult_le_compat_r; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
elim Fb; intros.
apply Rle_trans with (Z.pred (Zpos (vNum bo))); try apply Rle_IZR; auto with real zarith.
right; unfold Z.pred, Zminus; rewrite plus_IZR; rewrite pGivesBound; simpl.
rewrite Zpower_nat_Z_powerRZ; auto with real.
Qed.

Lemma LeExp: (powerRZ radix (Fexp ph)+powerRZ radix (Fexp uh) <= 2*powerRZ radix (Fexp z+1))%R.
apply Rle_trans with (powerRZ radix (Fexp z + 1)+
   powerRZ radix (Fexp z + 1))%R;[idtac|right; ring].
apply Rplus_le_compat; apply Rle_powerRZ; try apply Rle_IZR; auto with real zarith.
generalize LeExp1; generalize LeExp2; generalize LeExp3; intros.
case (Zle_lt_or_eq _ _ H0); intros; auto with zarith.
generalize LeExp1; generalize LeExp2; generalize LeExp3; intros; auto.
Qed.

Lemma vLe_aux: (Rabs (pl+ul) <= powerRZ radix (Fexp z)*radix)%R.
apply Rle_trans with (powerRZ radix (Fexp z+1));
  [idtac|rewrite powerRZ_add; try apply IZR_neq; auto with real zarith; simpl; right; ring].
apply Rle_trans with (Rabs pl+Rabs ul)%R;[apply Rabs_triang|idtac].
apply Rmult_le_reg_l with 2%R; auto with real.
apply Rle_trans with (powerRZ radix (Fexp ph)+powerRZ radix (Fexp uh))%R;
  [idtac|apply LeExp].
apply Rle_trans with (INR 2*Rabs pl+INR 2*Rabs ul)%R;[simpl; right; ring|idtac].
apply Rplus_le_compat.
apply Rle_trans with (Fulp bo radix p ph).
rewrite plDef; unfold FtoRradix; apply ClosestUlp; auto with zarith.
rewrite CanonicFulp; auto with zarith;[unfold FtoR; simpl; right; ring|left; auto].
apply Rle_trans with (Fulp bo radix p uh).
rewrite ulDef; unfold FtoRradix; apply ClosestUlp; auto with zarith.
rewrite CanonicFulp; auto with zarith;[unfold FtoR; simpl; right; ring|left; auto].
Qed.

Lemma vLe: (Rabs v <= powerRZ radix (Fexp z)*radix)%R.
assert (powerRZ radix (Fexp z) * radix=Float radix (Fexp z))%R.
unfold FtoRradix, FtoR; simpl; ring.
rewrite H.
unfold FtoRradix; apply RoundAbsMonotoner with bo p (Closest bo radix) (pl+ul)%R; auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
elim zDef; intros I1 I2; elim I1; intros; split; simpl; auto with zarith.
rewrite Z.abs_eq; auto with zarith.
apply Z.lt_trans with (pPred (vNum bo));
 [apply pPredMoreThanRadix with p|unfold pPred]; auto with zarith.
fold FtoRradix; rewrite <- H; apply vLe_aux.
Qed.

Lemma tLe: (Rabs t <= powerRZ radix (Fexp z)*(radix+1))%R.
replace (FtoRradix t) with (uh-z)%R.
replace (uh-z)%R with ((a*x+b-z)+(-(a*x-ph)+-((ph+b)-uh)))%R;[idtac|ring].
apply Rle_trans with (Rabs (a*x+b-z)+Rabs ( - (a * x - ph) + - (ph + b - uh)))%R;
   [apply Rabs_triang|idtac].
apply Rle_trans with (powerRZ radix (Fexp z)+powerRZ radix (Fexp z)*radix)%R;[idtac|right; ring].
apply Rplus_le_compat.
apply Rle_trans with (Fulp bo radix p z).
apply Rlt_le; unfold FtoRradix; apply RoundedModeUlp with (Closest bo radix); auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
rewrite CanonicFulp; auto with zarith;[unfold FtoR; simpl; right; ring|left; auto].
apply Rle_trans with (powerRZ radix (Fexp z+1));
  [idtac|rewrite powerRZ_add; try apply IZR_neq; auto with real zarith; simpl; right; ring].
apply Rle_trans with (Rabs (-(a*x-ph))+Rabs (-(ph+b-uh)))%R;[apply Rabs_triang|idtac].
rewrite Rabs_Ropp; rewrite Rabs_Ropp.
apply Rmult_le_reg_l with 2%R; auto with real.
apply Rle_trans with (powerRZ radix (Fexp ph)+powerRZ radix (Fexp uh))%R;
  [idtac|apply LeExp].
apply Rle_trans with (INR 2*Rabs (a*x-ph)+INR 2*Rabs (ph+b-uh))%R;[simpl; right; ring|idtac].
apply Rplus_le_compat.
apply Rle_trans with (Fulp bo radix p ph).
unfold FtoRradix; apply ClosestUlp; auto with zarith.
rewrite CanonicFulp; auto with zarith;[unfold FtoR; simpl; right; ring|left; auto].
apply Rle_trans with (Fulp bo radix p uh).
unfold FtoRradix; apply ClosestUlp; auto with zarith.
rewrite CanonicFulp; auto with zarith;[unfold FtoR; simpl; right; ring|left; auto].
elim tBounded with bo radix p a x b ph pl uh z; auto with zarith.
fold FtoRradix; intros t' T; elim T; intros; clear T.
rewrite <- H0; unfold FtoRradix.
apply RoundedModeProjectorIdemEq with bo p (Closest bo radix); auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
fold FtoRradix; rewrite H0; auto.
Qed.

Lemma wLe: (Rabs w <= powerRZ radix (Fexp z)*(2*radix+1))%R.
assert (powerRZ radix (Fexp z) * (2*radix+1)=Float (2*radix+1) (Fexp z))%R.
apply trans_eq with (powerRZ radix (Fexp z) * (2 * radix + 1)%Z)%R.
rewrite plus_IZR; rewrite mult_IZR; simpl; ring.
unfold FtoRradix, FtoR; simpl; ring.
rewrite H.
unfold FtoRradix; apply RoundAbsMonotoner with bo p (Closest bo radix) (t+v)%R; auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
elim zDef; intros I1 I2; elim I1; intros; split; [idtac|simpl; auto with zarith].
apply Z.le_lt_trans with (Z.abs (2*radix+1)); auto with zarith.
rewrite Z.abs_eq; auto with zarith.
apply Z.lt_le_trans with (2*radix+radix)%Z; auto with zarith.
apply Z.le_trans with (3*radix)%Z; auto with zarith.
rewrite pGivesBound; apply Z.le_trans with (Zpower_nat radix 3); auto with zarith.
unfold Zpower_nat; simpl (nat_rect _ _ _ _)%Z.
rewrite Zmult_comm; apply Zmult_le_compat_l; auto with zarith.
ring_simplify (radix*1)%Z; apply Z.le_trans with (2*2)%Z; auto with zarith.
apply Zmult_le_compat; auto with zarith.
apply Zpower_nat_monotone_le; omega.
fold FtoRradix; rewrite <- H.
apply Rle_trans with (Rabs t+Rabs v)%R;[apply Rabs_triang|idtac].
generalize tLe; generalize vLe; intros.
apply Rle_trans with ( powerRZ radix (Fexp z) * (radix + 1)+powerRZ radix (Fexp z) * radix)%R;
  [auto with real|right; ring].
Qed.

Theorem ErrFmaApprox_2_aux: (Rabs (z+w-(a*x+b)) <= (3*radix/2+/2)*powerRZ radix (2-2*p)*Rabs z)%R.
replace (z+w-(a*x+b))%R with ((t-(uh-z))+-(t+v-w)+-(pl+ul-v))%R;
  [idtac|rewrite ulDef; rewrite plDef; ring].
pattern (FtoRradix t) at 1; replace (FtoRradix t) with (uh-z)%R.
2: elim tBounded with bo radix p a x b ph pl uh z; auto with zarith.
2: fold FtoRradix; intros t' T; elim T; intros; clear T.
2: rewrite <- H0; unfold FtoRradix.
2: apply RoundedModeProjectorIdemEq with bo p (Closest bo radix); auto with zarith.
2: apply ClosestRoundedModeP with p; auto with zarith.
2: fold FtoRradix; rewrite H0; auto.
replace (uh - z - (uh - z) + - (t + v - w) + - (pl + ul - v))%R with
   (-((t + v - w)+(pl + ul - v)))%R;[rewrite Rabs_Ropp|ring].
apply Rle_trans with (Rabs (t+v-w)+Rabs (pl+ul-v))%R;[apply Rabs_triang|idtac].
apply Rmult_le_reg_l with 2%R; auto with real.
apply Rle_trans with (INR 2*Rabs (t + v - w) + INR 2*Rabs (pl + ul - v))%R;[right; simpl; ring|idtac].
apply Rle_trans with ((Fulp bo radix p w)+(Fulp bo radix p v))%R.
apply Rplus_le_compat; unfold FtoRradix; apply ClosestUlp; auto with zarith.
apply Rle_trans with (Rabs w*powerRZ radix (Z.succ (-p))+Rabs v*powerRZ radix (Z.succ (-p)))%R.
apply Rplus_le_compat; unfold FtoRradix; apply FulpLe2; auto with zarith.
elim wDef; auto.
rewrite FcanonicFnormalizeEq; auto with zarith; left; auto.
elim vDef; auto.
rewrite FcanonicFnormalizeEq; auto with zarith; left; auto.
apply Rle_trans with (powerRZ radix (Z.succ (-p))*(Rabs w+Rabs v))%R;[right; ring|idtac].
generalize wLe; generalize vLe; intros.
apply Rle_trans with (powerRZ radix (Z.succ (-p))*(powerRZ radix (Fexp z) * (2 * radix + 1) +
         powerRZ radix (Fexp z) *radix))%R; [apply Rmult_le_compat_l; auto with real zarith|idtac].
apply powerRZ_le, Rlt_IZR; omega.
apply Rle_trans with ((3*radix+1)*(powerRZ radix (Z.succ (- p))*(powerRZ radix (Fexp z))))%R;
   [right; ring|idtac].
apply Rle_trans with ((3 * radix + 1) * (powerRZ radix (2- 2*p) *Rabs z))%R;
   [idtac|right; field; auto with real].
apply Rmult_le_compat_l; auto with real zarith.
apply Rle_trans with (0+1)%R; auto with real.
assert (0 < 3*radix)%R; auto with real.
apply Rmult_lt_0_compat; try apply Rlt_IZR; auto with real zarith.
replace (2-2*p)%Z with (Z.succ (-p)+Z.succ (-p))%Z;[idtac|unfold Z.succ; ring].
rewrite powerRZ_add; try apply IZR_neq; auto with real zarith.
rewrite Rmult_assoc; apply Rmult_le_compat_l; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
apply Rle_trans with (Fulp bo radix p z).
unfold Fulp;rewrite FcanonicFnormalizeEq; auto with zarith real; left; auto.
unfold FtoRradix; rewrite Rmult_comm; apply FulpLe2; auto with zarith.
elim zDef; auto.
rewrite FcanonicFnormalizeEq; auto with zarith real; left; auto.
Qed.

End uhInexact.

Section uhInexact2.
Variable bo : Fbound.
Variable radix : Z.
Variable p : nat.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.
Hypothesis radixMoreThanOne : (1 < radix)%Z.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ radixMoreThanOne).

Hypothesis pGivesBound : Zpos (vNum bo) = Zpower_nat radix p.
Hypothesis precisionGreaterThanOne : 4 <= p.

Variables a x b ph pl uh ul z t v w:float.

Hypothesis Fb: Fbounded bo b.
Hypothesis Fa: Fbounded bo a.
Hypothesis Fx: Fbounded bo x.
Hypothesis Cb: Fcanonic radix bo b.

Hypothesis Nph: Fnormal radix bo ph \/ (FtoRradix ph=0).
Hypothesis Nuh: Fnormal radix bo uh \/ (FtoRradix uh=0).
Hypothesis Nz: Fnormal radix bo z \/ (FtoRradix z =0).
Hypothesis Nw: Fnormal radix bo w \/ (FtoRradix w =0).
Hypothesis Nv: Fnormal radix bo v \/ (FtoRradix v =0).

Hypothesis Exp1: (- dExp bo <= Fexp a+Fexp x)%Z.

Hypothesis zDef : Closest bo radix (a*x+b)%R z.
Hypothesis phDef: Closest bo radix (a*x)%R ph.
Hypothesis plDef: (FtoRradix pl=a*x-ph)%R.
Hypothesis uhDef: Closest bo radix (ph+b)%R uh.
Hypothesis ulDef: (FtoRradix ul=ph+b-uh)%R.
Hypothesis tDef : Closest bo radix (uh-z)%R t.
Hypothesis vDef : Closest bo radix (pl+ul)%R v.
Hypothesis wDef : Closest bo radix (t+v)%R w.

Hypothesis Case2: ~(FtoRradix ul=0)%R.

Theorem ErrFmaApprox_2: (Rabs (z+w-(a*x+b)) <= (3*radix/2+/2)*powerRZ radix (2-2*p)*Rabs z)%R.
case Nv; intros I5.
case Nz; intros I3.
case Nph; intros I1.
case Nuh; intros I2.
case Nw; intros I4.
unfold FtoRradix; apply ErrFmaApprox_2_aux with bo ph pl uh ul t v; auto.
replace (z + w - (a * x + b))%R with (-(ul+pl-v))%R.
apply Rmult_le_reg_l with (INR 2); auto with real zarith; rewrite Rabs_Ropp.
apply Rle_trans with (Fulp bo radix p v).
unfold FtoRradix; apply ClosestUlp; auto with zarith.
rewrite Rplus_comm; auto with real.
apply Rle_trans with (Rabs v * powerRZ radix (Z.succ (- p)))%R;
 [unfold FtoRradix; apply FulpLe2; auto with zarith|idtac].
elim vDef; auto.
rewrite FcanonicFnormalizeEq; auto with zarith; left; auto.
assert (Rabs v <= powerRZ radix (Fexp z) * radix)%R.
unfold FtoRradix; apply vLe with bo p a x b ph pl uh ul; auto.
apply Rle_trans with ((powerRZ radix (Fexp z) * radix)* powerRZ radix (Z.succ (- p)))%R;
   [apply Rmult_le_compat_r; auto with real zarith|idtac].
apply powerRZ_le, Rlt_IZR; omega.
apply Rle_trans with ((Rabs z * powerRZ radix (Z.succ (- p)))* radix * powerRZ radix (Z.succ (- p)))%R.
apply Rmult_le_compat_r; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
apply Rmult_le_compat_r; auto with real zarith.
apply Rle_IZR; omega.
apply Rle_trans with (Fulp bo radix p z).
unfold Fulp; rewrite FcanonicFnormalizeEq; auto with zarith real; left; auto.
unfold FtoRradix; apply FulpLe2; auto with zarith.
elim zDef; auto.
rewrite FcanonicFnormalizeEq; auto with zarith real; left; auto.
apply Rle_trans with (radix*powerRZ radix (2 - 2 * p) * Rabs z)%R.
replace (2-2*p)%Z with (1+1-p-p)%Z;[idtac|ring].
unfold Z.succ, Zminus; repeat rewrite powerRZ_add; try apply IZR_neq; auto with real zarith; simpl.
ring_simplify (radix*1)%R; right; ring.
rewrite <- Rmult_assoc; apply Rmult_le_compat_r; auto with real.
apply Rabs_pos.
rewrite <- Rmult_assoc; apply Rmult_le_compat_r; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
apply Rle_trans with (3*radix+1)%R;[idtac|right; simpl; field; auto with real].
apply Rle_trans with (1*radix+0)%R;[right; ring|apply Rplus_le_compat; auto with real zarith].
apply Rmult_le_compat_r; try apply Rle_IZR; omega.
rewrite ulDef; rewrite plDef.
assert (t+v=0)%R.
apply ClosestZero1 with bo radix p w (Fplus radix t v); auto with zarith.
rewrite Fplus_correct; fold FtoRradix; auto.
simpl; apply Zmin_Zle; auto with zarith.
elim tDef; intros Y1 Y2; elim Y1; auto.
elim vDef; intros Y1 Y2; elim Y1; auto.
cut (uh-z=t)%R;[intros H'|idtac].
rewrite I4; simpl; rewrite <- H; rewrite <- H'; ring.
elim tBounded with bo radix p a x b ph pl uh z; auto with zarith.
fold FtoRradix; intros t' T; elim T; intros; clear T.
rewrite <- H1; unfold FtoRradix.
apply RoundedModeProjectorIdemEq with bo p (Closest bo radix); auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
fold FtoRradix; rewrite H1; auto.
assert (ph+b=0)%R.
apply ClosestZero1 with bo radix p uh (Fplus radix ph b); auto with zarith.
rewrite Fplus_correct; fold FtoRradix; auto.
simpl; apply Zmin_Zle; auto with zarith.
elim phDef; intros Y1 Y2; elim Y1; auto.
apply Fb.
absurd (FtoRradix ul=0)%R; auto.
rewrite ulDef; rewrite I2; rewrite H; simpl; ring.
absurd (FtoRradix ul=0)%R; auto.
rewrite ulDef; rewrite I1.
cut (FtoRradix b=uh)%R; [intros Y; rewrite Y; simpl; ring|idtac].
unfold FtoRradix; apply RoundedModeProjectorIdemEq with bo p (Closest bo radix); auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
replace (FtoR radix b) with (ph+b)%R; auto; rewrite I1; auto with real.
assert (FtoRradix uh=0)%R.
unfold FtoRradix; apply ClosestZero2 with bo p (ph+b)%R; auto with real zarith.
cut (FtoRradix ph= (Fopp b));
    [intros I;rewrite I; unfold FtoRradix; rewrite Fopp_correct; ring|idtac].
unfold FtoRradix; apply sym_eq.
apply RoundedModeProjectorIdemEq with bo p (Closest bo radix); auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
apply oppBounded; auto.
replace (FtoR radix (Fopp b)) with (a*x)%R; auto.
rewrite Fopp_correct; fold FtoRradix; assert (a*x+b=0)%R; auto with real.
apply ClosestZero1 with bo radix p z (Fplus radix (Fmult a x) b); auto with zarith.
rewrite Fplus_correct; auto; rewrite Fmult_correct; fold FtoRradix; auto.
simpl; apply Zmin_Zle; auto with zarith.
apply Fb.
apply Rplus_eq_reg_l with (FtoRradix b); rewrite Rplus_comm; rewrite H;ring.
assert (ph+b=0)%R.
apply ClosestZero1 with bo radix p uh (Fplus radix ph b); auto with zarith.
rewrite Fplus_correct; fold FtoRradix; auto.
simpl; apply Zmin_Zle; auto with zarith.
elim phDef; intros Y1 Y2; elim Y1; auto.
apply Fb.
absurd (FtoRradix ul=0)%R; auto.
rewrite ulDef; rewrite H0; rewrite H; ring.
replace (z+w-(a*x+b))%R with 0%R.
rewrite Rabs_R0; apply Rle_trans with (0*Rabs z)%R;
   [right; ring|apply Rmult_le_compat_r; auto with real].
apply Rabs_pos.
apply Rlt_le; apply Rmult_lt_0_compat; auto with real zarith.
apply Rle_lt_trans with (0+0)%R; auto with real; apply Rplus_lt_compat; auto with real.
unfold Rdiv;repeat apply Rmult_lt_0_compat; auto with real zarith.
apply Rlt_IZR; omega.
apply powerRZ_lt, Rlt_IZR; omega.
apply sym_eq; apply trans_eq with (z+w-uh-(pl+ul))%R;[rewrite plDef; rewrite ulDef; ring|idtac].
assert (pl+ul=0)%R.
elim errorBoundedPlus with bo radix p ph b uh; auto with zarith.
2: elim phDef; auto.
fold FtoRradix; intros pl' T; elim T; intros Y1 T'; elim T'; intros Y2 Y3; clear T T'.
elim errorBoundedMult with bo radix p (Closest bo radix) a x ph; auto with zarith.
2: apply ClosestRoundedModeP with p; auto with zarith.
fold FtoRradix; intros ul' T; elim T; intros Y1' T'; elim T'; intros Y2' Y3'; clear T T'.
apply ClosestZero1 with bo radix p v (Fplus radix pl' ul'); auto with zarith.
rewrite Fplus_correct; fold FtoRradix; auto.
rewrite ulDef; rewrite plDef; rewrite Y1; rewrite Y1'; ring.
simpl; apply Zmin_Zle; auto with zarith.
apply Y2.
rewrite H.
replace (FtoRradix w) with (FtoRradix t).
replace (FtoRradix t) with (uh-z)%R; [ring|idtac].
elim tBounded with bo radix p a x b ph pl uh z; auto with zarith.
fold FtoRradix; intros t' T; elim T; intros; clear T.
rewrite <- H1; unfold FtoRradix.
apply RoundedModeProjectorIdemEq with bo p (Closest bo radix); auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
fold FtoRradix; rewrite H1; auto.
unfold FtoRradix; apply RoundedModeProjectorIdemEq with bo p (Closest bo radix); auto with zarith.
apply ClosestRoundedModeP with p; auto with zarith.
elim tDef; auto.
replace (FtoR radix t) with (t+v)%R; auto; rewrite I5; auto with real.
Qed.

End uhInexact2.

Section Total.
Variable bo : Fbound.
Variable radix : Z.
Variable p : nat.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.
Hypothesis radixMoreThanOne : (1 < radix)%Z.

Let radixMoreThanZERO := Zlt_1_O _ (Zlt_le_weak _ _ radixMoreThanOne).

Hypothesis pGivesBound : Zpos (vNum bo) = Zpower_nat radix p.
Hypothesis precisionGreaterThanOne : 4 <= p.

Variables a x b ph pl uh ul z t v w:float.

Hypothesis Fb: Fbounded bo b.
Hypothesis Fa: Fbounded bo a.
Hypothesis Fx: Fbounded bo x.

Hypothesis Nph: Fnormal radix bo ph \/ (FtoRradix ph=0).
Hypothesis Nuh: Fnormal radix bo uh \/ (FtoRradix uh=0).
Hypothesis Nz: Fnormal radix bo z \/ (FtoRradix z =0).
Hypothesis Nw: Fnormal radix bo w \/ (FtoRradix w =0).
Hypothesis Nv: Fnormal radix bo v \/ (FtoRradix v =0).

Hypothesis Exp1: (- dExp bo <= Fexp a+Fexp x)%Z.

Hypothesis zDef : Closest bo radix (a*x+b)%R z.
Hypothesis phDef: Closest bo radix (a*x)%R ph.
Hypothesis plDef: (FtoRradix pl=a*x-ph)%R.
Hypothesis uhDef: Closest bo radix (ph+b)%R uh.
Hypothesis ulDef: (FtoRradix ul=ph+b-uh)%R.
Hypothesis tDef : Closest bo radix (uh-z)%R t.
Hypothesis vDef : Closest bo radix (pl+ul)%R v.
Hypothesis wDef : Closest bo radix (t+v)%R w.

Theorem ErrFmaApprox: (Rabs (z+w-(a*x+b)) <= (3*radix/2+/2)*powerRZ radix (2-2*p)*Rabs z)%R.
case (Req_dec ul 0); intros.
elim errorBoundedMult with bo radix p (Closest bo radix) a x ph; auto with zarith.
2: apply ClosestRoundedModeP with p; auto with zarith.
fold FtoRradix; intros pl' T; elim T; intros M1 T'; elim T'; intros M2 M3; clear T T'.
unfold FtoRradix; rewrite <- FnormalizeCorrect with radix bo p b; auto.
apply ErrFmaApprox_1 with bo ph pl' uh ul t v; auto with zarith;
   try rewrite FnormalizeCorrect; auto with real.
apply FnormalizeBounded; auto with zarith.
apply FnormalizeCanonic; auto with zarith.
fold FtoRradix; rewrite M1; rewrite <- plDef; auto.
unfold FtoRradix; rewrite <- FnormalizeCorrect with radix bo p b; auto.
apply ErrFmaApprox_2 with bo ph pl uh ul t v; auto;
    try rewrite FnormalizeCorrect; auto with real.
apply FnormalizeBounded; auto with zarith.
apply FnormalizeCanonic; auto with zarith.
Qed.

End Total.


Section EFast.
Variable b : Fbound.
Variable precision : nat.

Let radix := 2%Z.

Let TMTO : (1 < radix)%Z := TwoMoreThanOne.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.

Hypothesis precisionGreaterThanOne : 1 < precision.
Hypothesis pGivesBound : Zpos (vNum b) = Zpower_nat radix precision.
Variable Iplus : float -> float -> float.
Hypothesis
  IplusCorrect :
    forall p q : float,
    Fbounded b p -> Fbounded b q -> Closest b radix (p + q) (Iplus p q).
Hypothesis
  IplusComp :
    forall p q r s : float,
    Fbounded b p ->
    Fbounded b q ->
    Fbounded b r ->
    Fbounded b s -> p = r :>R -> q = s :>R -> Iplus p q = Iplus r s :>R.
Hypothesis IplusSym : forall p q : float, Iplus p q = Iplus q p.
Hypothesis
  IplusOp : forall p q : float, Fopp (Iplus p q) = Iplus (Fopp p) (Fopp q).
Variable Iminus : float -> float -> float.
Hypothesis IminusPlus : forall p q : float, Iminus p q = Iplus p (Fopp q).

Theorem IplusOl :
 forall p q : float,
 Fbounded b p -> Fbounded b q -> p = 0%R :>R -> Iplus p q = q :>R.
intros p q H' H'0 H'1.
rewrite IplusSym.
apply (IplusOr b); auto.
Qed.

Let IminusCorrect := IminusCorrect b Iplus IplusCorrect Iminus IminusPlus.

Let IplusBounded := IplusBounded b Iplus IplusCorrect.

Let IminusBounded := IminusBounded b Iplus IplusCorrect Iminus IminusPlus.

Let IminusId := IminusId b Iplus IplusCorrect Iminus IminusPlus.

Theorem MKnuth :
 forall p q : float,
 Fbounded b p ->
 Fbounded b q ->
 Iminus (Iplus p q) p = (Iplus p q - p)%R :>R ->
 Iplus (Iminus p (Iminus (Iplus p q) (Iminus (Iplus p q) p)))
   (Iminus q (Iminus (Iplus p q) p)) = (p + q - Iplus p q)%R :>R.
intros p q H' H'0 H'1.
rewrite IplusOl; auto with zarith.
case
 (errorBoundedPlus b radix precision)
  with (p := p) (q := q) (pq := Iplus p q); auto.
intros x H'2; elim H'2; intros H'3 H'4; elim H'4; intros H'5 H'6;
 clear H'4 H'2.
unfold FtoRradix in |- *; rewrite <- H'3.
apply sym_eq; apply (ClosestIdem b); auto.
apply
 (ClosestCompatible b radix (q - Iminus (Iplus p q) p)%R (
    FtoR radix x) (Iminus q (Iminus (Iplus p q) p)));
 auto.
rewrite H'3.
rewrite H'1.
unfold FtoRradix in |- *; ring; ring.
replace 0%R with (FtoRradix (Iminus p p)).
repeat rewrite IminusPlus; apply IplusComp; repeat rewrite <- IminusPlus;
 auto with zarith.
apply oppBounded; auto.
apply oppBounded; auto.
unfold FtoRradix in |- *; repeat rewrite Fopp_correct.
replace (FtoR radix (Iminus (Iplus p q) (Iminus (Iplus p q) p))) with
 (FtoR radix p); auto.
apply (ClosestIdem b); auto.
apply
 (ClosestCompatible b radix (Iplus p q - Iminus (Iplus p q) p)%R
    (FtoR radix p) (Iminus (Iplus p q) (Iminus (Iplus p q) p)));
 auto.
rewrite H'1; auto.
unfold FtoRradix in |- *; ring; ring.
apply IminusId; auto.
Qed.

Theorem IplusCorrectEq :
 forall (p q pq : float) (r : R),
 Fbounded b p ->
 Fbounded b q ->
 Fbounded b pq -> r = pq :>R -> (p + q)%R = pq :>R -> Iplus p q = r :>R.
intros p q pq r H' H'0 H'1 H'2 H'3; rewrite H'2.
unfold FtoRradix in |- *; apply sym_eq; apply (ClosestIdem b radix); auto.
apply (ClosestCompatible b radix (p + q)%R pq (Iplus p q)); auto.
Qed.

Theorem IminusCorrectEq :
 forall (p q pq : float) (r : R),
 Fbounded b p ->
 Fbounded b q ->
 Fbounded b pq -> r = pq :>R -> (p - q)%R = pq :>R -> Iminus p q = r :>R.
intros p q pq r H' H'0 H'1 H'2 H'3; rewrite H'2.
unfold FtoRradix in |- *; apply sym_eq; apply (ClosestIdem b radix); auto.
apply (ClosestCompatible b radix (p - q)%R pq (Iminus p q)); auto.
Qed.

Theorem MKnuth1 :
 forall p q : float,
 Fbounded b p ->
 Fbounded b q ->
 Iminus q (Iminus (Iplus p q) p) = (q - Iminus (Iplus p q) p)%R :>R ->
 Iminus (Iplus p q) (Iminus (Iplus p q) p) =
 (Iplus p q - Iminus (Iplus p q) p)%R :>R ->
 Iplus (Iminus p (Iminus (Iplus p q) (Iminus (Iplus p q) p)))
   (Iminus q (Iminus (Iplus p q) p)) = (p + q - Iplus p q)%R :>R.
intros p q H' H'0 H'1 H'2.
cut
 (Iminus p (Iminus (Iplus p q) (Iminus (Iplus p q) p)) =
  (p - Iminus (Iplus p q) (Iminus (Iplus p q) p))%R :>R);
 [ intros H'3 | idtac ].
case
 (errorBoundedPlus b radix precision)
  with (p := p) (q := q) (pq := Iplus p q); auto.
intros pq H'4; elim H'4; intros H'5 H'6; elim H'6; intros H'7 H'8;
 clear H'6 H'4.
apply IplusCorrectEq with (pq := pq); auto.
repeat rewrite H'3; repeat rewrite H'2; repeat rewrite H'1.
unfold FtoRradix in |- *; rewrite H'5.
ring.
rewrite H'2.
case
 (errorBoundedPlus b radix precision)
  with (p := Iplus p q) (q := Fopp p) (pq := Iminus (Iplus p q) p);
 auto with zarith.
apply oppBounded; auto.
apply
 (ClosestCompatible b radix (Iplus p q - p)%R (Iplus p q + Fopp p)%R
    (Iminus (Iplus p q) p)); auto.
rewrite (Fopp_correct radix); auto.
intros pq H'3; elim H'3; intros H'4 H'5; elim H'5; intros H'6 H'7;
 clear H'5 H'3.
rewrite (IminusPlus p); apply IplusCorrectEq with (pq := Fopp pq);
 auto with zarith.
apply oppBounded; auto.
apply oppBounded; auto.
unfold FtoRradix in |- *; rewrite (Fopp_correct radix); auto; rewrite H'4;
 rewrite (Fopp_correct radix); auto; ring; ring.
unfold FtoRradix in |- *; repeat rewrite (Fopp_correct radix); auto;
 rewrite H'4; rewrite (Fopp_correct radix); auto.
unfold FtoRradix in H'2; rewrite H'2; ring.
Qed.

Theorem MKnuth2 :
 forall p q : float,
 (Rabs q <= Rabs p)%R ->
 Fbounded b p ->
 Fbounded b q ->
 Iplus (Iminus p (Iminus (Iplus p q) (Iminus (Iplus p q) p)))
   (Iminus q (Iminus (Iplus p q) p)) = (p + q - Iplus p q)%R :>R.
intros p q H' H'0 H'1.
apply MKnuth; auto.
apply (MDekker b precision); auto.
Qed.

Theorem IminusOp :
 forall p q : float, Fopp (Iminus p q) = Iminus (Fopp p) (Fopp q).
intros p q; repeat rewrite IminusPlus; repeat rewrite <- IplusOp; auto.
Qed.

Theorem MKnuthOpp :
 forall p q : float,
 Iplus (Iminus p (Iminus (Iplus p q) (Iminus (Iplus p q) p)))
   (Iminus q (Iminus (Iplus p q) p)) =
 Fopp
   (Iplus
      (Iminus (Fopp p)
         (Iminus (Iplus (Fopp p) (Fopp q))
            (Iminus (Iplus (Fopp p) (Fopp q)) (Fopp p))))
      (Iminus (Fopp q) (Iminus (Iplus (Fopp p) (Fopp q)) (Fopp p)))) :>R.
intros p q; repeat rewrite <- IplusOp || rewrite <- IminusOp;
 rewrite Fopp_Fopp; auto.
Qed.

Theorem MKnuth3 :
 forall p q : float,
 (0 <= q)%R ->
 (q <= radix * - p)%R ->
 (- p <= q)%R ->
 Fbounded b p ->
 Fbounded b q ->
 Iplus (Iminus p (Iminus (Iplus p q) (Iminus (Iplus p q) p)))
   (Iminus q (Iminus (Iplus p q) p)) = (p + q - Iplus p q)%R :>R.
intros p q H' H'0 H'1 H'2 H'3.
apply MKnuth; auto.
cut (FtoRradix (Iplus p q) = (p + q)%R :>R); [ intros Eq1 | idtac ].
apply IminusCorrectEq with (pq := q); auto.
rewrite Eq1; ring.
rewrite Eq1; ring.
apply IplusCorrectEq with (pq := Fminus radix q (Fopp p)); auto.
apply Sterbenz; auto with zarith.
apply oppBounded; auto.
2: repeat rewrite (Fopp_correct radix); auto.
2: rewrite (Fminus_correct radix); repeat rewrite (Fopp_correct radix);
    unfold FtoRradix in |- *; auto; ring; ring.
2: rewrite (Fminus_correct radix); repeat rewrite (Fopp_correct radix);
    unfold FtoRradix in |- *; auto; ring; ring.
apply Rle_trans with (2 := H'1).
apply Rmult_le_reg_l with (r := IZR radix); auto with real.
rewrite <- Rmult_assoc; rewrite Rinv_r; auto with real arith;
 rewrite Rmult_1_l.
rewrite (Fopp_correct radix).
apply Rle_trans with (1 := H'1); auto.
Qed.

Theorem MKnuth4 :
 forall p q : float,
 (0 < - p)%R ->
 (0 < q)%R ->
 (radix * - p < q)%R ->
 Fbounded b p ->
 Fbounded b q ->
 Iplus (Iminus p (Iminus (Iplus p q) (Iminus (Iplus p q) p)))
   (Iminus q (Iminus (Iplus p q) p)) = (p + q - Iplus p q)%R :>R.
intros p q H' H'0 H'1 H'2 H'3.
cut (/ radix * q <= Iplus p q)%R; [ intros Rle1 | idtac ].
cut (Iplus p q <= Iminus (Iplus p q) p)%R; [ intros Rle2 | idtac ].
cut (Iminus (Iplus p q) p <= radix * Iplus p q)%R; [ intros Rle3 | idtac ].
cut (radix * Iplus p q <= radix * q)%R; [ intros Rle4 | idtac ].
apply MKnuth1; auto with zarith.
apply IminusCorrectEq with (pq := Fminus radix q (Iminus (Iplus p q) p));
 auto with zarith.
rewrite <- (Fopp_Fopp (Fminus radix q (Iminus (Iplus p q) p))).
apply oppBounded.
rewrite Fopp_Fminus.
apply Sterbenz; auto.
apply Rle_trans with (1 := Rle1); auto.
apply Rle_trans with (1 := Rle3); auto.
apply sym_eq; apply (Fminus_correct radix); auto.
apply sym_eq; apply (Fminus_correct radix); auto.
apply
 IminusCorrectEq with (pq := Fminus radix (Iplus p q) (Iminus (Iplus p q) p));
 auto.
rewrite <- (Fopp_Fopp (Fminus radix (Iplus p q) (Iminus (Iplus p q) p))).
apply oppBounded.
rewrite Fopp_Fminus.
apply Sterbenz; auto.
apply Rle_trans with (2 := Rle2); auto.
apply Rmult_le_reg_l with (r := IZR radix); auto with real.
rewrite <- Rmult_assoc; rewrite Rinv_r; auto with real arith;
 rewrite Rmult_1_l.
apply Rledouble.
apply Rle_trans with (2 := Rle1).
apply Rmult_le_reg_l with (r := IZR radix); auto with real.
rewrite <- Rmult_assoc; rewrite Rinv_r; auto with real zarith;
 rewrite Rmult_1_l; rewrite Rmult_0_r; apply Rlt_le;
 auto.
apply sym_eq; apply (Fminus_correct radix); auto.
apply sym_eq; apply (Fminus_correct radix); auto.
apply Rmult_le_reg_l with (r := (/ radix)%R); auto with real.
repeat rewrite <- Rmult_assoc; repeat rewrite Rinv_l; auto with real zarith;
 repeat rewrite Rmult_1_l; auto with real.
cut (MonotoneP radix (Closest b radix)); [ intros Cp | idtac ].
apply (Cp (p + q)%R q); auto.
apply Rplus_lt_reg_l with (r := (- q)%R); auto.
replace (- q + (p + q))%R with (- - p)%R; [ idtac | ring ].
replace (- q + q)%R with (-0)%R; [ auto with real | ring ].
apply (RoundedModeProjectorIdem b radix (Closest b radix)); auto.
apply ClosestRoundedModeP with (precision := precision); auto.
apply ClosestMonotone; auto.
apply
 (RoundedModeMult b radix)
  with (P := Closest b radix) (r := (Iplus p q - p)%R);
 auto.
apply ClosestRoundedModeP with (precision := precision); auto.
apply Rplus_le_reg_l with (r := (- Iplus p q)%R); auto.
replace (- Iplus p q + (Iplus p q - p))%R with (- p)%R; [ idtac | ring ].
replace (- FtoRradix (Iplus p q) + radix * FtoR radix (Iplus p q))%R with
 (FtoRradix (Iplus p q));
 [ idtac | unfold FtoRradix in |- *; simpl in |- *; unfold radix; ring; ring ].
rewrite <- (Fopp_correct radix); auto.
cut (MonotoneP radix (Closest b radix)); [ intros Cp | idtac ].
apply (Cp (Fopp p) (p + q)%R); auto.
rewrite (Fopp_correct radix); auto.
apply Rplus_lt_reg_l with (r := (- p)%R); auto.
replace (- p + - FtoR radix p)%R with (radix * - p)%R;
 [ idtac | simpl in |- *; fold FtoRradix; unfold radix; ring ].
replace (- p + (p + q))%R with (FtoRradix q); [ auto | simpl in |- *; ring ].
apply (RoundedModeProjectorIdem b radix (Closest b radix)); auto.
apply ClosestRoundedModeP with (precision := precision); auto.
apply oppBounded; auto.
apply ClosestMonotone; auto.
cut (MonotoneP radix (Closest b radix)); [ intros Cp | idtac ].
apply (Cp (Iplus p q) (Iplus p q - p)%R); auto.
apply Rplus_lt_reg_l with (r := (- Iplus p q)%R); auto.
replace (- Iplus p q + Iplus p q)%R with 0%R; [ auto | simpl in |- *; ring ].
replace (- Iplus p q + (Iplus p q - p))%R with (- p)%R;
 [ auto | simpl in |- *; ring ].
apply (RoundedModeProjectorIdem b radix (Closest b radix)); auto.
apply ClosestRoundedModeP with (precision := precision); auto.
apply ClosestMonotone; auto.
apply (FmultRadixInv b radix precision) with (y := (p + q)%R); auto.
apply Rmult_lt_reg_l with (r := IZR radix); auto with real.
repeat rewrite <- Rmult_assoc; repeat rewrite Rinv_r; auto with real arith;
 repeat rewrite Rmult_1_l; auto.
apply Rplus_lt_reg_l with (r := (- q)%R); auto.
apply Rplus_lt_reg_l with (r := (radix * - p)%R); auto.
replace (radix * - p + (- q + FtoR radix q))%R with (radix * - p)%R;
 [ auto | unfold FtoRradix in |- *; simpl in |- *; ring ].
replace (radix * - p + (- q + radix * (p + q)))%R with (FtoRradix q);
 [ auto | simpl in |- *; unfold radix; ring ].
Qed.

Theorem MKnuth5 :
 forall p q : float,
 (0 < p)%R ->
 (p < q)%R ->
 Fbounded b p ->
 Fbounded b q ->
 Iplus (Iminus p (Iminus (Iplus p q) (Iminus (Iplus p q) p)))
   (Iminus q (Iminus (Iplus p q) p)) = (p + q - Iplus p q)%R :>R.
intros p q H' H'0 H'1 H'2.
cut
 (Iminus (Iplus p q) (Iminus (Iplus p q) p) =
  (Iplus p q - Iminus (Iplus p q) p)%R :>R); [ intros Eq1 | idtac ].
apply MKnuth1; auto.
cut (Iminus (Iplus p q) p <= Iplus p q)%R; [ intros Rle1 | idtac ].
cut (Iplus p q <= radix * q)%R; [ intros Rle2 | idtac ].
case (Rle_or_lt q (Iminus (Iplus p q) p)); intros Rle3.
apply IminusCorrectEq with (pq := Fminus radix q (Iminus (Iplus p q) p));
 auto with zarith.
rewrite <- (Fopp_Fopp (Fminus radix q (Iminus (Iplus p q) p))).
apply oppBounded.
rewrite Fopp_Fminus.
apply Sterbenz; auto.
apply Rle_trans with (2 := Rle3); auto.
apply Rmult_le_reg_l with (r := IZR radix); auto with real.
rewrite <- Rmult_assoc; rewrite Rinv_r; auto with real arith;
 rewrite Rmult_1_l.
apply Rledouble.
apply Rlt_le; apply Rlt_trans with (1 := H'); auto.
apply Rle_trans with (1 := Rle1); auto.
apply sym_eq; apply (Fminus_correct radix); auto.
apply sym_eq; apply (Fminus_correct radix); auto.
case
 (ExactMinusInterval b radix precision)
  with
    (P := Closest b radix)
    (p := Iminus (Iplus p q) p)
    (q := Iplus p q)
    (r := q); auto.
apply ClosestRoundedModeP with (precision := precision); auto.
apply
 (RleRoundedR0 b radix precision)
  with (P := Closest b radix) (r := (Iplus p q - p)%R);
 auto.
apply ClosestRoundedModeP with (precision := precision); auto.
apply Rplus_le_reg_l with (r := FtoRradix p); auto.
replace (p + 0)%R with (FtoRradix p); [ idtac | ring ].
replace (p + (Iplus p q - p))%R with (FtoRradix (Iplus p q));
 [ idtac | ring ].
cut (MonotoneP radix (Closest b radix)); [ intros Cp | idtac ].
apply (Cp p (p + q)%R); auto.
apply Rplus_lt_reg_l with (r := (- p)%R); auto.
replace (- p + p)%R with 0%R; [ idtac | ring ].
replace (- p + (p + q))%R with (FtoRradix q);
 [ apply Rlt_trans with (1 := H'); auto | ring ].
apply (RoundedModeProjectorIdem b radix (Closest b radix)); auto.
apply ClosestRoundedModeP with (precision := precision); auto.
apply ClosestMonotone; auto.
exists (Iminus (Iplus p q) (Iminus (Iplus p q) p)); split; auto.
apply Rlt_le; auto.
cut (MonotoneP radix (Closest b radix)); [ intros Cp | idtac ].
apply (Cp q (p + q)%R); auto.
apply Rplus_lt_reg_l with (r := (- q)%R); auto.
replace (- q + q)%R with 0%R; [ idtac | ring ].
replace (- q + (p + q))%R with (FtoRradix p); [ auto | ring ].
apply (RoundedModeProjectorIdem b radix (Closest b radix)); auto.
apply ClosestRoundedModeP with (precision := precision); auto.
apply ClosestMonotone; auto.
intros x H'4; elim H'4; intros H'5 H'6; clear H'4.
apply IminusCorrectEq with (pq := x); auto.
apply (RoundedModeMult b radix) with (P := Closest b radix) (r := (p + q)%R);
 auto.
apply ClosestRoundedModeP with (precision := precision); auto.
apply Rplus_le_reg_l with (r := (- q)%R); auto.
replace (- q + (p + q))%R with (FtoRradix p); [ idtac | ring ].
replace (- FtoRradix q + radix * FtoR radix q)%R with (FtoRradix q);
 [ apply Rlt_le; auto | unfold FtoRradix in |- *; simpl in |- *; unfold radix; ring ].
cut (MonotoneP radix (Closest b radix)); [ intros Cp | idtac ].
apply (Cp (Iplus p q - p)%R (Iplus p q)); auto.
apply Rplus_lt_reg_l with (r := (- Iplus p q)%R); auto.
replace (- Iplus p q + (Iplus p q - p))%R with (- p)%R; [ idtac | ring ].
replace (- Iplus p q + Iplus p q)%R with (-0)%R; [ auto with real | ring ].
apply (RoundedModeProjectorIdem b radix (Closest b radix)); auto.
apply ClosestRoundedModeP with (precision := precision); auto.
apply ClosestMonotone; auto.
case
 (minusRoundRep b radix precision)
  with
    (P := Closest b radix)
    (p := p)
    (q := Iplus p q)
    (qmp := Iminus (Iplus p q) p); auto.
apply ClosestRoundedModeP with (precision := precision); auto.
apply Rlt_le; auto.
cut (MonotoneP radix (Closest b radix)); [ intros Cp | idtac ].
apply (Cp (FtoRradix p) (p + q)%R); auto.
apply Rplus_lt_reg_l with (r := (- p)%R); auto.
replace (- p + p)%R with 0%R; [ idtac | ring ].
replace (- p + (p + q))%R with (FtoRradix q);
 [ apply Rlt_trans with (1 := H'); auto | ring ].
apply (RoundedModeProjectorIdem b radix (Closest b radix)); auto.
apply ClosestRoundedModeP with (precision := precision); auto.
apply ClosestMonotone; auto.
intros x H'4; elim H'4; intros H'5 H'6; clear H'4; auto with zarith.
apply IminusCorrectEq with (pq := x); auto.
Qed.

Theorem MKnuth6 :
 forall p q : float,
 Iplus p q = (p + q)%R :>R ->
 Fbounded b p ->
 Fbounded b q ->
 Iplus (Iminus p (Iminus (Iplus p q) (Iminus (Iplus p q) p)))
   (Iminus q (Iminus (Iplus p q) p)) = (p + q - Iplus p q)%R :>R.
intros p q H' H'0 H'1.
apply MKnuth; auto.
apply IminusCorrectEq with (pq := q); auto.
rewrite H'; ring.
rewrite H'; ring.
Qed.

Theorem MKnuth7 :
 forall p q : float,
 (Rabs p < q)%R ->
 Fbounded b p ->
 Fbounded b q ->
 Iplus (Iminus p (Iminus (Iplus p q) (Iminus (Iplus p q) p)))
   (Iminus q (Iminus (Iplus p q) p)) = (p + q - Iplus p q)%R :>R.
intros p q H' H'0 H'1.
cut (0 < q)%R;
 [ intros Rlt1 | apply Rle_lt_trans with (2 := H'); auto with real ].
case (Rle_or_lt 0 p); intros Rle1; [ Casec Rle1; intros Rle1 | idtac ].
apply MKnuth5; auto.
rewrite <- (Rabs_right p); auto with real; apply Rle_ge; apply Rlt_le; auto.
apply MKnuth6; auto.
apply IplusCorrectEq with (pq := q); auto; rewrite <- Rle1; ring.
case (Rle_or_lt q (radix * - p)); intros Rle2.
apply MKnuth3; auto.
apply Rlt_le; auto.
rewrite <- (Rabs_left1 p); auto; apply Rlt_le; auto.
apply MKnuth4; auto.
replace 0%R with (-0)%R; [ auto with real | ring ].
apply Rabs_pos.
Qed.

Theorem Knuth :
 forall p q : float,
 Fbounded b p ->
 Fbounded b q ->
 Iplus (Iminus p (Iminus (Iplus p q) (Iminus (Iplus p q) p)))
   (Iminus q (Iminus (Iplus p q) p)) = (p + q - Iplus p q)%R :>R.
intros p q H' H'0; case (Rle_or_lt (Rabs q) (Rabs p)); intros Rle1.
apply MKnuth2; auto.
case (Rle_or_lt 0 q); intros Rle2.
apply MKnuth7; auto.
rewrite <- (Rabs_right q); auto with real; apply Rle_ge; apply Rlt_le; auto.
rewrite MKnuthOpp.
rewrite (Fopp_correct radix).
replace (p + q - Iplus p q)%R with
 (- (Fopp p + Fopp q - Iplus (Fopp p) (Fopp q)))%R.
cut (forall x y : R, x = y -> (- x)%R = (- y)%R);
 [ intros tmp; apply tmp; clear tmp | intros x y tmp; rewrite tmp; auto ].
apply MKnuth7; auto with zarith.
repeat rewrite (Fopp_correct radix); auto.
rewrite <- (Rabs_left1 (FtoR radix q)); try apply Rlt_le; auto.
rewrite Rabs_Ropp; auto.
apply oppBounded; auto.
apply oppBounded; auto.
rewrite <- IplusOp; repeat rewrite (Fopp_correct radix);
 unfold FtoRradix in |- *; ring.
Qed.
End EFast.


Section Round.
Variable b : Fbound.
Variable radix : Z.
Variable p : nat.

Hypothesis radixMoreThanOne : (1 < radix)%Z.
Hypothesis pGreaterThanOne : 1 < p.
Hypothesis pGivesBound : Zpos (vNum b) = Zpower_nat radix p.

Let FtoRradix := FtoR radix.
Local Coercion FtoRradix : float >-> R.

Various lemmas about exp, ln

Theorem exp_ln_powerRZ :
 forall u v : Z, (0 < u)%Z -> (exp (ln u * v) = powerRZ u v)%R.
intros u v H1.
cut (forall e f : nat, (0 < e)%Z -> exp (ln e * f) = powerRZ e f).
intros H2.
case (Zle_or_lt 0 v); intros H3.
replace u with (Z_of_nat (Z.abs_nat u));
 [ idtac | apply inj_abs; auto with zarith ].
replace v with (Z_of_nat (Z.abs_nat v)); [ idtac | apply inj_abs; auto ].
repeat rewrite <- INR_IZR_INZ; apply H2.
rewrite inj_abs; auto with zarith.
replace v with (- Z.abs_nat v)%Z.
rewrite <- Rinv_powerRZ; auto with zarith real.
replace u with (Z_of_nat (Z.abs_nat u));
 [ idtac | apply inj_abs; auto with zarith ].
rewrite Ropp_Ropp_IZR; rewrite Ropp_mult_distr_r_reverse; rewrite exp_Ropp;
 repeat rewrite <- INR_IZR_INZ.
rewrite H2; auto with real.
rewrite inj_abs; auto with zarith.
apply Rgt_not_eq, Rlt_IZR; auto with real.
rewrite <- Zabs_absolu; rewrite <- Zabs_Zopp.
rewrite Z.abs_eq; auto with zarith.
intros e f H2.
induction f as [| f Hrecf].
simpl in |- *; ring_simplify (ln e * 0)%R; apply exp_0.
replace (ln e * S f)%R with (ln e * f + ln e)%R.
rewrite exp_plus; rewrite Hrecf; rewrite exp_ln; auto with real zarith.
replace (Z_of_nat (S f)) with (f + 1)%Z.
rewrite powerRZ_add; try apply IZR_neq; simpl; auto with real zarith.
rewrite inj_S; auto with zarith.
replace (INR (S f)) with (f + 1)%R; [ ring | idtac ].
apply trans_eq with (IZR (f + 1)).
rewrite plus_IZR; simpl in |- *; rewrite <- INR_IZR_INZ; auto with real.
apply trans_eq with (IZR (Z.succ f)); auto with zarith real.
rewrite <- inj_S; rewrite <- INR_IZR_INZ; auto with zarith real.
Qed.

Theorem ln_radix_pos : (0 < ln radix)%R.
rewrite <- ln_1.
apply ln_increasing; try apply Rlt_IZR; auto with real zarith.
Qed.

Theorem exp_le_inv : forall x y : R, (exp x <= exp y)%R -> (x <= y)%R.
intros x y H; case H; intros H2.
left; apply exp_lt_inv; auto.
right; apply exp_inv; auto.
Qed.

Theorem exp_monotone : forall x y : R, (x <= y)%R -> (exp x <= exp y)%R.
intros x y H; case H; intros H2.
left; apply exp_increasing; auto.
right; auto with real.
Qed.

Theorem firstNormalPos_eq :
 FtoRradix (firstNormalPos radix b p) = powerRZ radix (Z.pred p + - dExp b).
Proof.
unfold firstNormalPos, nNormMin, FtoRradix, FtoR in |- *; simpl.
rewrite Zpower_nat_Z_powerRZ; rewrite powerRZ_add; try apply IZR_neq; auto with real zarith.
replace (Z_of_nat (pred p)) with (Z.pred p);
 [ ring | rewrite inj_pred; auto with zarith ].
Qed.

Results about the ineger rounding down

Definition IRNDD (r : R) := Z.pred (up r).

Theorem IRNDD_correct1 : forall r : R, (IRNDD r <= r)%R.
intros r; unfold IRNDD in |- *.
generalize (archimed r); intros T; elim T; intros H1 H2; clear T.
unfold Z.pred in |- *; apply Rplus_le_reg_l with (1 + - r)%R.
ring_simplify (1 + - r + r)%R.
apply Rle_trans with (2 := H2).
rewrite plus_IZR; right; simpl in |- *; ring.
Qed.

Theorem IRNDD_correct2 : forall r : R, (r < Z.succ (IRNDD r))%R.
intros r; unfold IRNDD in |- *.
generalize (archimed r); intros T; elim T; intros H1 H2; clear T.
rewrite <- Zsucc_pred; auto.
Qed.

Theorem IRNDD_correct3 : forall r : R, (r - 1 < IRNDD r)%R.
intros r; unfold IRNDD in |- *.
generalize (archimed r); intros T; elim T; intros H1 H2; clear T.
unfold Z.pred, Rminus in |- *; rewrite plus_IZR; simpl in |- *; auto with real.
Qed.

Theorem IRNDD_pos : forall r : R, (0 <= r)%R -> (0 <= IRNDD r)%R.
intros r H1; unfold IRNDD in |- *.
generalize (archimed r); intros T; elim T; intros H3 H2; clear T.
replace 0%R with (IZR 0); auto with real; apply IZR_le.
apply Zle_Zpred.
apply lt_IZR; apply Rle_lt_trans with r; auto with real zarith.
Qed.

Theorem IRNDD_eq :
 forall (r : R) (z : Z), (z <= r)%R -> (r < Z.succ z)%R -> IRNDD r = z.
intros r z H1 H2.
cut (IRNDD r - z < 1)%Z;
 [ intros H3 | apply lt_IZR; rewrite <- Z_R_minus; simpl in |- * ].
cut (z - IRNDD r < 1)%Z;
 [ intros H4; auto with zarith
 | apply lt_IZR; rewrite <- Z_R_minus; simpl in |- * ].
unfold Rminus in |- *; apply Rle_lt_trans with (r + - IRNDD r)%R;
 auto with real.
apply Rlt_le_trans with (r + - (r - 1))%R; auto with real.
2: right; ring.
apply Rplus_lt_compat_l, Ropp_lt_contravar.
apply IRNDD_correct3.
unfold Rminus in |- *; apply Rle_lt_trans with (r + - z)%R; auto with real.
apply Rplus_le_compat_r; apply IRNDD_correct1.
apply Rlt_le_trans with (Z.succ z + - z)%R; auto with real; right;
 unfold Z.succ in |- *; rewrite plus_IZR; simpl in |- *;
 ring.
Qed.

Theorem IRNDD_projector : forall z : Z, IRNDD z = z.
intros z.
apply IRNDD_eq; auto with zarith real.
apply Rlt_IZR; omega.
Qed.

Rounding down of a positive real

Definition RND_Min_Pos (r : R) :=
  match Rle_dec (firstNormalPos radix b p) r with
  | left _ =>
      let e := IRNDD (ln r / ln radix + (- Z.pred p)%Z) in
      Float (IRNDD (r * powerRZ radix (- e))) e
  | right _ => Float (IRNDD (r * powerRZ radix (dExp b))) (- dExp b)
  end.

Theorem RND_Min_Pos_bounded_aux :
 forall (r : R) (e : Z),
 (0 <= r)%R ->
 (- dExp b <= e)%Z ->
 (r < powerRZ radix (e + p))%R ->
 Fbounded b (Float (IRNDD (r * powerRZ radix (- e))) e).
intros r e H1 H2 H3.
split; auto.
simpl in |- *; rewrite pGivesBound; apply lt_IZR.
rewrite Zpower_nat_Z_powerRZ; rewrite <- Rabs_Zabs.
rewrite Rabs_right;
 [ idtac
 | apply Rle_ge; apply IRNDD_pos; apply Rmult_le_pos; auto with real zarith ].
apply Rle_lt_trans with (1 := IRNDD_correct1 (r * powerRZ radix (- e))).
apply Rmult_lt_reg_l with (powerRZ radix e); auto with zarith real.
apply powerRZ_lt, Rlt_IZR; omega.
rewrite Rmult_comm; rewrite Rmult_assoc.
rewrite <- powerRZ_add; try apply IZR_neq; auto with zarith real.
rewrite <- powerRZ_add; try apply IZR_neq; auto with zarith real.
apply Rle_lt_trans with (2 := H3); ring_simplify (- e + e)%Z; simpl in |- *; right;
 ring.
apply powerRZ_le, Rlt_IZR; omega.
Qed.

Theorem RND_Min_Pos_canonic :
 forall r : R, (0 <= r)%R -> Fcanonic radix b (RND_Min_Pos r).
intros r H1; unfold RND_Min_Pos in |- *.
generalize ln_radix_pos; intros W.
case (Rle_dec (firstNormalPos radix b p) r); intros H2.
cut (0 < r)%R; [ intros V | idtac ].
2: apply Rlt_le_trans with (2 := H2); rewrite firstNormalPos_eq;
    auto with real zarith.
left; split.
apply RND_Min_Pos_bounded_aux; auto.
apply Zgt_succ_le; apply Z.lt_gt.
apply lt_IZR.
apply
 Rle_lt_trans with (2 := IRNDD_correct2 (ln r / ln radix + (- Z.pred p)%Z)).
apply Rplus_le_reg_l with (Z.pred p).
apply Rmult_le_reg_l with (ln radix).
apply ln_radix_pos.
apply Rle_trans with (ln r).
apply exp_le_inv.
rewrite exp_ln; auto.
replace (Z.pred p + (- dExp b)%Z)%R with (IZR (Z.pred p + - dExp b)).
rewrite exp_ln_powerRZ; auto with zarith.
apply Rle_trans with (2 := H2).
rewrite firstNormalPos_eq; auto with real.
rewrite plus_IZR; rewrite Ropp_Ropp_IZR; ring.
rewrite Ropp_Ropp_IZR; right; field; auto with real.
rewrite <- exp_ln_powerRZ; auto with zarith.
pattern r at 1 in |- *; rewrite <- (exp_ln r); auto.
apply exp_increasing.
rewrite plus_IZR.
apply
 Rle_lt_trans with (ln radix * (ln r / ln radix + (- Z.pred p)%Z - 1 + p))%R.
rewrite Ropp_Ropp_IZR; unfold Z.pred in |- *; rewrite plus_IZR; simpl in |- *.
repeat rewrite <- INR_IZR_INZ; right; field; auto with real.
apply Rmult_lt_compat_l; auto with real.
repeat rewrite <- INR_IZR_INZ.
apply Rplus_lt_compat_r; auto with real.
apply IRNDD_correct3.
simpl in |- *; rewrite pGivesBound; apply le_IZR; simpl in |- *.
2: apply powerRZ_lt, Rlt_IZR; omega.
rewrite Zpower_nat_Z_powerRZ; rewrite Z.abs_eq.
2: apply le_IZR; rewrite mult_IZR; simpl in |- *.
2: apply Rmult_le_pos.
2: apply Rle_IZR; auto with real zarith.
2: apply IRNDD_pos; apply Rmult_le_pos; auto with real zarith.
2: apply powerRZ_le, Rlt_IZR; omega.
rewrite mult_IZR; pattern (Z_of_nat p) at 1 in |- *;
 replace (Z_of_nat p) with (1 + Z.pred p)%Z.
2: unfold Z.pred in |- *; ring.
rewrite powerRZ_add; try apply IZR_neq; auto with zarith real; simpl in |- *; ring_simplify (radix * 1)%R.
apply Rmult_le_compat_l; auto with zarith real.
apply Rle_IZR; omega.
rewrite <- inj_pred; auto with zarith.
rewrite <- Zpower_nat_Z_powerRZ; apply IZR_le.
apply Zgt_succ_le; apply Z.lt_gt; apply lt_IZR; rewrite Ropp_Ropp_IZR.
apply
 Rle_lt_trans
  with (r * powerRZ radix (- IRNDD (ln r / ln radix + - pred p)))%R.
2: repeat rewrite <- INR_IZR_INZ; apply IRNDD_correct2.
rewrite <- exp_ln_powerRZ; auto with zarith.
rewrite Zpower_nat_Z_powerRZ; rewrite Ropp_Ropp_IZR.
apply Rle_trans with (r * exp (ln radix * - (ln r / ln radix + - pred p)))%R.
pattern r at 1 in |- *; rewrite <- (exp_ln r); auto; rewrite <- exp_plus.
replace (ln r + ln radix * - (ln r / ln radix + - pred p))%R with
 (ln radix * pred p)%R.
2: field; auto with real.
rewrite INR_IZR_INZ; rewrite exp_ln_powerRZ; auto with real zarith.
apply Rmult_le_compat_l; auto with real.
apply exp_monotone; auto with real.
apply Rmult_le_compat_l; auto with real.
apply Ropp_le_contravar; apply IRNDD_correct1.
cut (r < powerRZ radix (Z.pred p + - dExp b))%R; [ intros H3 | idtac ].
2: rewrite <- firstNormalPos_eq; auto with real.
right; split.
pattern (dExp b) at 2 in |- *;
 replace (Z_of_N (dExp b)) with (- - dExp b)%Z; auto with zarith.
apply RND_Min_Pos_bounded_aux; auto with zarith.
apply Rlt_trans with (1 := H3); apply Rlt_powerRZ; try apply Rlt_IZR; auto with real zarith.
split; [ simpl in |- *; auto | idtac ].
simpl in |- *; rewrite pGivesBound; apply lt_IZR.
rewrite Zpower_nat_Z_powerRZ; rewrite <- Rabs_Zabs.
rewrite mult_IZR; rewrite Rabs_right;
 [ idtac
 | apply Rle_ge; apply Rmult_le_pos].
2: apply Rle_IZR; omega.
2: apply IRNDD_pos; apply Rmult_le_pos; auto with real zarith.
2: apply powerRZ_le, Rlt_IZR; omega.
apply Rle_lt_trans with (radix * (r * powerRZ radix (dExp b)))%R;
 auto with real zarith.
apply Rmult_le_compat_l; try apply Rle_IZR; auto with zarith.
apply IRNDD_correct1.
apply
 Rlt_le_trans
  with
    (radix * (powerRZ radix (Z.pred p + - dExp b) * powerRZ radix (dExp b)))%R;
 auto with real zarith.
apply Rmult_lt_compat_l.
apply Rlt_IZR; omega.
apply Rmult_lt_compat_r; auto with real.
apply powerRZ_lt, Rlt_IZR; omega.
rewrite <- powerRZ_add; try apply IZR_neq; auto with zarith real.
pattern (IZR radix) at 1 in |- *; replace (IZR radix) with (powerRZ radix 1);
 [ rewrite <- powerRZ_add; try apply IZR_neq | simpl in |- * ]; auto with zarith real;
 unfold Z.pred in |- *.
ring_simplify (1 + (p + -1 + - dExp b + dExp b))%Z; auto with real.
Qed.

Theorem RND_Min_Pos_Rle : forall r : R, (0 <= r)%R -> (RND_Min_Pos r <= r)%R.
intros r H.
unfold RND_Min_Pos in |- *; case (Rle_dec (firstNormalPos radix b p) r);
 intros H2.
unfold FtoRradix, FtoR in |- *; simpl in |- *.
apply
 Rle_trans
  with
    (r * powerRZ radix (- IRNDD (ln r / ln radix + (- Z.pred p)%Z)) *
     powerRZ radix (IRNDD (ln r / ln radix + (- Z.pred p)%Z)))%R;
 auto with real.
apply Rmult_le_compat_r.
apply powerRZ_le, Rlt_IZR; omega.
apply IRNDD_correct1.
rewrite Rmult_assoc; rewrite <- powerRZ_add; try apply IZR_neq; auto with real zarith.
ring_simplify
    (- IRNDD (ln r / ln radix + (- Z.pred p)%Z) +
     IRNDD (ln r / ln radix + (- Z.pred p)%Z))%Z; simpl in |- *;
 auto with real.
unfold FtoRradix, FtoR in |- *; simpl in |- *.
apply
 Rle_trans with (r * powerRZ radix (dExp b) * powerRZ radix (- dExp b))%R;
 auto with real.
apply Rmult_le_compat_r.
apply powerRZ_le, Rlt_IZR; omega.
apply IRNDD_correct1.
rewrite Rmult_assoc; rewrite <- powerRZ_add; try apply IZR_neq; auto with real zarith.
ring_simplify (dExp b + - dExp b)%Z; simpl in |- *; auto with real.
Qed.

Theorem RND_Min_Pos_monotone :
 forall r s : R,
 (0 <= r)%R -> (r <= s)%R -> (RND_Min_Pos r <= RND_Min_Pos s)%R.
intros r s V1 H.
cut (0 <= s)%R;
 [ intros V2 | apply Rle_trans with (1 := V1); auto with real ].
rewrite <-
 FPredSuc
          with
          (x := RND_Min_Pos s)
         (precision := p)
         (b := b)
         (radix := radix); auto with zarith.
2: apply RND_Min_Pos_canonic; auto.
unfold FtoRradix in |- *; apply FPredProp; auto with zarith;
 fold FtoRradix in |- *.
apply RND_Min_Pos_canonic; auto.
apply FSuccCanonic; auto with zarith; apply RND_Min_Pos_canonic; auto.
apply Rle_lt_trans with r; auto with real.
apply RND_Min_Pos_Rle; auto.
apply Rle_lt_trans with (1 := H).
replace (FtoRradix (FSucc b radix p (RND_Min_Pos s))) with
 (RND_Min_Pos s + powerRZ radix (Fexp (RND_Min_Pos s)))%R.
unfold RND_Min_Pos in |- *; case (Rle_dec (firstNormalPos radix b p) s);
 intros H1.
unfold FtoRradix, FtoR in |- *; simpl in |- *.
apply
 Rle_lt_trans
  with
    ((s * powerRZ radix (- IRNDD (ln s / ln radix + (- Z.pred p)%Z)) - 1) *
     powerRZ radix (IRNDD (ln s / ln radix + (- Z.pred p)%Z)) +
     powerRZ radix (IRNDD (ln s / ln radix + (- Z.pred p)%Z)))%R;
 auto with real.
right; ring_simplify.
rewrite Rmult_assoc; rewrite <- powerRZ_add; try apply IZR_neq; auto with zarith real.
ring_simplify
    (-IRNDD (ln s / ln radix + (- Z.pred p)%Z) +
      IRNDD (ln s / ln radix + (- Z.pred p)%Z))%Z;
   simpl; ring.
apply Rplus_lt_compat_r; apply Rmult_lt_compat_r.
apply powerRZ_lt, Rlt_IZR; omega.
apply IRNDD_correct3.
unfold FtoRradix, FtoR in |- *; simpl in |- *.
apply
 Rle_lt_trans
  with
    ((s * powerRZ radix (dExp b) - 1) * powerRZ radix (- dExp b) +
     powerRZ radix (- dExp b))%R; auto with real.
right; ring_simplify.
rewrite Rmult_assoc; rewrite <- powerRZ_add; try apply IZR_neq; auto with zarith real.
ring_simplify (dExp b + -dExp b)%Z; simpl in |- *; ring.
apply Rplus_lt_compat_r; apply Rmult_lt_compat_r.
apply powerRZ_lt, Rlt_IZR; omega.
apply IRNDD_correct3.
replace (powerRZ radix (Fexp (RND_Min_Pos s))) with
 (FtoR radix (Float 1%nat (Fexp (RND_Min_Pos s))));
 [ idtac | unfold FtoR in |- *; simpl in |- *; ring ].
rewrite <- FSuccDiff1 with b radix p (RND_Min_Pos s); auto with zarith.
rewrite Fminus_correct; auto with zarith; fold FtoRradix in |- *; ring.
cut (- nNormMin radix p < Fnum (RND_Min_Pos s))%Z; auto with zarith.
apply Z.lt_le_trans with 0%Z.
replace 0%Z with (- (0))%Z; unfold nNormMin in |- *; auto with arith zarith.
apply Zlt_Zopp, Zpower_nat_less; omega.
apply le_IZR; unfold RND_Min_Pos in |- *;
 case (Rle_dec (firstNormalPos radix b p) s); intros H1;
 simpl in |- *; apply IRNDD_pos; apply Rmult_le_pos;
 auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
apply powerRZ_le, Rlt_IZR; omega.
Qed.

Theorem RND_Min_Pos_projector :
 forall f : float,
 (0 <= f)%R -> Fcanonic radix b f -> FtoRradix (RND_Min_Pos f) = f.
intros f H1 H2.
unfold RND_Min_Pos in |- *; case (Rle_dec (firstNormalPos radix b p) f);
 intros H3.
replace (IRNDD (ln f / ln radix + (- Z.pred p)%Z)) with (Fexp f).
replace (f * powerRZ radix (- Fexp f))%R with (IZR (Fnum f)).
rewrite IRNDD_projector; unfold FtoRradix, FtoR in |- *; simpl in |- *; ring.
unfold FtoRradix, FtoR in |- *; simpl in |- *.
rewrite Rmult_assoc; rewrite <- powerRZ_add; try apply IZR_neq; auto with real zarith.
ring_simplify (Fexp f + - Fexp f)%Z; simpl in |- *; ring.
generalize ln_radix_pos; intros V1.
cut (0 < Fnum f)%R; [ intros V2 | idtac ].
apply sym_eq; apply IRNDD_eq.
unfold FtoRradix, FtoR in |- *; simpl in |- *.
rewrite ln_mult; auto with real zarith.
unfold Rdiv in |- *; rewrite Rmult_plus_distr_r.
apply Rle_trans with (Z.pred p + Fexp f + (- Z.pred p)%Z)%R;
 [ rewrite Ropp_Ropp_IZR; right; ring | idtac ].
apply Rplus_le_compat_r; apply Rplus_le_compat.
apply Rmult_le_reg_l with (ln radix); [ auto with real | idtac ].
apply Rle_trans with (ln (Fnum f)); [ idtac | right; field; auto with real ].
apply exp_le_inv.
rewrite exp_ln; auto.
rewrite exp_ln_powerRZ; auto with zarith.
case H2; intros T; elim T; intros C1 C2.
apply Rmult_le_reg_l with radix; auto with real zarith.
apply Rlt_IZR; omega.
apply Rle_trans with (IZR (Zpos (vNum b)));
 [ right; rewrite pGivesBound; rewrite Zpower_nat_Z_powerRZ | idtac ].
pattern (Z_of_nat p) at 2 in |- *; replace (Z_of_nat p) with (1 + Z.pred p)%Z;
 [ rewrite powerRZ_add; try apply IZR_neq; auto with real zarith; simpl in |- *; ring
 | unfold Z.pred in |- *; ring ].
rewrite <- (Rabs_right (radix * Fnum f)); auto with real zarith.
rewrite <- mult_IZR; rewrite Rabs_Zabs; auto with real zarith.
apply Rle_IZR; auto.
apply Rle_ge; apply Rmult_le_pos; auto with real.
apply Rle_IZR; auto with real zarith.
contradict H3; apply Rlt_not_le; unfold FtoRradix in |- *;
 apply FsubnormalLtFirstNormalPos; auto with zarith.
apply Rmult_le_reg_l with (ln radix); [ auto with real | idtac ].
apply Rle_trans with (ln (powerRZ radix (Fexp f)));
 [ idtac | right; field; auto with real ].
rewrite <- exp_ln_powerRZ; auto with zarith.
rewrite ln_exp; auto with real.
apply powerRZ_lt, Rlt_IZR; omega.
unfold FtoRradix, FtoR in |- *; simpl in |- *.
rewrite ln_mult; auto with real zarith.
rewrite <- exp_ln_powerRZ; auto with zarith.
rewrite ln_exp; auto with real.
unfold Rdiv in |- *; rewrite Rmult_plus_distr_r.
apply Rlt_le_trans with (p + Fexp f + (- Z.pred p)%Z)%R.
2: rewrite Ropp_Ropp_IZR; unfold Z.succ, Z.pred in |- *;
    repeat rewrite plus_IZR; repeat rewrite <- INR_IZR_INZ;
    simpl in |- *; right; ring.
replace (ln radix * Fexp f * / ln radix)%R with (IZR (Fexp f));
 [ idtac | field; auto with real ].
apply Rplus_lt_compat_r; apply Rplus_lt_compat_r.
apply Rmult_lt_reg_l with (ln radix); [ auto with real | idtac ].
apply Rle_lt_trans with (ln (Fnum f));
 [ right; field; auto with real | idtac ].
apply exp_lt_inv.
rewrite exp_ln; auto.
rewrite INR_IZR_INZ; rewrite exp_ln_powerRZ; auto with zarith.
apply Rlt_le_trans with (IZR (Zpos (vNum b))).
rewrite <- (Rabs_right (IZR (Fnum f))); auto with real.
rewrite Rabs_Zabs; apply Rlt_IZR; cut (Fbounded b f);
 auto with real zarith.
intros T; apply T.
apply FcanonicBound with radix; auto.
apply Rle_ge; auto with real.
right; rewrite pGivesBound; rewrite Zpower_nat_Z_powerRZ; auto with real zarith.
apply powerRZ_lt, Rlt_IZR; omega.
apply Rlt_IZR.
apply LtR0Fnum with radix; auto with zarith; fold FtoRradix in |- *.
apply Rlt_le_trans with (2 := H3); rewrite firstNormalPos_eq;
 auto with real zarith.
apply powerRZ_lt, Rlt_IZR; omega.
case H2; intros T; elim T; intros C1 C2.
absurd (firstNormalPos radix b p <= f)%R; auto with real.
unfold FtoRradix in |- *; apply FnormalLtFirstNormalPos; auto with zarith.
elim C2; intros C3 C4.
replace (f * powerRZ radix (dExp b))%R with (IZR (Fnum f)).
rewrite IRNDD_projector; rewrite <- C3; unfold FtoRradix, FtoR in |- *;
 simpl in |- *; ring.
unfold FtoRradix, FtoR in |- *; rewrite C3.
rewrite Rmult_assoc; rewrite <- powerRZ_add; try apply IZR_neq; auto with real zarith.
ring_simplify (- dExp b + dExp b)%Z; simpl in |- *; ring.
Qed.

Theorem RND_Min_Pos_correct :
 forall r : R, (0 <= r)%R -> isMin b radix r (RND_Min_Pos r).
intros r H1.
split.
apply FcanonicBound with radix; apply RND_Min_Pos_canonic; auto.
split.
apply RND_Min_Pos_Rle; auto.
fold FtoRradix in |- *; intros f H2 H3.
case (Rle_or_lt 0 f); intros H4.
unfold FtoRradix in |- *; rewrite <- FnormalizeCorrect with radix b p f; auto;
 fold FtoRradix in |- *.
rewrite <- RND_Min_Pos_projector.
apply RND_Min_Pos_monotone; unfold FtoRradix in |- *;
 rewrite FnormalizeCorrect; auto.
unfold FtoRradix in |- *; rewrite FnormalizeCorrect; auto.
apply FnormalizeCanonic; auto with zarith.
apply Rle_trans with 0%R; auto with real.
unfold RND_Min_Pos in |- *; case (Rle_dec (firstNormalPos radix b p) r);
 intros H5; unfold FtoRradix, FtoR in |- *; simpl in |- *;
 apply Rmult_le_pos; auto with real zarith; try apply powerRZ_le, Rlt_IZR; try omega.
apply IRNDD_pos; apply Rmult_le_pos; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
apply IRNDD_pos; apply Rmult_le_pos; auto with real zarith.
apply powerRZ_le, Rlt_IZR; omega.
Qed.

Easily deduced, the rounding up of a positive real

Definition RND_Max_Pos (r : R) :=
  match Req_EM_T r (RND_Min_Pos r) with
  | left _ => RND_Min_Pos r
  | right _ => FSucc b radix p (RND_Min_Pos r)
  end.

Theorem RND_Max_Pos_canonic :
 forall r : R, (0 <= r)%R -> Fcanonic radix b (RND_Max_Pos r).
intros r H; unfold RND_Max_Pos in |- *.
case (Req_EM_T r (RND_Min_Pos r)); intros H1.
apply RND_Min_Pos_canonic; auto.
apply FSuccCanonic; auto with zarith; apply RND_Min_Pos_canonic; auto.
Qed.

Theorem RND_Max_Pos_Rle : forall r : R, (0 <= r)%R -> (r <= RND_Max_Pos r)%R.
intros r H.
unfold RND_Max_Pos in |- *; case (Req_EM_T r (RND_Min_Pos r)); intros H1.
rewrite <- H1; auto with real.
case (Rle_or_lt r (FSucc b radix p (RND_Min_Pos r))); auto; intros H2.
generalize (RND_Min_Pos_correct r H); intros T; elim T; intros H3 T1; elim T1;
 intros H4 H5; clear T T1.
absurd (FSucc b radix p (RND_Min_Pos r) <= RND_Min_Pos r)%R;
 auto with zarith real.
apply Rlt_not_le; auto with zarith.
unfold FtoRradix in |- *; apply FSuccLt; auto with zarith.
apply H5; auto with real.
apply FBoundedSuc; auto with zarith.
Qed.

Theorem RND_Max_Pos_correct :
 forall r : R, (0 <= r)%R -> isMax b radix r (RND_Max_Pos r).
intros r H.
split.
apply FcanonicBound with radix; apply RND_Max_Pos_canonic; auto.
split.
apply RND_Max_Pos_Rle; auto.
unfold RND_Max_Pos in |- *; case (Req_EM_T r (RND_Min_Pos r)); intros H1.
fold FtoRradix in |- *; intros f H2 H3; rewrite <- H1; auto with real.
fold FtoRradix in |- *; intros f H2 H3.
case H3; intros V.
case (Rle_or_lt (FSucc b radix p (RND_Min_Pos r)) f); auto; intros H4.
absurd (f < f)%R; auto with real.
apply Rle_lt_trans with (RND_Min_Pos r).
rewrite <- FPredSuc with b radix p (RND_Min_Pos r); auto with zarith.
2: apply RND_Min_Pos_canonic; auto.
unfold FtoRradix in |- *; rewrite <- FnormalizeCorrect with radix b p f; auto.
apply FPredProp; auto with arith zarith.
apply FnormalizeCanonic; auto with zarith.
apply FSuccCanonic; auto with zarith; apply RND_Min_Pos_canonic; auto.
rewrite FnormalizeCorrect; auto with real.
apply Rle_lt_trans with (2 := V).
apply RND_Min_Pos_Rle; auto.
contradict H1.
rewrite V; unfold FtoRradix in |- *;
 rewrite <- FnormalizeCorrect with radix b p f; auto.
fold FtoRradix in |- *; apply sym_eq; apply RND_Min_Pos_projector;
 auto with zarith.
unfold FtoRradix in |- *; rewrite FnormalizeCorrect; fold FtoRradix in |- *;
 auto with real.
apply Rle_trans with r; auto with real.
apply FnormalizeCanonic; auto with zarith.
Qed.

Roundings up and down of any real

Definition RND_Min (r : R) :=
  match Rle_dec 0 r with
  | left _ => RND_Min_Pos r
  | right _ => Fopp (RND_Max_Pos (- r))
  end.

Theorem RND_Min_canonic : forall r : R, Fcanonic radix b (RND_Min r).
intros r.
unfold RND_Min in |- *; case (Rle_dec 0 r); intros H.
apply RND_Min_Pos_canonic; auto.
apply FcanonicFopp; apply RND_Max_Pos_canonic; auto with real.
rewrite <- Ropp_0; apply Ropp_le_contravar; auto with real.
Qed.

Theorem RND_Min_correct : forall r : R, isMin b radix r (RND_Min r).
intros r.
unfold RND_Min in |- *; case (Rle_dec 0 r); intros H.
apply RND_Min_Pos_correct; auto.
pattern r at 1 in |- *; rewrite <- (Ropp_involutive r).
apply MaxOppMin; apply RND_Max_Pos_correct; auto with real.
rewrite <- Ropp_0; apply Ropp_le_contravar; auto with real.
Qed.

Definition RND_Max (r : R) :=
  match Rle_dec 0 r with
  | left _ => RND_Max_Pos r
  | right _ => Fopp (RND_Min_Pos (- r))
  end.

Theorem RND_Max_canonic : forall r : R, Fcanonic radix b (RND_Max r).
intros r.
unfold RND_Max in |- *; case (Rle_dec 0 r); intros H.
apply RND_Max_Pos_canonic; auto.
apply FcanonicFopp; apply RND_Min_Pos_canonic; auto with real.
rewrite <- Ropp_0; apply Ropp_le_contravar; auto with real.
Qed.

Theorem RND_Max_correct : forall r : R, isMax b radix r (RND_Max r).
intros r.
unfold RND_Max in |- *; case (Rle_dec 0 r); intros H.
apply RND_Max_Pos_correct; auto.
pattern r at 1 in |- *; rewrite <- (Ropp_involutive r).
apply MinOppMax; apply RND_Min_Pos_correct; auto with real.
rewrite <- Ropp_0; apply Ropp_le_contravar; auto with real.
Qed.

Rounding to the nearest of any real First, ClosestUp (when equality, the biggest is returned) Then, EvenClosest (when equality, the even is returned)

Definition RND_EvenClosest (r : R) :=
  match Rle_dec (Rabs (RND_Max r - r)) (Rabs (RND_Min r - r)) with
  | left H =>
      match
        Rle_lt_or_eq_dec (Rabs (RND_Max r - r)) (Rabs (RND_Min r - r)) H
      with
      | left _ => RND_Max r
      | right _ =>
          match OddEvenDec (Fnum (RND_Min r)) with
          | left _ => RND_Max r
          | right _ => RND_Min r
          end
      end
  | right _ => RND_Min r
  end.

Theorem RND_EvenClosest_canonic :
 forall r : R, Fcanonic radix b (RND_EvenClosest r).
intros r; unfold RND_EvenClosest in |- *.
case (Rle_dec (Rabs (RND_Max r - r)) (Rabs (RND_Min r - r))); intros H1.
case (Rle_lt_or_eq_dec (Rabs (RND_Max r - r)) (Rabs (RND_Min r - r)) H1);
 intros H2.
apply RND_Max_canonic.
case (OddEvenDec (Fnum (RND_Min r))); intros H3.
apply RND_Max_canonic.
apply RND_Min_canonic.
apply RND_Min_canonic.
Qed.

Theorem RND_EvenClosest_correct :
 forall r : R, EvenClosest b radix p r (RND_EvenClosest r).
intros r.
cut (RND_Min r <= r)%R; [ intros V1 | idtac ].
2: generalize (RND_Min_correct r); intros T; elim T; intros T1 T2; elim T2;
    intros T3 T4; auto with real.
cut (r <= RND_Max r)%R; [ intros V2 | idtac ].
2: generalize (RND_Max_correct r); intros T; elim T; intros T1 T2; elim T2;
    intros T3 T4; auto with real.
cut (forall v w : R, (v <= w)%R -> (0 <= w - v)%R); [ intros V3 | idtac ].
2: intros v w H; apply Rplus_le_reg_l with v; ring_simplify; auto with real.
cut (forall v w : R, (v <= w)%R -> (v - w <= 0)%R); [ intros V4 | idtac ].
2: intros v w H; apply Rplus_le_reg_l with w; ring_simplify; auto with real.
unfold RND_EvenClosest in |- *;
 case (Rle_dec (Rabs (RND_Max r - r)) (Rabs (RND_Min r - r)));
 intros H1.
case (Rle_lt_or_eq_dec (Rabs (RND_Max r - r)) (Rabs (RND_Min r - r)) H1);
 intros H2.
split.
split.
apply FcanonicBound with radix; apply RND_Max_canonic.
intros f H3; fold FtoRradix in |- *.
rewrite Rabs_right in H1; [ idtac | apply Rle_ge; apply V3; auto with real ].
rewrite Rabs_left1 in H1; [ idtac | apply V4; auto with real ].
rewrite Rabs_right; [ idtac | apply Rle_ge; apply V3; auto with real ].
case (Rle_or_lt f r); intros H4.
rewrite Rabs_left1; [ idtac | apply V4; auto with real ].
apply Rle_trans with (1 := H1); apply Ropp_le_contravar;
 unfold Rminus in |- *; apply Rplus_le_compat_r.
generalize (RND_Min_correct r); intros T; elim T; intros T1 T2; elim T2;
 intros T3 T4; auto with real.
rewrite Rabs_right; [ idtac | apply Rle_ge; apply V3; auto with real ].
unfold Rminus in |- *; apply Rplus_le_compat_r.
generalize (RND_Max_correct r); intros T; elim T; intros T1 T2; elim T2;
 intros T3 T4; auto with real.
right; intros q H3.
generalize (ClosestMinOrMax b radix); unfold MinOrMaxP in |- *; intros T.
case (T r q); auto; intros H4; clear T.
contradict H2; apply Rle_not_lt.
replace (FtoRradix (RND_Min r)) with (FtoRradix q).
elim H3; intros T1 T2; unfold FtoRradix in |- *; apply T2.
apply FcanonicBound with radix; apply RND_Max_canonic.
generalize (MinUniqueP b radix); unfold UniqueP in |- *; intros T;
 apply T with r; auto.
apply RND_Min_correct.
generalize (MaxUniqueP b radix); unfold UniqueP in |- *; intros T;
 apply T with r; auto.
apply RND_Max_correct.
case (OddEvenDec (Fnum (RND_Min r))); intros H3.
split.
split.
apply FcanonicBound with radix; apply RND_Max_canonic.
intros f H4; fold FtoRradix in |- *.
rewrite Rabs_right in H1; [ idtac | apply Rle_ge; apply V3; auto with real ].
rewrite Rabs_left1 in H1; [ idtac | apply V4; auto with real ].
rewrite Rabs_right; [ idtac | apply Rle_ge; apply V3; auto with real ].
case (Rle_or_lt f r); intros H5.
rewrite Rabs_left1; [ idtac | apply V4; auto with real ].
apply Rle_trans with (1 := H1); apply Ropp_le_contravar;
 unfold Rminus in |- *; apply Rplus_le_compat_r.
generalize (RND_Min_correct r); intros T; elim T; intros T1 T2; elim T2;
 intros T3 T4; auto with real.
rewrite Rabs_right; [ idtac | apply Rle_ge; apply V3; auto with real ].
unfold Rminus in |- *; apply Rplus_le_compat_r.
generalize (RND_Max_correct r); intros T; elim T; intros T1 T2; elim T2;
 intros T3 T4; auto with real.
case (Req_EM_T (RND_Max r) (RND_Min r)); intros W.
right; intros q H4.
generalize (ClosestMinOrMax b radix); unfold MinOrMaxP in |- *; intros T.
case (T r q); auto; intros H5; clear T.
fold FtoRradix in |- *; rewrite W; generalize (MinUniqueP b radix);
 unfold UniqueP in |- *; intros T; apply T with r;
 auto.
apply RND_Min_correct.
generalize (MaxUniqueP b radix); unfold UniqueP in |- *; intros T;
 apply T with r; auto.
apply RND_Max_correct.
left; unfold FNeven in |- *.
rewrite FcanonicFnormalizeEq; auto with zarith;
 [ idtac | apply RND_Max_canonic ].
replace (RND_Max r) with (FSucc b radix p (RND_Min r)).
apply FoddSuc; auto.
unfold RND_Max, RND_Min in |- *; case (Rle_dec 0 r); intros W1.
unfold RND_Max_Pos in |- *.
case (Req_EM_T r (RND_Min_Pos r)); intros W2; auto.
contradict W.
pattern r at 1 in |- *; rewrite W2.
apply sym_eq; unfold FtoRradix in |- *;
 apply RoundedModeProjectorIdemEq with b p (isMax b radix);
 auto.
apply MaxRoundedModeP with p; auto.
apply FcanonicBound with radix; apply RND_Min_canonic.
replace (FtoR radix (RND_Min r)) with (FtoR radix (RND_Min_Pos r));
 [ fold FtoRradix in |- *; rewrite <- W2; apply RND_Max_correct | idtac ].
fold FtoRradix in |- *; unfold RND_Min in |- *; auto with real.
case (Rle_dec 0 r); auto with real; intros W3; contradict W1; auto with real.
unfold RND_Max_Pos in |- *.
case (Req_EM_T (- r) (RND_Min_Pos (- r))); intros W2; auto.
contradict W.
cut (r = FtoRradix (Fopp (RND_Min_Pos (- r)))); [ intros W3 | idtac ].
pattern r at 1 in |- *; rewrite W3.
apply sym_eq; unfold FtoRradix in |- *;
 apply RoundedModeProjectorIdemEq with b p (isMax b radix);
 auto.
apply MaxRoundedModeP with p; auto.
apply FcanonicBound with radix; apply RND_Min_canonic.
replace (FtoR radix (RND_Min r)) with (- FtoR radix (RND_Min_Pos (- r)))%R;
 [ fold FtoRradix in |- *; rewrite <- W3 | idtac ].
rewrite <- W2; rewrite Ropp_involutive; apply RND_Max_correct.
fold FtoRradix in |- *; unfold RND_Min in |- *; auto with real.
case (Rle_dec 0 r).
intros W4; contradict W1; auto with real.
intros W4; unfold RND_Max_Pos in |- *;
 case (Req_EM_T (- r) (RND_Min_Pos (- r))); intros W5.
unfold FtoRradix in |- *; rewrite Fopp_correct; ring.
contradict W2; auto with real.
unfold FtoRradix in |- *; rewrite Fopp_correct; fold FtoRradix in |- *;
 rewrite <- W2; ring.
pattern (RND_Min_Pos (- r)) at 1 in |- *;
 rewrite <- (Fopp_Fopp (RND_Min_Pos (- r))).
rewrite <- FPredFopFSucc; auto with zarith.
apply FSucPred; auto with zarith.
apply FcanonicFopp; apply RND_Min_Pos_canonic; auto with real.
rewrite <- Ropp_0; apply Ropp_le_contravar; auto with real.
split.
split.
apply FcanonicBound with radix; apply RND_Min_canonic.
intros f H4; fold FtoRradix in |- *.
rewrite Rabs_right in H1; [ idtac | apply Rle_ge; apply V3; auto with real ].
rewrite Rabs_left1 in H1; [ idtac | apply V4; auto with real ].
case (Rle_or_lt f r); intros H5.
rewrite Rabs_left1; [ idtac | apply V4; auto with real ].
rewrite Rabs_left1; [ idtac | apply V4; auto with real ].
apply Ropp_le_contravar; unfold Rminus in |- *; apply Rplus_le_compat_r.
generalize (RND_Min_correct r); intros T; elim T; intros T1 T2; elim T2;
 intros T3 T4; auto with real.
rewrite <- H2.
rewrite Rabs_right; [ idtac | apply Rle_ge; apply V3; auto with real ].
rewrite Rabs_right; [ idtac | apply Rle_ge; apply V3; auto with real ].
unfold Rminus in |- *; apply Rplus_le_compat_r.
generalize (RND_Max_correct r); intros T; elim T; intros T1 T2; elim T2;
 intros T3 T4; auto with real.
left; unfold FNeven in |- *.
rewrite FcanonicFnormalizeEq; auto with zarith; apply RND_Min_canonic.
cut (Rabs (RND_Min r - r) < Rabs (RND_Max r - r))%R; auto with real;
 intros H'.
cut (Rabs (RND_Min r - r) < Rabs (RND_Max r - r))%R; auto with real;
 intros H''.
rewrite Rabs_left1 in H'; [ idtac | apply V4; auto with real ].
rewrite Rabs_right in H'; [ idtac | apply Rle_ge; apply V3; auto with real ].
split.
split.
apply FcanonicBound with radix; apply RND_Min_canonic.
intros f W1.
rewrite Rabs_left1; [ idtac | apply V4; auto with real ].
case (Rle_or_lt f r); intros H2.
rewrite Rabs_left1; [ idtac | apply V4; auto with real ].
apply Ropp_le_contravar; unfold Rminus in |- *; apply Rplus_le_compat_r.
generalize (RND_Min_correct r); intros T; elim T; intros T1 T2; elim T2;
 intros T3 T4; auto with real.
rewrite Rabs_right; [ idtac | apply Rle_ge; apply V3; auto with real ].
apply Rle_trans with (RND_Max r - r)%R; auto with real; unfold Rminus in |- *;
 apply Rplus_le_compat_r.
generalize (RND_Max_correct r); intros T; elim T; intros T1 T2; elim T2;
 intros T3 T4; auto with real.
right; intros q H3.
generalize (ClosestMinOrMax b radix); unfold MinOrMaxP in |- *; intros T.
case (T r q); auto; intros H4; clear T.
generalize (MinUniqueP b radix); unfold UniqueP in |- *; intros T;
 apply T with r; auto.
apply RND_Min_correct.
contradict H''; apply Rle_not_lt.
replace (FtoRradix (RND_Max r)) with (FtoRradix q).
elim H3; intros T1 T2; unfold FtoRradix in |- *; apply T2.
apply FcanonicBound with radix; apply RND_Min_canonic.
generalize (MaxUniqueP b radix); unfold UniqueP in |- *; intros T;
 apply T with r; auto.
apply RND_Max_correct.
Qed.

End Round.