Built with Alectryon, running Coq+SerAPI v8.10.0+0.7.0. Coq sources are in this panel; goals and messages will appear in the other. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑ Ctrl+↓ to navigate, Ctrl+🖱️ to focus.
This file is part of the Flocq formalization of floating-point arithmetic in Coq: http://flocq.gforge.inria.fr/
Copyright (C) 2009-2018 Sylvie Boldo
Copyright (C) 2009-2018 Guillaume Melquiond
This library is free software; you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation; either version 3 of the License, or (at your option) any later version.
This library is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the COPYING file for more details.

Unit in the Last Place: our definition using fexp and its properties, successor and predecessor

Require Import Reals Psatz.
Require Import Raux Defs Round_pred Generic_fmt Float_prop.

Section Fcore_ulp.

Variable beta : radix.

Notation bpow e := (bpow beta e).

Variable fexp : Z -> Z.
Definition and basic properties about the minimal exponent, when it exists
beta:radix
fexp:Z -> Z

forall x y : Z, (x <= y)%Z \/ ~ (x <= y)%Z
beta:radix
fexp:Z -> Z

forall x y : Z, (x <= y)%Z \/ ~ (x <= y)%Z
beta:radix
fexp:Z -> Z
x, y:Z

(x <= y)%Z \/ ~ (x <= y)%Z
beta:radix
fexp:Z -> Z
x, y:Z
l:(x <= y)%Z

(x <= y)%Z \/ ~ (x <= y)%Z
beta:radix
fexp:Z -> Z
x, y:Z
n:~ (x <= y)%Z
(x <= y)%Z \/ ~ (x <= y)%Z
beta:radix
fexp:Z -> Z
x, y:Z
n:~ (x <= y)%Z

(x <= y)%Z \/ ~ (x <= y)%Z
now right. Qed.
negligible_exp is either none (as in FLX) or Some n such that n <= fexp n.
Definition negligible_exp: option Z :=
  match (LPO_Z _ (fun z => Z_le_dec_aux z (fexp z))) with
   | inleft N => Some (proj1_sig N)
   | inright _ => None
 end.


Inductive negligible_exp_prop: option Z -> Prop :=
  | negligible_None: (forall n, (fexp n < n)%Z) -> negligible_exp_prop None
  | negligible_Some: forall n, (n <= fexp n)%Z -> negligible_exp_prop (Some n).


beta:radix
fexp:Z -> Z

negligible_exp_prop negligible_exp
beta:radix
fexp:Z -> Z

negligible_exp_prop negligible_exp
beta:radix
fexp:Z -> Z
n:Z
Hn:(n <= fexp n)%Z

negligible_exp_prop (Some (proj1_sig (exist (fun n0 : Z => (n0 <= fexp n0)%Z) n Hn)))
beta:radix
fexp:Z -> Z
Hn:forall n : Z, ~ (n <= fexp n)%Z
negligible_exp_prop None
beta:radix
fexp:Z -> Z
Hn:forall n : Z, ~ (n <= fexp n)%Z

negligible_exp_prop None
beta:radix
fexp:Z -> Z
Hn:forall n : Z, ~ (n <= fexp n)%Z

forall n : Z, (fexp n < n)%Z
intros n; specialize (Hn n); omega. Qed.
beta:radix
fexp:Z -> Z

negligible_exp = None /\ (forall n : Z, (fexp n < n)%Z) \/ (exists n : Z, negligible_exp = Some n /\ (n <= fexp n)%Z)
beta:radix
fexp:Z -> Z

negligible_exp = None /\ (forall n : Z, (fexp n < n)%Z) \/ (exists n : Z, negligible_exp = Some n /\ (n <= fexp n)%Z)
beta:radix
fexp:Z -> Z
n:Z
Hn:(n <= fexp n)%Z

Some (proj1_sig (exist (fun n0 : Z => (n0 <= fexp n0)%Z) n Hn)) = None /\ (forall n0 : Z, (fexp n0 < n0)%Z) \/ (exists n0 : Z, Some (proj1_sig (exist (fun n1 : Z => (n1 <= fexp n1)%Z) n Hn)) = Some n0 /\ (n0 <= fexp n0)%Z)
beta:radix
fexp:Z -> Z
Hn:forall n : Z, ~ (n <= fexp n)%Z
None = None /\ (forall n : Z, (fexp n < n)%Z) \/ (exists n : Z, None = Some n /\ (n <= fexp n)%Z)
beta:radix
fexp:Z -> Z
Hn:forall n : Z, ~ (n <= fexp n)%Z

None = None /\ (forall n : Z, (fexp n < n)%Z) \/ (exists n : Z, None = Some n /\ (n <= fexp n)%Z)
beta:radix
fexp:Z -> Z
Hn:forall n : Z, ~ (n <= fexp n)%Z

forall n : Z, (fexp n < n)%Z
intros n; specialize (Hn n); omega. Qed. Context { valid_exp : Valid_exp fexp }.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall n m : Z, (n <= fexp n)%Z -> (m <= fexp m)%Z -> fexp n = fexp m
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall n m : Z, (n <= fexp n)%Z -> (m <= fexp m)%Z -> fexp n = fexp m
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
n, m:Z
Hn:(n <= fexp n)%Z
Hm:(m <= fexp m)%Z

fexp n = fexp m
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
n, m:Z
Hn:(n <= fexp n)%Z
Hm:(m <= fexp m)%Z
H:(n <= m)%Z

fexp n = fexp m
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
n, m:Z
Hn:(n <= fexp n)%Z
Hm:(m <= fexp m)%Z
H:(m < n)%Z
fexp n = fexp m
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
n, m:Z
Hn:(n <= fexp n)%Z
Hm:(m <= fexp m)%Z
H:(m < n)%Z

fexp n = fexp m
apply sym_eq, valid_exp; omega. Qed.
Definition and basic properties about the ulp
Now includes a nice ulp(0): ulp(0) is now 0 when there is no minimal exponent, such as in FLX, and beta^(fexp n) when there is a n such that n <= fexp n. For instance, the value of ulp(O) is then beta^emin in FIX and FLT. The main lemma to use is ulp_neq_0 that is equivalent to the previous "unfold ulp" provided the value is not zero.
Definition ulp x := match Req_bool x 0 with
  | true   => match negligible_exp with
            | Some n => bpow (fexp n)
            | None => 0%R
            end
  | false  => bpow (cexp beta fexp x)
 end.

beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, x <> 0%R -> ulp x = bpow (cexp beta fexp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, x <> 0%R -> ulp x = bpow (cexp beta fexp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:x <> 0%R

ulp x = bpow (cexp beta fexp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:x <> 0%R

x = 0%R -> match negligible_exp with | Some n => bpow (fexp n) | None => 0%R end = bpow (cexp beta fexp x)
intros H; now contradict H. Qed. Notation F := (generic_format beta fexp).
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, ulp (- x) = ulp x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, ulp (- x) = ulp x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R

ulp (- x) = ulp x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R

(if Req_bool (- x) 0 then match negligible_exp with | Some n => bpow (fexp n) | None => 0%R end else bpow (cexp beta fexp (- x))) = (if Req_bool x 0 then match negligible_exp with | Some n => bpow (fexp n) | None => 0%R end else bpow (cexp beta fexp x))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
H1:(- x)%R = 0%R

match negligible_exp with | Some n => bpow (fexp n) | None => 0%R end = (if Req_bool x 0 then match negligible_exp with | Some n => bpow (fexp n) | None => 0%R end else bpow (cexp beta fexp x))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
H1:(- x)%R <> 0%R
bpow (cexp beta fexp (- x)) = (if Req_bool x 0 then match negligible_exp with | Some n => bpow (fexp n) | None => 0%R end else bpow (cexp beta fexp x))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
H1:(- x)%R = 0%R

x = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
H1:(- x)%R <> 0%R
bpow (cexp beta fexp (- x)) = (if Req_bool x 0 then match negligible_exp with | Some n => bpow (fexp n) | None => 0%R end else bpow (cexp beta fexp x))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
H1:(- x)%R <> 0%R

bpow (cexp beta fexp (- x)) = (if Req_bool x 0 then match negligible_exp with | Some n => bpow (fexp n) | None => 0%R end else bpow (cexp beta fexp x))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
H1:(- x)%R <> 0%R

bpow (cexp beta fexp (- x)) = bpow (cexp beta fexp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
H1:(- x)%R <> 0%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
H1:(- x)%R <> 0%R

x <> 0%R
intros H2; apply H1; rewrite H2; ring. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, ulp (Rabs x) = ulp x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, ulp (Rabs x) = ulp x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R

ulp (Rabs x) = ulp x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
H1:x = 0%R

(if Req_bool (Rabs x) 0 then match negligible_exp with | Some n => bpow (fexp n) | None => 0%R end else bpow (cexp beta fexp (Rabs x))) = match negligible_exp with | Some n => bpow (fexp n) | None => 0%R end
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
H1:x <> 0%R
(if Req_bool (Rabs x) 0 then match negligible_exp with | Some n => bpow (fexp n) | None => 0%R end else bpow (cexp beta fexp (Rabs x))) = bpow (cexp beta fexp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
H1:x = 0%R

Rabs x = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
H1:x <> 0%R
(if Req_bool (Rabs x) 0 then match negligible_exp with | Some n => bpow (fexp n) | None => 0%R end else bpow (cexp beta fexp (Rabs x))) = bpow (cexp beta fexp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
H1:x <> 0%R

(if Req_bool (Rabs x) 0 then match negligible_exp with | Some n => bpow (fexp n) | None => 0%R end else bpow (cexp beta fexp (Rabs x))) = bpow (cexp beta fexp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
H1:x <> 0%R

bpow (cexp beta fexp (Rabs x)) = bpow (cexp beta fexp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
H1:x <> 0%R
Rabs x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
H1:x <> 0%R

Rabs x <> 0%R
now apply Rabs_no_R0. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, (0 <= ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, (0 <= ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
H:x = 0%R

(0 <= match negligible_exp with | Some n => bpow (fexp n) | None => 0 end)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
H:x <> 0%R
(0 <= bpow (cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
H:x = 0%R
z:Z

(0 <= bpow (fexp z))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
H:x = 0%R
(0 <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
H:x <> 0%R
(0 <= bpow (cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
H:x = 0%R

(0 <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
H:x <> 0%R
(0 <= bpow (cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
H:x <> 0%R

(0 <= bpow (cexp beta fexp x))%R
apply bpow_ge_0. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, (0 < x)%R -> F x -> (ulp x <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, (0 < x)%R -> F x -> (ulp x <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x

(ulp x <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x

(1 * ulp x <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x

(1 * ulp x <= F2R {| Fnum := Ztrunc (scaled_mantissa beta fexp x); Fexp := cexp beta fexp x |})%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x

(1 * bpow (cexp beta fexp x) <= F2R {| Fnum := Ztrunc (scaled_mantissa beta fexp x); Fexp := cexp beta fexp x |})%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x

(1 * bpow (cexp beta fexp x) <= F2R {| Fnum := Ztrunc (scaled_mantissa beta fexp x); Fexp := cexp beta fexp x |})%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x

(1 * bpow (cexp beta fexp x) <= IZR (Ztrunc (scaled_mantissa beta fexp x)) * bpow (cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x

(0 <= bpow (cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
(1 <= IZR (Ztrunc (scaled_mantissa beta fexp x)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x

(1 <= IZR (Ztrunc (scaled_mantissa beta fexp x)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x

(0 < Ztrunc (scaled_mantissa beta fexp x))%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x

(0 < F2R {| Fnum := Ztrunc (scaled_mantissa beta fexp x); Fexp := cexp beta fexp x |})%R
now rewrite <- Fx. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, x <> 0%R -> F x -> (ulp x <= Rabs x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, x <> 0%R -> F x -> (ulp x <= Rabs x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:x <> 0%R
Fx:F x

(ulp x <= Rabs x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:x <> 0%R
Fx:F x

(ulp (Rabs x) <= Rabs x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:x <> 0%R
Fx:F x

(0 < Rabs x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:x <> 0%R
Fx:F x
F (Rabs x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:x <> 0%R
Fx:F x

F (Rabs x)
now apply generic_format_abs. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, ~ F x -> round beta fexp Zceil x = (round beta fexp Zfloor x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, ~ F x -> round beta fexp Zceil x = (round beta fexp Zfloor x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x

round beta fexp Zceil x = (round beta fexp Zfloor x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x

round beta fexp Zceil x = (round beta fexp Zfloor x + bpow (cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x

F2R {| Fnum := Zceil (scaled_mantissa beta fexp x); Fexp := cexp beta fexp x |} = (F2R {| Fnum := Zfloor (scaled_mantissa beta fexp x); Fexp := cexp beta fexp x |} + bpow (cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x

F2R {| Fnum := Zceil (scaled_mantissa beta fexp x); Fexp := cexp beta fexp x |} = (F2R {| Fnum := Zfloor (scaled_mantissa beta fexp x); Fexp := cexp beta fexp x |} + bpow (cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x

(IZR (Fnum {| Fnum := Zceil (scaled_mantissa beta fexp x); Fexp := cexp beta fexp x |}) * bpow (Fexp {| Fnum := Zceil (scaled_mantissa beta fexp x); Fexp := cexp beta fexp x |}))%R = (IZR (Fnum {| Fnum := Zfloor (scaled_mantissa beta fexp x); Fexp := cexp beta fexp x |}) * bpow (Fexp {| Fnum := Zfloor (scaled_mantissa beta fexp x); Fexp := cexp beta fexp x |}) + bpow (cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x

(IZR (Zceil (scaled_mantissa beta fexp x)) * bpow (cexp beta fexp x))%R = (IZR (Zfloor (scaled_mantissa beta fexp x)) * bpow (cexp beta fexp x) + bpow (cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x

(IZR (Zfloor (scaled_mantissa beta fexp x) + 1) * bpow (cexp beta fexp x))%R = (IZR (Zfloor (scaled_mantissa beta fexp x)) * bpow (cexp beta fexp x) + bpow (cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
IZR (Zfloor (scaled_mantissa beta fexp x)) <> scaled_mantissa beta fexp x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x

((IZR (Zfloor (scaled_mantissa beta fexp x)) + 1) * bpow (cexp beta fexp x))%R = (IZR (Zfloor (scaled_mantissa beta fexp x)) * bpow (cexp beta fexp x) + bpow (cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
IZR (Zfloor (scaled_mantissa beta fexp x)) <> scaled_mantissa beta fexp x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x

((IZR (Zfloor (scaled_mantissa beta fexp x)) + 1) * bpow (cexp beta fexp x))%R = (IZR (Zfloor (scaled_mantissa beta fexp x)) * bpow (cexp beta fexp x) + bpow (cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
IZR (Zfloor (scaled_mantissa beta fexp x)) <> scaled_mantissa beta fexp x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x

IZR (Zfloor (scaled_mantissa beta fexp x)) <> scaled_mantissa beta fexp x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
H:IZR (Zfloor (scaled_mantissa beta fexp x)) = scaled_mantissa beta fexp x

False
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
H:IZR (Zfloor (scaled_mantissa beta fexp x)) = scaled_mantissa beta fexp x

F x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
H:IZR (Zfloor (scaled_mantissa beta fexp x)) = scaled_mantissa beta fexp x

x = (IZR (Fnum {| Fnum := Ztrunc (scaled_mantissa beta fexp x); Fexp := cexp beta fexp x |}) * bpow (Fexp {| Fnum := Ztrunc (scaled_mantissa beta fexp x); Fexp := cexp beta fexp x |}))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
H:IZR (Zfloor (scaled_mantissa beta fexp x)) = scaled_mantissa beta fexp x

x = (IZR (Ztrunc (scaled_mantissa beta fexp x)) * bpow (cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
H:IZR (Zfloor (scaled_mantissa beta fexp x)) = scaled_mantissa beta fexp x

x = (IZR (Ztrunc (IZR (Zfloor (scaled_mantissa beta fexp x)))) * bpow (cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
H:IZR (Zfloor (scaled_mantissa beta fexp x)) = scaled_mantissa beta fexp x

x = (IZR (Zfloor (scaled_mantissa beta fexp x)) * bpow (cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
H:IZR (Zfloor (scaled_mantissa beta fexp x)) = scaled_mantissa beta fexp x

x = (scaled_mantissa beta fexp x * bpow (cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x

x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
V:x = 0%R

F x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
V:x = 0%R

F 0
apply generic_format_0. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall m e : Z, m <> 0%Z -> canonical beta fexp {| Fnum := m; Fexp := e |} -> ulp (F2R {| Fnum := m; Fexp := e |}) = bpow e
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall m e : Z, m <> 0%Z -> canonical beta fexp {| Fnum := m; Fexp := e |} -> ulp (F2R {| Fnum := m; Fexp := e |}) = bpow e
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
m, e:Z
Hm:m <> 0%Z
Hc:canonical beta fexp {| Fnum := m; Fexp := e |}

ulp (F2R {| Fnum := m; Fexp := e |}) = bpow e
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
m, e:Z
Hm:m <> 0%Z
Hc:canonical beta fexp {| Fnum := m; Fexp := e |}

bpow (cexp beta fexp (F2R {| Fnum := m; Fexp := e |})) = bpow e
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
m, e:Z
Hm:m <> 0%Z
Hc:canonical beta fexp {| Fnum := m; Fexp := e |}

cexp beta fexp (F2R {| Fnum := m; Fexp := e |}) = e
now apply sym_eq. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall e : Z, ulp (bpow e) = bpow (fexp (e + 1))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall e : Z, ulp (bpow e) = bpow (fexp (e + 1))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e:Z

ulp (bpow e) = bpow (fexp (e + 1))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e:Z

bpow (cexp beta fexp (bpow e)) = bpow (fexp (e + 1))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e:Z
bpow e <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e:Z

cexp beta fexp (bpow e) = fexp (e + 1)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e:Z
bpow e <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e:Z

(bpow (e + 1 - 1) <= Rabs (bpow e) < bpow (e + 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e:Z
bpow e <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e:Z

(bpow (e + 1 - 1) <= bpow e < bpow (e + 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e:Z
(0 <= bpow e)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e:Z
bpow e <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e:Z

(bpow (e + 1 - 1) <= bpow e)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e:Z
(bpow e < bpow (e + 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e:Z
(0 <= bpow e)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e:Z
bpow e <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e:Z

(bpow e <= bpow e)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e:Z
(bpow e < bpow (e + 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e:Z
(0 <= bpow e)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e:Z
bpow e <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e:Z

(bpow e < bpow (e + 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e:Z
(0 <= bpow e)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e:Z
bpow e <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e:Z

(e < e + 1)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e:Z
(0 <= bpow e)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e:Z
bpow e <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e:Z

(0 <= bpow e)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e:Z
bpow e <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e:Z

bpow e <> 0%R
apply Rgt_not_eq, Rlt_gt, bpow_gt_0. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

F (ulp 0)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

F (ulp 0)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

F (if Req_bool 0 0 then match negligible_exp with | Some n => bpow (fexp n) | None => 0%R end else bpow (cexp beta fexp 0))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

F match negligible_exp with | Some n => bpow (fexp n) | None => 0 end
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

(forall n : Z, (fexp n < n)%Z) -> F 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
forall n : Z, (n <= fexp n)%Z -> F (bpow (fexp n))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall n : Z, (n <= fexp n)%Z -> F (bpow (fexp n))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
n:Z
H1:(n <= fexp n)%Z

F (bpow (fexp n))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
n:Z
H1:(n <= fexp n)%Z

(fexp (fexp n + 1) <= fexp n)%Z
now apply valid_exp. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall e : Z, (ulp 0 <= bpow e)%R -> F (bpow e)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall e : Z, (ulp 0 <= bpow e)%R -> F (bpow e)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e:Z

((if Req_bool 0 0 then match negligible_exp with | Some n => bpow (fexp n) | None => 0 end else bpow (cexp beta fexp 0)) <= bpow e)%R -> F (bpow e)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e:Z

(match negligible_exp with | Some n => bpow (fexp n) | None => 0 end <= bpow e)%R -> F (bpow e)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e:Z

(forall n : Z, (fexp n < n)%Z) -> (0 <= bpow e)%R -> F (bpow e)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e:Z
forall n : Z, (n <= fexp n)%Z -> (bpow (fexp n) <= bpow e)%R -> F (bpow e)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e:Z
H1:forall n : Z, (fexp n < n)%Z

F (bpow e)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e:Z
forall n : Z, (n <= fexp n)%Z -> (bpow (fexp n) <= bpow e)%R -> F (bpow e)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e:Z
H1:forall n : Z, (fexp n < n)%Z

(fexp (e + 1) <= e)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e:Z
forall n : Z, (n <= fexp n)%Z -> (bpow (fexp n) <= bpow e)%R -> F (bpow e)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e:Z

forall n : Z, (n <= fexp n)%Z -> (bpow (fexp n) <= bpow e)%R -> F (bpow e)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e, n:Z
H1:(n <= fexp n)%Z
H2:(bpow (fexp n) <= bpow e)%R

F (bpow e)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e, n:Z
H1:(n <= fexp n)%Z
H2:(bpow (fexp n) <= bpow e)%R

(fexp (e + 1) <= e)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e, n:Z
H1:(n <= fexp n)%Z
H2:(bpow (fexp n) <= bpow e)%R
H4:(e + 1 <= fexp (e + 1))%Z

(fexp (e + 1) <= e)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e, n:Z
H1:(n <= fexp n)%Z
H2:(bpow (fexp n) <= bpow e)%R
H4:(fexp (e + 1) < e + 1)%Z
(fexp (e + 1) <= e)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e, n:Z
H1:(n <= fexp n)%Z
H2:(bpow (fexp n) <= bpow e)%R
H4:(e + 1 <= fexp (e + 1))%Z

~ (e + 1 <= e)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e, n:Z
H1:(n <= fexp n)%Z
H2:(bpow (fexp n) <= bpow e)%R
H4:(e + 1 <= fexp (e + 1))%Z
(e + 1 <= e)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e, n:Z
H1:(n <= fexp n)%Z
H2:(bpow (fexp n) <= bpow e)%R
H4:(fexp (e + 1) < e + 1)%Z
(fexp (e + 1) <= e)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e, n:Z
H1:(n <= fexp n)%Z
H2:(bpow (fexp n) <= bpow e)%R
H4:(e + 1 <= fexp (e + 1))%Z

(e + 1 <= e)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e, n:Z
H1:(n <= fexp n)%Z
H2:(bpow (fexp n) <= bpow e)%R
H4:(fexp (e + 1) < e + 1)%Z
(fexp (e + 1) <= e)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e, n:Z
H1:(n <= fexp n)%Z
H2:(bpow (fexp n) <= bpow e)%R
H4:(e + 1 <= fexp (e + 1))%Z

(fexp (e + 1) <= e)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e, n:Z
H1:(n <= fexp n)%Z
H2:(bpow (fexp n) <= bpow e)%R
H4:(fexp (e + 1) < e + 1)%Z
(fexp (e + 1) <= e)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e, n:Z
H1:(n <= fexp n)%Z
H2:(bpow (fexp n) <= bpow e)%R
H4:(e + 1 <= fexp (e + 1))%Z

(fexp n <= e)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e, n:Z
H1:(n <= fexp n)%Z
H2:(bpow (fexp n) <= bpow e)%R
H4:(e + 1 <= fexp (e + 1))%Z
fexp n = fexp (e + 1)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e, n:Z
H1:(n <= fexp n)%Z
H2:(bpow (fexp n) <= bpow e)%R
H4:(fexp (e + 1) < e + 1)%Z
(fexp (e + 1) <= e)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e, n:Z
H1:(n <= fexp n)%Z
H2:(bpow (fexp n) <= bpow e)%R
H4:(e + 1 <= fexp (e + 1))%Z

fexp n = fexp (e + 1)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e, n:Z
H1:(n <= fexp n)%Z
H2:(bpow (fexp n) <= bpow e)%R
H4:(fexp (e + 1) < e + 1)%Z
(fexp (e + 1) <= e)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e, n:Z
H1:(n <= fexp n)%Z
H2:(bpow (fexp n) <= bpow e)%R
H4:(fexp (e + 1) < e + 1)%Z

(fexp (e + 1) <= e)%Z
omega. Qed.
The three following properties are equivalent: Exp_not_FTZ ; forall x, F (ulp x) ; forall x, ulp 0 <= ulp x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

Exp_not_FTZ fexp -> forall x : R, F (ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

Exp_not_FTZ fexp -> forall x : R, F (ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:forall e : Z, (fexp (fexp e + 1) <= fexp e)%Z
x:R

F (ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:forall e : Z, (fexp (fexp e + 1) <= fexp e)%Z
x:R
Hx:x = 0%R

F (ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:forall e : Z, (fexp (fexp e + 1) <= fexp e)%Z
x:R
Hx:x <> 0%R
F (ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:forall e : Z, (fexp (fexp e + 1) <= fexp e)%Z
x:R
Hx:x <> 0%R

F (ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:forall e : Z, (fexp (fexp e + 1) <= fexp e)%Z
x:R
Hx:x <> 0%R

F (bpow (cexp beta fexp x))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:forall e : Z, (fexp (fexp e + 1) <= fexp e)%Z
x:R
Hx:x <> 0%R

(fexp (cexp beta fexp x + 1) <= cexp beta fexp x)%Z
apply H. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

(forall x : R, F (ulp x)) -> Exp_not_FTZ fexp
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

(forall x : R, F (ulp x)) -> Exp_not_FTZ fexp
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:forall x : R, F (ulp x)
e:Z

(fexp (fexp e + 1) <= fexp e)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e:Z
H:F (ulp (bpow (e - 1)))

(fexp (fexp e + 1) <= fexp e)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e:Z
H:F (bpow (cexp beta fexp (bpow (e - 1))))

(fexp (fexp e + 1) <= fexp e)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e:Z
H:F (ulp (bpow (e - 1)))
bpow (e - 1) <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e:Z
H:F (bpow (cexp beta fexp (bpow (e - 1))))

(fexp (fexp e + 1) <= fexp e)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e:Z
H:F (bpow (fexp (mag beta (bpow (e - 1)))))

(fexp (fexp e + 1) <= fexp e)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e:Z
H:F (bpow (fexp (e - 1 + 1)))

(fexp (fexp e + 1) <= fexp e)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e:Z
H:(fexp (fexp (e - 1 + 1) + 1) <= fexp (e - 1 + 1))%Z

(fexp (fexp e + 1) <= fexp e)%Z
now replace (e-1+1)%Z with e in H by ring. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

Exp_not_FTZ fexp -> forall x : R, (ulp 0 <= ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

Exp_not_FTZ fexp -> forall x : R, (ulp 0 <= ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:forall e : Z, (fexp (fexp e + 1) <= fexp e)%Z
x:R

(ulp 0 <= ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:forall e : Z, (fexp (fexp e + 1) <= fexp e)%Z
x:R
Hx:x = 0%R

(ulp 0 <= ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:forall e : Z, (fexp (fexp e + 1) <= fexp e)%Z
x:R
Hx:x <> 0%R
(ulp 0 <= ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:forall e : Z, (fexp (fexp e + 1) <= fexp e)%Z
x:R
Hx:x <> 0%R

(ulp 0 <= ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:forall e : Z, (fexp (fexp e + 1) <= fexp e)%Z
x:R
Hx:x <> 0%R

((if Req_bool 0 0 then match negligible_exp with | Some n => bpow (fexp n) | None => 0 end else bpow (cexp beta fexp 0)) <= ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:forall e : Z, (fexp (fexp e + 1) <= fexp e)%Z
x:R
Hx:x <> 0%R

(match negligible_exp with | Some n => bpow (fexp n) | None => 0 end <= ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:forall e : Z, (fexp (fexp e + 1) <= fexp e)%Z
x:R
Hx:x <> 0%R

negligible_exp = None /\ (forall n : Z, (fexp n < n)%Z) -> (match negligible_exp with | Some n => bpow (fexp n) | None => 0 end <= ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:forall e : Z, (fexp (fexp e + 1) <= fexp e)%Z
x:R
Hx:x <> 0%R
(exists n : Z, negligible_exp = Some n /\ (n <= fexp n)%Z) -> (match negligible_exp with | Some n => bpow (fexp n) | None => 0 end <= ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:forall e : Z, (fexp (fexp e + 1) <= fexp e)%Z
x:R
Hx:x <> 0%R

(exists n : Z, negligible_exp = Some n /\ (n <= fexp n)%Z) -> (match negligible_exp with | Some n => bpow (fexp n) | None => 0 end <= ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:forall e : Z, (fexp (fexp e + 1) <= fexp e)%Z
x:R
Hx:x <> 0%R
n:Z
H1:negligible_exp = Some n
H2:(n <= fexp n)%Z

(bpow (fexp n) <= ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:forall e : Z, (fexp (fexp e + 1) <= fexp e)%Z
x:R
Hx:x <> 0%R
n:Z
H1:negligible_exp = Some n
H2:(n <= fexp n)%Z

(bpow (fexp n) <= bpow (cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:forall e : Z, (fexp (fexp e + 1) <= fexp e)%Z
x:R
Hx:x <> 0%R
n:Z
H1:negligible_exp = Some n
H2:(n <= fexp n)%Z

(fexp n <= fexp (mag beta x))%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:forall e : Z, (fexp (fexp e + 1) <= fexp e)%Z
x:R
Hx:x <> 0%R
n:Z
H1:negligible_exp = Some n
H2:(n <= fexp n)%Z
l:mag_prop beta x

(fexp n <= fexp l)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:forall e : Z, (fexp (fexp e + 1) <= fexp e)%Z
x:R
Hx:x <> 0%R
n:Z
H1:negligible_exp = Some n
H2:(n <= fexp n)%Z
l:mag_prop beta x
Hl:(l <= fexp l)%Z

(fexp n <= fexp l)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:forall e : Z, (fexp (fexp e + 1) <= fexp e)%Z
x:R
Hx:x <> 0%R
n:Z
H1:negligible_exp = Some n
H2:(n <= fexp n)%Z
l:mag_prop beta x
Hl:(fexp l < l)%Z
(fexp n <= fexp l)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:forall e : Z, (fexp (fexp e + 1) <= fexp e)%Z
x:R
Hx:x <> 0%R
n:Z
H1:negligible_exp = Some n
H2:(n <= fexp n)%Z
l:mag_prop beta x
Hl:(fexp l < l)%Z

(fexp n <= fexp l)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:forall e : Z, (fexp (fexp e + 1) <= fexp e)%Z
x:R
Hx:x <> 0%R
n:Z
H1:negligible_exp = Some n
H2:(n <= fexp n)%Z
l:mag_prop beta x
Hl:(fexp l < l)%Z
K:(fexp l < fexp n)%Z

(fexp n <= fexp l)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:forall e : Z, (fexp (fexp e + 1) <= fexp e)%Z
x:R
Hx:x <> 0%R
n:Z
H1:negligible_exp = Some n
H2:(n <= fexp n)%Z
l:mag_prop beta x
Hl:(fexp l < l)%Z
K:(fexp l < fexp n)%Z

~ (fexp n <= fexp l)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:forall e : Z, (fexp (fexp e + 1) <= fexp e)%Z
x:R
Hx:x <> 0%R
n:Z
H1:negligible_exp = Some n
H2:(n <= fexp n)%Z
l:mag_prop beta x
Hl:(fexp l < l)%Z
K:(fexp l < fexp n)%Z
(fexp n <= fexp l)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:forall e : Z, (fexp (fexp e + 1) <= fexp e)%Z
x:R
Hx:x <> 0%R
n:Z
H1:negligible_exp = Some n
H2:(n <= fexp n)%Z
l:mag_prop beta x
Hl:(fexp l < l)%Z
K:(fexp l < fexp n)%Z

(fexp n <= fexp l)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:forall e : Z, (fexp (fexp e + 1) <= fexp e)%Z
x:R
Hx:x <> 0%R
n:Z
H1:negligible_exp = Some n
H2:(n <= fexp n)%Z
l:mag_prop beta x
Hl:(fexp l < l)%Z
K:(fexp l < fexp n)%Z

(fexp n <= fexp (fexp l + 1))%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:forall e : Z, (fexp (fexp e + 1) <= fexp e)%Z
x:R
Hx:x <> 0%R
n:Z
H1:negligible_exp = Some n
H2:(n <= fexp n)%Z
l:mag_prop beta x
Hl:(fexp l < l)%Z
K:(fexp l < fexp n)%Z

(fexp l + 1 <= fexp n)%Z
omega. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

(forall x : R, (ulp 0 <= ulp x)%R) -> Exp_not_FTZ fexp
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

(forall x : R, (ulp 0 <= ulp x)%R) -> Exp_not_FTZ fexp
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:forall x : R, (ulp 0 <= ulp x)%R
e:Z

(fexp (fexp e + 1) <= fexp e)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:forall x : R, (ulp 0 <= ulp x)%R
e:Z

F (bpow (fexp e))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:forall x : R, (ulp 0 <= ulp x)%R
e:Z

(ulp 0 <= bpow (fexp e))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:forall x : R, (ulp 0 <= ulp x)%R
e:Z

(ulp 0 <= bpow (fexp (e - 1 + 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:forall x : R, (ulp 0 <= ulp x)%R
e:Z

(ulp 0 <= ulp (bpow (e - 1)))%R
apply H. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

Monotone_exp fexp -> forall x y : R, (0 <= x)%R -> (x <= y)%R -> (ulp x <= ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

Monotone_exp fexp -> forall x y : R, (0 <= x)%R -> (x <= y)%R -> (ulp x <= ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
x, y:R
Hx:(0 <= x)%R
Hxy:(x <= y)%R

(ulp x <= ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R

(ulp x <= ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
x, y:R
Hx:0%R = x
Hxy:(x <= y)%R
(ulp x <= ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R

(bpow (cexp beta fexp x) <= ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
x, y:R
Hx:0%R = x
Hxy:(x <= y)%R
(ulp x <= ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R

(bpow (cexp beta fexp x) <= bpow (cexp beta fexp y))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
y <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
x, y:R
Hx:0%R = x
Hxy:(x <= y)%R
(ulp x <= ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R

(cexp beta fexp x <= cexp beta fexp y)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
y <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
x, y:R
Hx:0%R = x
Hxy:(x <= y)%R
(ulp x <= ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R

(mag beta x <= mag beta y)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
y <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
x, y:R
Hx:0%R = x
Hxy:(x <= y)%R
(ulp x <= ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R

y <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
x, y:R
Hx:0%R = x
Hxy:(x <= y)%R
(ulp x <= ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R

(0 < y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
x, y:R
Hx:0%R = x
Hxy:(x <= y)%R
(ulp x <= ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R

x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
x, y:R
Hx:0%R = x
Hxy:(x <= y)%R
(ulp x <= ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
x, y:R
Hx:0%R = x
Hxy:(x <= y)%R

(ulp x <= ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
x, y:R
Hx:0%R = x
Hxy:(x <= y)%R

(ulp 0 <= ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
x, y:R
Hx:0%R = x
Hxy:(x <= y)%R

Exp_not_FTZ fexp
apply monotone_exp_not_FTZ... Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

Monotone_exp fexp -> forall x y : R, (Rabs x <= Rabs y)%R -> (ulp x <= ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

Monotone_exp fexp -> forall x y : R, (Rabs x <= Rabs y)%R -> (ulp x <= ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
x, y:R
Hxy:(Rabs x <= Rabs y)%R

(ulp x <= ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
x, y:R
Hxy:(Rabs x <= Rabs y)%R

(ulp (Rabs x) <= ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
x, y:R
Hxy:(Rabs x <= Rabs y)%R

(ulp (Rabs x) <= ulp (Rabs y))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
x, y:R
Hxy:(Rabs x <= Rabs y)%R

(0 <= Rabs x)%R
apply Rabs_pos. Qed.
Properties when there is no minimal exponent
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

negligible_exp = None -> forall rnd : R -> Z, Valid_rnd rnd -> forall x : R, round beta fexp rnd x = 0%R -> x = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

negligible_exp = None -> forall rnd : R -> Z, Valid_rnd rnd -> forall x : R, round beta fexp rnd x = 0%R -> x = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:negligible_exp = None
rnd:R -> Z
Vr:Valid_rnd rnd
x:R
Hx:round beta fexp rnd x = 0%R

x = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:negligible_exp = None
rnd:R -> Z
Vr:Valid_rnd rnd
x:R
Hx:round beta fexp rnd x = 0%R
Hx2:x <> 0%R

x = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:negligible_exp = None
rnd:R -> Z
Vr:Valid_rnd rnd
x:R
Hx:round beta fexp rnd x = 0%R
Hx2:x <> 0%R

Rabs (round beta fexp rnd x) <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:negligible_exp = None
rnd:R -> Z
Vr:Valid_rnd rnd
x:R
Hx:round beta fexp rnd x = 0%R
Hx2:x <> 0%R
Rabs (round beta fexp rnd x) = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:negligible_exp = None
rnd:R -> Z
Vr:Valid_rnd rnd
x:R
Hx:round beta fexp rnd x = 0%R
Hx2:x <> 0%R

Rabs (round beta fexp rnd x) <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:negligible_exp = None
rnd:R -> Z
Vr:Valid_rnd rnd
x:R
Hx:round beta fexp rnd x = 0%R
Hx2:x <> 0%R

(Rabs (round beta fexp rnd x) > 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:negligible_exp = None
rnd:R -> Z
Vr:Valid_rnd rnd
x:R
Hx:round beta fexp rnd x = 0%R
Hx2:x <> 0%R

(0 < bpow (mag beta x - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:negligible_exp = None
rnd:R -> Z
Vr:Valid_rnd rnd
x:R
Hx:round beta fexp rnd x = 0%R
Hx2:x <> 0%R
(bpow (mag beta x - 1) <= Rabs (round beta fexp rnd x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:negligible_exp = None
rnd:R -> Z
Vr:Valid_rnd rnd
x:R
Hx:round beta fexp rnd x = 0%R
Hx2:x <> 0%R

(bpow (mag beta x - 1) <= Rabs (round beta fexp rnd x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:negligible_exp = None
rnd:R -> Z
Vr:Valid_rnd rnd
x:R
Hx:round beta fexp rnd x = 0%R
Hx2:x <> 0%R

F (bpow (mag beta x - 1))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:negligible_exp = None
rnd:R -> Z
Vr:Valid_rnd rnd
x:R
Hx:round beta fexp rnd x = 0%R
Hx2:x <> 0%R
(bpow (mag beta x - 1) <= Rabs x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:negligible_exp = None
rnd:R -> Z
Vr:Valid_rnd rnd
x:R
Hx:round beta fexp rnd x = 0%R
Hx2:x <> 0%R

(fexp (mag beta x - 1 + 1) <= mag beta x - 1)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:negligible_exp = None
rnd:R -> Z
Vr:Valid_rnd rnd
x:R
Hx:round beta fexp rnd x = 0%R
Hx2:x <> 0%R
(bpow (mag beta x - 1) <= Rabs x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:negligible_exp = None
rnd:R -> Z
Vr:Valid_rnd rnd
x:R
Hx:round beta fexp rnd x = 0%R
Hx2:x <> 0%R
K1:negligible_exp = None
K2:forall n : Z, (fexp n < n)%Z

(fexp (mag beta x - 1 + 1) <= mag beta x - 1)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:negligible_exp = None
rnd:R -> Z
Vr:Valid_rnd rnd
x:R
Hx:round beta fexp rnd x = 0%R
Hx2:x <> 0%R
(exists n : Z, negligible_exp = Some n /\ (n <= fexp n)%Z) -> (fexp (mag beta x - 1 + 1) <= mag beta x - 1)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:negligible_exp = None
rnd:R -> Z
Vr:Valid_rnd rnd
x:R
Hx:round beta fexp rnd x = 0%R
Hx2:x <> 0%R
(bpow (mag beta x - 1) <= Rabs x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:negligible_exp = None
rnd:R -> Z
Vr:Valid_rnd rnd
x:R
Hx:round beta fexp rnd x = 0%R
Hx2:x <> 0%R
K1:negligible_exp = None
K2:forall n : Z, (fexp n < n)%Z

(fexp (mag beta x) <= mag beta x - 1)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:negligible_exp = None
rnd:R -> Z
Vr:Valid_rnd rnd
x:R
Hx:round beta fexp rnd x = 0%R
Hx2:x <> 0%R
(exists n : Z, negligible_exp = Some n /\ (n <= fexp n)%Z) -> (fexp (mag beta x - 1 + 1) <= mag beta x - 1)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:negligible_exp = None
rnd:R -> Z
Vr:Valid_rnd rnd
x:R
Hx:round beta fexp rnd x = 0%R
Hx2:x <> 0%R
(bpow (mag beta x - 1) <= Rabs x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:negligible_exp = None
rnd:R -> Z
Vr:Valid_rnd rnd
x:R
Hx:round beta fexp rnd x = 0%R
Hx2:x <> 0%R

(exists n : Z, negligible_exp = Some n /\ (n <= fexp n)%Z) -> (fexp (mag beta x - 1 + 1) <= mag beta x - 1)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:negligible_exp = None
rnd:R -> Z
Vr:Valid_rnd rnd
x:R
Hx:round beta fexp rnd x = 0%R
Hx2:x <> 0%R
(bpow (mag beta x - 1) <= Rabs x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:negligible_exp = None
rnd:R -> Z
Vr:Valid_rnd rnd
x:R
Hx:round beta fexp rnd x = 0%R
Hx2:x <> 0%R
n:Z
Hn1:negligible_exp = Some n
Hn2:(n <= fexp n)%Z

(fexp (mag beta x - 1 + 1) <= mag beta x - 1)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:negligible_exp = None
rnd:R -> Z
Vr:Valid_rnd rnd
x:R
Hx:round beta fexp rnd x = 0%R
Hx2:x <> 0%R
(bpow (mag beta x - 1) <= Rabs x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:negligible_exp = None
rnd:R -> Z
Vr:Valid_rnd rnd
x:R
Hx:round beta fexp rnd x = 0%R
Hx2:x <> 0%R

(bpow (mag beta x - 1) <= Rabs x)%R
now apply bpow_mag_le. Qed.
Definition and properties of pred and succ
Definition pred_pos x :=
  if Req_bool x (bpow (mag beta x - 1)) then
    (x - bpow (fexp (mag beta x - 1)))%R
  else
    (x - ulp x)%R.

Definition succ x :=
  if (Rle_bool 0 x) then
    (x+ulp x)%R
  else
    (- pred_pos (-x))%R.

Definition pred x := (- succ (-x))%R.

beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, (0 <= x)%R -> pred x = pred_pos x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, (0 <= x)%R -> pred x = pred_pos x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 <= x)%R

(- (if Rle_bool 0 (- x) then - x + ulp (- x) else - pred_pos (- - x)))%R = pred_pos x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 <= x)%R
Hx':(0 <= - x)%R

(- (- x + ulp (- x)))%R = pred_pos x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 <= x)%R
Hx':(- x < 0)%R
(- - pred_pos (- - x))%R = pred_pos x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 <= x)%R
Hx':(0 <= - x)%R

x = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 <= x)%R
Hx':(0 <= - x)%R
K:x = 0%R
(- (- x + ulp (- x)))%R = pred_pos x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 <= x)%R
Hx':(- x < 0)%R
(- - pred_pos (- - x))%R = pred_pos x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 <= x)%R
Hx':(0 <= - x)%R

(x <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 <= x)%R
Hx':(0 <= - x)%R
K:x = 0%R
(- (- x + ulp (- x)))%R = pred_pos x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 <= x)%R
Hx':(- x < 0)%R
(- - pred_pos (- - x))%R = pred_pos x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 <= x)%R
Hx':(0 <= - x)%R

(- 0 <= - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 <= x)%R
Hx':(0 <= - x)%R
K:x = 0%R
(- (- x + ulp (- x)))%R = pred_pos x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 <= x)%R
Hx':(- x < 0)%R
(- - pred_pos (- - x))%R = pred_pos x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 <= x)%R
Hx':(0 <= - x)%R
K:x = 0%R

(- (- x + ulp (- x)))%R = pred_pos x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 <= x)%R
Hx':(- x < 0)%R
(- - pred_pos (- - x))%R = pred_pos x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 <= x)%R
Hx':(0 <= - x)%R
K:x = 0%R

(- (- 0 + ulp (- 0)))%R = (if Req_bool 0 (bpow (mag beta 0 - 1)) then (0 - bpow (fexp (mag beta 0 - 1)))%R else (0 - ulp 0)%R)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 <= x)%R
Hx':(- x < 0)%R
(- - pred_pos (- - x))%R = pred_pos x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 <= x)%R
Hx':(0 <= - x)%R
K:x = 0%R

(- (- 0 + ulp (- 0)))%R = (0 - ulp 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 <= x)%R
Hx':(0 <= - x)%R
K:x = 0%R
0%R <> bpow (mag beta 0 - 1)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 <= x)%R
Hx':(- x < 0)%R
(- - pred_pos (- - x))%R = pred_pos x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 <= x)%R
Hx':(0 <= - x)%R
K:x = 0%R

(- (- 0 + ulp (- 0)))%R = (0 - ulp 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 <= x)%R
Hx':(- x < 0)%R
(- - pred_pos (- - x))%R = pred_pos x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 <= x)%R
Hx':(- x < 0)%R

(- - pred_pos (- - x))%R = pred_pos x
now rewrite 2!Ropp_involutive. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, (0 <= x)%R -> succ x = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, (0 <= x)%R -> succ x = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 <= x)%R

(if Rle_bool 0 x then (x + ulp x)%R else (- pred_pos (- x))%R) = (x + ulp x)%R
now rewrite Rle_bool_true. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, succ (- x) = (- pred x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, succ (- x) = (- pred x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R

succ (- x) = (- pred x)%R
now apply sym_eq, Ropp_involutive. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, pred (- x) = (- succ x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, pred (- x) = (- succ x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R

pred (- x) = (- succ x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R

(- succ (- - x))%R = (- succ x)%R
now rewrite Ropp_involutive. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall e : Z, pred (bpow e) = (bpow e - bpow (fexp e))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall e : Z, pred (bpow e) = (bpow e - bpow (fexp e))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e:Z

pred (bpow e) = (bpow e - bpow (fexp e))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e:Z

pred_pos (bpow e) = (bpow e - bpow (fexp e))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e:Z

(if Req_bool (bpow e) (bpow (mag beta (bpow e) - 1)) then (bpow e - bpow (fexp (mag beta (bpow e) - 1)))%R else (bpow e - ulp (bpow e))%R) = (bpow e - bpow (fexp e))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e:Z

(if Req_bool (bpow e) (bpow (e + 1 - 1)) then (bpow e - bpow (fexp (e + 1 - 1)))%R else (bpow e - ulp (bpow e))%R) = (bpow e - bpow (fexp e))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
e:Z

(if Req_bool (bpow e) (bpow e) then (bpow e - bpow (fexp e))%R else (bpow e - ulp (bpow e))%R) = (bpow e - bpow (fexp e))%R
now rewrite Req_bool_true. Qed.
pred and succ are in the format
(* cannont be x <> ulp 0, due to the counter-example 1-bit FP format fexp: e -> e-1 *)
(* was pred_ge_bpow *)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall (x : R) (e : Z), F x -> x <> ulp x -> (bpow e < x)%R -> (bpow e <= x - ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall (x : R) (e : Z), F x -> x <> ulp x -> (bpow e < x)%R -> (bpow e <= x - ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx':x <> ulp x
Hx:(bpow e < x)%R

(bpow e <= x - ulp x)%R
(* *)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx':x <> ulp x
Hx:(bpow e < x)%R

(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx':x <> ulp x
Hx:(bpow e < x)%R
H:(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
(bpow e <= x - ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx':x <> ulp x
Hx:(bpow e < x)%R

(0 < Ztrunc (scaled_mantissa beta fexp x))%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx':x <> ulp x
Hx:(bpow e < x)%R
H:(0 < Ztrunc (scaled_mantissa beta fexp x))%Z
(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx':x <> ulp x
Hx:(bpow e < x)%R
H:(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
(bpow e <= x - ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx':x <> ulp x
Hx:(bpow e < x)%R

(0 < F2R {| Fnum := Ztrunc (scaled_mantissa beta fexp x); Fexp := cexp beta fexp x |})%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx':x <> ulp x
Hx:(bpow e < x)%R
H:(0 < Ztrunc (scaled_mantissa beta fexp x))%Z
(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx':x <> ulp x
Hx:(bpow e < x)%R
H:(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
(bpow e <= x - ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx':x <> ulp x
Hx:(bpow e < x)%R

(0 < x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx':x <> ulp x
Hx:(bpow e < x)%R
H:(0 < Ztrunc (scaled_mantissa beta fexp x))%Z
(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx':x <> ulp x
Hx:(bpow e < x)%R
H:(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
(bpow e <= x - ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx':x <> ulp x
Hx:(bpow e < x)%R

(0 <= bpow e)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx':x <> ulp x
Hx:(bpow e < x)%R
H:(0 < Ztrunc (scaled_mantissa beta fexp x))%Z
(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx':x <> ulp x
Hx:(bpow e < x)%R
H:(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
(bpow e <= x - ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx':x <> ulp x
Hx:(bpow e < x)%R
H:(0 < Ztrunc (scaled_mantissa beta fexp x))%Z

(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx':x <> ulp x
Hx:(bpow e < x)%R
H:(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
(bpow e <= x - ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx':x <> ulp x
Hx:(bpow e < x)%R
H:(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z

(bpow e <= x - ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx':x <> ulp x
Hx:(bpow e < x)%R
H:(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
Hm:(1 < Ztrunc (scaled_mantissa beta fexp x))%Z

(bpow e <= x - ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx':x <> ulp x
Hx:(bpow e < x)%R
H:(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
Hm:1%Z = Ztrunc (scaled_mantissa beta fexp x)
(bpow e <= x - ulp x)%R
(* *)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx':x <> ulp x
Hx:(bpow e < x)%R
H:(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
Hm:(1 < Ztrunc (scaled_mantissa beta fexp x))%Z

(bpow e <= F2R {| Fnum := Ztrunc (scaled_mantissa beta fexp x); Fexp := cexp beta fexp x |} - ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx':x <> ulp x
Hx:(bpow e < x)%R
H:(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
Hm:1%Z = Ztrunc (scaled_mantissa beta fexp x)
(bpow e <= x - ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx':x <> ulp x
Hx:(bpow e < x)%R
H:(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
Hm:(1 < Ztrunc (scaled_mantissa beta fexp x))%Z

(bpow e <= F2R {| Fnum := Ztrunc (scaled_mantissa beta fexp x); Fexp := cexp beta fexp x |} - bpow (cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx':x <> ulp x
Hx:(bpow e < x)%R
H:(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
Hm:(1 < Ztrunc (scaled_mantissa beta fexp x))%Z
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx':x <> ulp x
Hx:(bpow e < x)%R
H:(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
Hm:1%Z = Ztrunc (scaled_mantissa beta fexp x)
(bpow e <= x - ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx':x <> ulp x
Hx:(bpow e < x)%R
H:(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
Hm:(1 < Ztrunc (scaled_mantissa beta fexp x))%Z

(bpow e <= IZR (Fnum {| Fnum := Ztrunc (scaled_mantissa beta fexp x); Fexp := cexp beta fexp x |}) * bpow (Fexp {| Fnum := Ztrunc (scaled_mantissa beta fexp x); Fexp := cexp beta fexp x |}) - bpow (cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx':x <> ulp x
Hx:(bpow e < x)%R
H:(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
Hm:(1 < Ztrunc (scaled_mantissa beta fexp x))%Z
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx':x <> ulp x
Hx:(bpow e < x)%R
H:(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
Hm:1%Z = Ztrunc (scaled_mantissa beta fexp x)
(bpow e <= x - ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx':x <> ulp x
Hx:(bpow e < x)%R
H:(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
Hm:(1 < Ztrunc (scaled_mantissa beta fexp x))%Z

(bpow e <= IZR (Ztrunc (scaled_mantissa beta fexp x)) * bpow (cexp beta fexp x) - bpow (cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx':x <> ulp x
Hx:(bpow e < x)%R
H:(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
Hm:(1 < Ztrunc (scaled_mantissa beta fexp x))%Z
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx':x <> ulp x
Hx:(bpow e < x)%R
H:(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
Hm:1%Z = Ztrunc (scaled_mantissa beta fexp x)
(bpow e <= x - ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx':x <> ulp x
Hx:(bpow e < x)%R
H:(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
Hm:(1 < Ztrunc (scaled_mantissa beta fexp x))%Z

(bpow e <= IZR (Ztrunc (scaled_mantissa beta fexp x)) * bpow (cexp beta fexp x) - 1 * bpow (cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx':x <> ulp x
Hx:(bpow e < x)%R
H:(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
Hm:(1 < Ztrunc (scaled_mantissa beta fexp x))%Z
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx':x <> ulp x
Hx:(bpow e < x)%R
H:(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
Hm:1%Z = Ztrunc (scaled_mantissa beta fexp x)
(bpow e <= x - ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx':x <> ulp x
Hx:(bpow e < x)%R
H:(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
Hm:(1 < Ztrunc (scaled_mantissa beta fexp x))%Z

(bpow e <= (IZR (Ztrunc (scaled_mantissa beta fexp x)) - 1) * bpow (cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx':x <> ulp x
Hx:(bpow e < x)%R
H:(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
Hm:(1 < Ztrunc (scaled_mantissa beta fexp x))%Z
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx':x <> ulp x
Hx:(bpow e < x)%R
H:(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
Hm:1%Z = Ztrunc (scaled_mantissa beta fexp x)
(bpow e <= x - ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx':x <> ulp x
Hx:(bpow e < x)%R
H:(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
Hm:(1 < Ztrunc (scaled_mantissa beta fexp x))%Z

(bpow e <= IZR (Ztrunc (scaled_mantissa beta fexp x) - 1) * bpow (cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx':x <> ulp x
Hx:(bpow e < x)%R
H:(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
Hm:(1 < Ztrunc (scaled_mantissa beta fexp x))%Z
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx':x <> ulp x
Hx:(bpow e < x)%R
H:(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
Hm:1%Z = Ztrunc (scaled_mantissa beta fexp x)
(bpow e <= x - ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx':x <> ulp x
Hx:(bpow e < x)%R
H:(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
Hm:(1 < Ztrunc (scaled_mantissa beta fexp x))%Z

(1 < Ztrunc (scaled_mantissa beta fexp x))%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx':x <> ulp x
Hx:(bpow e < x)%R
H:(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
Hm:(1 < Ztrunc (scaled_mantissa beta fexp x))%Z
(bpow e < F2R {| Fnum := Ztrunc (scaled_mantissa beta fexp x); Fexp := fexp (mag beta x) |})%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx':x <> ulp x
Hx:(bpow e < x)%R
H:(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
Hm:(1 < Ztrunc (scaled_mantissa beta fexp x))%Z
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx':x <> ulp x
Hx:(bpow e < x)%R
H:(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
Hm:1%Z = Ztrunc (scaled_mantissa beta fexp x)
(bpow e <= x - ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx':x <> ulp x
Hx:(bpow e < x)%R
H:(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
Hm:(1 < Ztrunc (scaled_mantissa beta fexp x))%Z

(bpow e < F2R {| Fnum := Ztrunc (scaled_mantissa beta fexp x); Fexp := fexp (mag beta x) |})%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx':x <> ulp x
Hx:(bpow e < x)%R
H:(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
Hm:(1 < Ztrunc (scaled_mantissa beta fexp x))%Z
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx':x <> ulp x
Hx:(bpow e < x)%R
H:(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
Hm:1%Z = Ztrunc (scaled_mantissa beta fexp x)
(bpow e <= x - ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx':x <> ulp x
Hx:(bpow e < x)%R
H:(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
Hm:(1 < Ztrunc (scaled_mantissa beta fexp x))%Z

x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx':x <> ulp x
Hx:(bpow e < x)%R
H:(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
Hm:1%Z = Ztrunc (scaled_mantissa beta fexp x)
(bpow e <= x - ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx':x <> ulp x
Hx:(bpow e < x)%R
H:(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
Hm:(1 < Ztrunc (scaled_mantissa beta fexp x))%Z

(0 < x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx':x <> ulp x
Hx:(bpow e < x)%R
H:(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
Hm:1%Z = Ztrunc (scaled_mantissa beta fexp x)
(bpow e <= x - ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx':x <> ulp x
Hx:(bpow e < x)%R
H:(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
Hm:1%Z = Ztrunc (scaled_mantissa beta fexp x)

(bpow e <= x - ulp x)%R
(* *)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx:(bpow e < x)%R
H:(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
Hm:1%Z = Ztrunc (scaled_mantissa beta fexp x)

x = ulp x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx:(bpow e < x)%R
H:(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
Hm:1%Z = Ztrunc (scaled_mantissa beta fexp x)

F2R {| Fnum := Ztrunc (scaled_mantissa beta fexp x); Fexp := cexp beta fexp x |} = ulp x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx:(bpow e < x)%R
H:(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
Hm:1%Z = Ztrunc (scaled_mantissa beta fexp x)

F2R {| Fnum := 1; Fexp := cexp beta fexp x |} = ulp x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx:(bpow e < x)%R
H:(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
Hm:1%Z = Ztrunc (scaled_mantissa beta fexp x)

F2R {| Fnum := 1; Fexp := cexp beta fexp x |} = bpow (cexp beta fexp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx:(bpow e < x)%R
H:(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
Hm:1%Z = Ztrunc (scaled_mantissa beta fexp x)
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx:(bpow e < x)%R
H:(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
Hm:1%Z = Ztrunc (scaled_mantissa beta fexp x)

(1 * bpow (cexp beta fexp x))%R = bpow (cexp beta fexp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx:(bpow e < x)%R
H:(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
Hm:1%Z = Ztrunc (scaled_mantissa beta fexp x)
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx:(bpow e < x)%R
H:(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
Hm:1%Z = Ztrunc (scaled_mantissa beta fexp x)

x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Fx:F x
Hx:(bpow e < x)%R
H:(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
Hm:1%Z = Ztrunc (scaled_mantissa beta fexp x)

(0 < x)%R
apply Rlt_trans with (2:=Hx), bpow_gt_0. Qed. (* was succ_le_bpow *)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall (x : R) (e : Z), (0 < x)%R -> F x -> (x < bpow e)%R -> (x + ulp x <= bpow e)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall (x : R) (e : Z), (0 < x)%R -> F x -> (x < bpow e)%R -> (x + ulp x <= bpow e)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Zx:(0 < x)%R
Fx:F x
Hx:(x < bpow e)%R

(x + ulp x <= bpow e)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Zx:(0 < x)%R
Fx:F x
Hx:(x < bpow e)%R

(F2R {| Fnum := Ztrunc (scaled_mantissa beta fexp x); Fexp := cexp beta fexp x |} + ulp x <= bpow e)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Zx:(0 < x)%R
Fx:F x
Hx:(x < bpow e)%R

(F2R {| Fnum := Ztrunc (scaled_mantissa beta fexp x); Fexp := cexp beta fexp x |} + bpow (cexp beta fexp x) <= bpow e)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Zx:(0 < x)%R
Fx:F x
Hx:(x < bpow e)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Zx:(0 < x)%R
Fx:F x
Hx:(x < bpow e)%R

(IZR (Fnum {| Fnum := Ztrunc (scaled_mantissa beta fexp x); Fexp := cexp beta fexp x |}) * bpow (Fexp {| Fnum := Ztrunc (scaled_mantissa beta fexp x); Fexp := cexp beta fexp x |}) + bpow (cexp beta fexp x) <= bpow e)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Zx:(0 < x)%R
Fx:F x
Hx:(x < bpow e)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Zx:(0 < x)%R
Fx:F x
Hx:(x < bpow e)%R

(IZR (Ztrunc (scaled_mantissa beta fexp x)) * bpow (cexp beta fexp x) + bpow (cexp beta fexp x) <= bpow e)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Zx:(0 < x)%R
Fx:F x
Hx:(x < bpow e)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Zx:(0 < x)%R
Fx:F x
Hx:(x < bpow e)%R

(IZR (Ztrunc (scaled_mantissa beta fexp x)) * bpow (cexp beta fexp x) + 1 * bpow (cexp beta fexp x) <= bpow e)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Zx:(0 < x)%R
Fx:F x
Hx:(x < bpow e)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Zx:(0 < x)%R
Fx:F x
Hx:(x < bpow e)%R

((IZR (Ztrunc (scaled_mantissa beta fexp x)) + 1) * bpow (cexp beta fexp x) <= bpow e)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Zx:(0 < x)%R
Fx:F x
Hx:(x < bpow e)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Zx:(0 < x)%R
Fx:F x
Hx:(x < bpow e)%R

(IZR (Ztrunc (scaled_mantissa beta fexp x) + 1) * bpow (cexp beta fexp x) <= bpow e)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Zx:(0 < x)%R
Fx:F x
Hx:(x < bpow e)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Zx:(0 < x)%R
Fx:F x
Hx:(x < bpow e)%R

(0 < Ztrunc (scaled_mantissa beta fexp x))%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Zx:(0 < x)%R
Fx:F x
Hx:(x < bpow e)%R
(F2R {| Fnum := Ztrunc (scaled_mantissa beta fexp x); Fexp := fexp (mag beta x) |} < bpow e)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Zx:(0 < x)%R
Fx:F x
Hx:(x < bpow e)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Zx:(0 < x)%R
Fx:F x
Hx:(x < bpow e)%R

(0 < F2R {| Fnum := Ztrunc (scaled_mantissa beta fexp x); Fexp := cexp beta fexp x |})%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Zx:(0 < x)%R
Fx:F x
Hx:(x < bpow e)%R
(F2R {| Fnum := Ztrunc (scaled_mantissa beta fexp x); Fexp := fexp (mag beta x) |} < bpow e)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Zx:(0 < x)%R
Fx:F x
Hx:(x < bpow e)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Zx:(0 < x)%R
Fx:F x
Hx:(x < bpow e)%R

(F2R {| Fnum := Ztrunc (scaled_mantissa beta fexp x); Fexp := fexp (mag beta x) |} < bpow e)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Zx:(0 < x)%R
Fx:F x
Hx:(x < bpow e)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
e:Z
Zx:(0 < x)%R
Fx:F x
Hx:(x < bpow e)%R

x <> 0%R
now apply Rgt_not_eq. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, (0 < x)%R -> F x -> x <> bpow (mag beta x - 1) -> F (x - ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, (0 < x)%R -> F x -> x <> bpow (mag beta x - 1) -> F (x - ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
Hx:x <> bpow (mag beta x - 1)

F (x - ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (Build_mag_prop beta x ex Ex - 1)

F (x - ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)

F (x - ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)

F (x - ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)

(bpow (ex - 1) < x < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
F (x - ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
Hx:x <> bpow (ex - 1)

(bpow (ex - 1) < x < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
F (x - ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
H:bpow (ex - 1) = x
H':(x < bpow ex)%R
Hx:x <> bpow (ex - 1)

(bpow (ex - 1) < x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
F (x - ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)

(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
F (x - ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R

F (x - ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R

(x - ulp x)%R = F2R {| Fnum := Ztrunc ((x - ulp x) * bpow (- fexp (mag beta (x - ulp x)))); Fexp := fexp (mag beta (x - ulp x)) |}
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R

(x - ulp x)%R = F2R {| Fnum := Ztrunc ((x - ulp x) * bpow (- fexp ex)); Fexp := fexp ex |}
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
(bpow (ex - 1) <= Rabs (x - ulp x) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R

(F2R {| Fnum := Ztrunc (scaled_mantissa beta fexp x); Fexp := cexp beta fexp x |} - ulp x)%R = F2R {| Fnum := Ztrunc ((F2R {| Fnum := Ztrunc (scaled_mantissa beta fexp x); Fexp := cexp beta fexp x |} - ulp x) * bpow (- fexp ex)); Fexp := fexp ex |}
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
(bpow (ex - 1) <= Rabs (x - ulp x) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R

(F2R {| Fnum := Ztrunc (scaled_mantissa beta fexp x); Fexp := cexp beta fexp x |} - bpow (cexp beta fexp x))%R = F2R {| Fnum := Ztrunc ((F2R {| Fnum := Ztrunc (scaled_mantissa beta fexp x); Fexp := cexp beta fexp x |} - bpow (cexp beta fexp x)) * bpow (- fexp ex)); Fexp := fexp ex |}
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
(bpow (ex - 1) <= Rabs (x - ulp x) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R

(F2R {| Fnum := Ztrunc (x * bpow (- cexp beta fexp x)); Fexp := cexp beta fexp x |} - bpow (cexp beta fexp x))%R = F2R {| Fnum := Ztrunc ((F2R {| Fnum := Ztrunc (x * bpow (- cexp beta fexp x)); Fexp := cexp beta fexp x |} - bpow (cexp beta fexp x)) * bpow (- fexp ex)); Fexp := fexp ex |}
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
(bpow (ex - 1) <= Rabs (x - ulp x) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R

(F2R {| Fnum := Ztrunc (x * bpow (- fexp ex)); Fexp := fexp ex |} - bpow (fexp ex))%R = F2R {| Fnum := Ztrunc ((F2R {| Fnum := Ztrunc (x * bpow (- fexp ex)); Fexp := fexp ex |} - bpow (fexp ex)) * bpow (- fexp ex)); Fexp := fexp ex |}
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
(bpow (ex - 1) <= Rabs (x - ulp x) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R

(IZR (Fnum {| Fnum := Ztrunc (x * bpow (- fexp ex)); Fexp := fexp ex |}) * bpow (Fexp {| Fnum := Ztrunc (x * bpow (- fexp ex)); Fexp := fexp ex |}) - bpow (fexp ex))%R = (IZR (Fnum {| Fnum := Ztrunc ((IZR (Fnum {| Fnum := Ztrunc (x * bpow (- fexp ex)); Fexp := fexp ex |}) * bpow (Fexp {| Fnum := Ztrunc (x * bpow (- fexp ex)); Fexp := fexp ex |}) - bpow (fexp ex)) * bpow (- fexp ex)); Fexp := fexp ex |}) * bpow (Fexp {| Fnum := Ztrunc ((IZR (Fnum {| Fnum := Ztrunc (x * bpow (- fexp ex)); Fexp := fexp ex |}) * bpow (Fexp {| Fnum := Ztrunc (x * bpow (- fexp ex)); Fexp := fexp ex |}) - bpow (fexp ex)) * bpow (- fexp ex)); Fexp := fexp ex |}))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
(bpow (ex - 1) <= Rabs (x - ulp x) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R

(IZR (Ztrunc (x * bpow (- fexp ex))) * bpow (fexp ex) - bpow (fexp ex))%R = (IZR (Ztrunc ((IZR (Ztrunc (x * bpow (- fexp ex))) * bpow (fexp ex) - bpow (fexp ex)) * bpow (- fexp ex))) * bpow (fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
(bpow (ex - 1) <= Rabs (x - ulp x) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R

(IZR (Ztrunc (x * bpow (- fexp ex))) * bpow (fexp ex) - bpow (fexp ex))%R = (IZR (Ztrunc (IZR (Ztrunc (x * bpow (- fexp ex))) * bpow (fexp ex) * bpow (- fexp ex) - bpow (fexp ex) * bpow (- fexp ex))) * bpow (fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
(bpow (ex - 1) <= Rabs (x - ulp x) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R

(IZR (Ztrunc (x * bpow (- fexp ex))) * bpow (fexp ex) - bpow (fexp ex))%R = (IZR (Ztrunc (IZR (Ztrunc (x * bpow (- fexp ex))) * (bpow (fexp ex) * bpow (- fexp ex)) - bpow (fexp ex) * bpow (- fexp ex))) * bpow (fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
(bpow (ex - 1) <= Rabs (x - ulp x) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R

(IZR (Ztrunc (x * bpow (- fexp ex))) * bpow (fexp ex) - bpow (fexp ex))%R = (IZR (Ztrunc (IZR (Ztrunc (x * bpow (- fexp ex))) - bpow 0)) * bpow (fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
(bpow (ex - 1) <= Rabs (x - ulp x) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R

(IZR (Ztrunc (x * bpow (- fexp ex))) * bpow (fexp ex) - bpow (fexp ex))%R = (IZR (Ztrunc (IZR (Ztrunc (x * bpow (- fexp ex))) - 1)) * bpow (fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
(bpow (ex - 1) <= Rabs (x - ulp x) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R

(IZR (Ztrunc (x * bpow (- fexp ex))) * bpow (fexp ex) - bpow (fexp ex))%R = (IZR (Ztrunc (IZR (Ztrunc (x * bpow (- fexp ex)) - 1))) * bpow (fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
(bpow (ex - 1) <= Rabs (x - ulp x) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R

(IZR (Ztrunc (x * bpow (- fexp ex))) * bpow (fexp ex) - bpow (fexp ex))%R = (IZR (Ztrunc (x * bpow (- fexp ex)) - 1) * bpow (fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
(bpow (ex - 1) <= Rabs (x - ulp x) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R

(IZR (Ztrunc (x * bpow (- fexp ex))) * bpow (fexp ex) - bpow (fexp ex))%R = ((IZR (Ztrunc (x * bpow (- fexp ex))) - 1) * bpow (fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
(bpow (ex - 1) <= Rabs (x - ulp x) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R

(IZR (Ztrunc (x * bpow (- fexp ex))) * bpow (fexp ex) - bpow (fexp ex))%R = (IZR (Ztrunc (x * bpow (- fexp ex))) * bpow (fexp ex) - 1 * bpow (fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
(bpow (ex - 1) <= Rabs (x - ulp x) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R

x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
(bpow (ex - 1) <= Rabs (x - ulp x) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R

(bpow (ex - 1) <= Rabs (x - ulp x) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R

(bpow (ex - 1) <= x - ulp x < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
(0 <= x - ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R

(bpow (ex - 1) <= x - ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
(x - ulp x < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
(0 <= x - ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R

x <> ulp x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
(bpow (ex - 1) < x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
(x - ulp x < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
(0 <= x - ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R

x <> bpow (cexp beta fexp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
(bpow (ex - 1) < x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
(x - ulp x < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
(0 <= x - ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
H:x = bpow (cexp beta fexp x)

False
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
(bpow (ex - 1) < x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
(x - ulp x < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
(0 <= x - ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
H:x = bpow (cexp beta fexp x)

(ex - 1 < cexp beta fexp x < ex)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
H:x = bpow (cexp beta fexp x)
H0:(ex - 1 < cexp beta fexp x < ex)%Z
False
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
(bpow (ex - 1) < x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
(x - ulp x < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
(0 <= x - ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
H:x = bpow (cexp beta fexp x)
H0:(ex - 1 < cexp beta fexp x < ex)%Z

False
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
(bpow (ex - 1) < x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
(x - ulp x < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
(0 <= x - ulp x)%R
beta:radix
fexp:Z -> Z
x:R
ex:Z
H0:(ex - 1 < cexp beta fexp x < ex)%Z

False
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
(bpow (ex - 1) < x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
(x - ulp x < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
(0 <= x - ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R

x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
(bpow (ex - 1) < x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
(x - ulp x < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
(0 <= x - ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R

(bpow (ex - 1) < x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
(x - ulp x < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
(0 <= x - ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R

(x - ulp x < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
(0 <= x - ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R

(x - ulp x <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
(0 <= x - ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R

(x - ulp x <= x + 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
(0 <= x - ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R

(- ulp x <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
(0 <= x - ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R

(- ulp x <= - 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
(0 <= x - ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R

(0 <= ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
(0 <= x - ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R

(0 <= x - ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R

(ulp x <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R

(ulp x <= F2R {| Fnum := Ztrunc (scaled_mantissa beta fexp x); Fexp := cexp beta fexp x |})%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R

(bpow (cexp beta fexp x) <= F2R {| Fnum := Ztrunc (scaled_mantissa beta fexp x); Fexp := cexp beta fexp x |})%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R

(bpow (cexp beta fexp x) <= IZR (Ztrunc (scaled_mantissa beta fexp x)) * bpow (cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R

(1 * bpow (cexp beta fexp x) <= IZR (Ztrunc (scaled_mantissa beta fexp x)) * bpow (cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R

(0 <= bpow (cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
(1 <= IZR (Ztrunc (scaled_mantissa beta fexp x)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R

(1 <= IZR (Ztrunc (scaled_mantissa beta fexp x)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R

(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R

(0 < Ztrunc (scaled_mantissa beta fexp x))%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
H:(0 < Ztrunc (scaled_mantissa beta fexp x))%Z
(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R

(0 < F2R {| Fnum := Ztrunc (scaled_mantissa beta fexp x); Fexp := cexp beta fexp x |})%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
H:(0 < Ztrunc (scaled_mantissa beta fexp x))%Z
(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R

(0 < x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
H:(0 < Ztrunc (scaled_mantissa beta fexp x))%Z
(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R

(0 <= bpow (ex - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
H:(0 < Ztrunc (scaled_mantissa beta fexp x))%Z
(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
H:(0 < Ztrunc (scaled_mantissa beta fexp x))%Z

(1 <= Ztrunc (scaled_mantissa beta fexp x))%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:x <> bpow (ex - 1)
Ex':(bpow (ex - 1) < x < bpow ex)%R

x <> 0%R
now apply Rgt_not_eq. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, (0 < x)%R -> F x -> let e := mag_val beta x (mag beta x) in x = bpow (e - 1) -> F (x - bpow (fexp (e - 1)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, (0 < x)%R -> F x -> let e := mag_val beta x (mag beta x) in x = bpow (e - 1) -> F (x - bpow (fexp (e - 1)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)

F (x - bpow (fexp (e - 1)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R

F (x - bpow (fexp (e - 1)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R

F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R

(fexp (e - 1) <= e - 1)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) <= e - 1)%Z
F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R

F (bpow (e - 1))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) <= e - 1)%Z
F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) <= e - 1)%Z

F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z

F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:fexp (e - 1) = (e - 1)%Z
F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z

f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:fexp (e - 1) = (e - 1)%Z
F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z

(bpow (e - 1) - bpow (fexp (e - 1)))%R = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:fexp (e - 1) = (e - 1)%Z
F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z

(bpow (e - 1) - bpow (fexp (e - 1)))%R = (IZR (beta ^ (e - 1 - fexp (e - 1)) - 1) * bpow (fexp (e - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:fexp (e - 1) = (e - 1)%Z
F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z

(bpow (e - 1) - bpow (fexp (e - 1)))%R = ((bpow (e - 1 - fexp (e - 1)) - 1) * bpow (fexp (e - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
(0 <= e - 1 - fexp (e - 1))%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:fexp (e - 1) = (e - 1)%Z
F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z

(bpow (e - 1) - bpow (fexp (e - 1)))%R = (bpow (e - 1 - fexp (e - 1)) * bpow (fexp (e - 1)) - bpow (fexp (e - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
(0 <= e - 1 - fexp (e - 1))%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:fexp (e - 1) = (e - 1)%Z
F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z

(bpow (e - 1) - bpow (fexp (e - 1)))%R = (bpow (e - 1 - fexp (e - 1) + fexp (e - 1)) - bpow (fexp (e - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
(0 <= e - 1 - fexp (e - 1))%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:fexp (e - 1) = (e - 1)%Z
F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z

(0 <= e - 1 - fexp (e - 1))%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:fexp (e - 1) = (e - 1)%Z
F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}

F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:fexp (e - 1) = (e - 1)%Z
F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}

F (F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |})
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:fexp (e - 1) = (e - 1)%Z
F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}

(beta ^ (e - 1 - fexp (e - 1)) - 1)%Z <> 0%Z -> (cexp beta fexp (F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}) <= fexp (e - 1))%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:fexp (e - 1) = (e - 1)%Z
F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}

(cexp beta fexp (F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}) <= fexp (e - 1))%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:fexp (e - 1) = (e - 1)%Z
F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}

cexp beta fexp (F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}) = fexp (e - 1)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:fexp (e - 1) = (e - 1)%Z
F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}

(bpow (e - 1 - 1) <= Rabs (F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}) < bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:fexp (e - 1) = (e - 1)%Z
F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}

(bpow (e - 1 - 1) <= Rabs f < bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:fexp (e - 1) = (e - 1)%Z
F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}

(bpow (e - 1 - 1) <= Rabs (bpow (e - 1) - bpow (fexp (e - 1))) < bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:fexp (e - 1) = (e - 1)%Z
F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}

(bpow (e - 1 - 1) <= bpow (e - 1) - bpow (fexp (e - 1)) < bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
(bpow (e - 1) - bpow (fexp (e - 1)) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:fexp (e - 1) = (e - 1)%Z
F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}

(bpow (e - 1 - 1) <= bpow (e - 1) - bpow (fexp (e - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
(bpow (e - 1) - bpow (fexp (e - 1)) < bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
(bpow (e - 1) - bpow (fexp (e - 1)) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:fexp (e - 1) = (e - 1)%Z
F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}

(bpow (fexp (e - 1)) + bpow (e - 1 - 1) <= bpow (fexp (e - 1)) + (bpow (e - 1) - bpow (fexp (e - 1))))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
(bpow (e - 1) - bpow (fexp (e - 1)) < bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
(bpow (e - 1) - bpow (fexp (e - 1)) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:fexp (e - 1) = (e - 1)%Z
F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}

(bpow (fexp (e - 1)) + bpow (e - 1 - 1) <= bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
(bpow (e - 1) - bpow (fexp (e - 1)) < bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
(bpow (e - 1) - bpow (fexp (e - 1)) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:fexp (e - 1) = (e - 1)%Z
F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}

(bpow (fexp (e - 1)) + bpow (e - 1 - 1) <= bpow (e - 2) + bpow (e - 2))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
(bpow (e - 2) + bpow (e - 2) <= bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
(bpow (e - 1) - bpow (fexp (e - 1)) < bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
(bpow (e - 1) - bpow (fexp (e - 1)) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:fexp (e - 1) = (e - 1)%Z
F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}

(bpow (e - 2) + bpow (e - 2) <= bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
(bpow (e - 1) - bpow (fexp (e - 1)) < bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
(bpow (e - 1) - bpow (fexp (e - 1)) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:fexp (e - 1) = (e - 1)%Z
F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}

(2 * bpow (e - 2) <= bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
(bpow (e - 1) - bpow (fexp (e - 1)) < bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
(bpow (e - 1) - bpow (fexp (e - 1)) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:fexp (e - 1) = (e - 1)%Z
F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}

(2 * bpow (e - 2) <= bpow 1 * bpow (e - 2))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
(bpow 1 * bpow (e - 2) <= bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
(bpow (e - 1) - bpow (fexp (e - 1)) < bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
(bpow (e - 1) - bpow (fexp (e - 1)) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:fexp (e - 1) = (e - 1)%Z
F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}

(0 <= bpow (e - 2))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
(2 <= bpow 1)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
(bpow 1 * bpow (e - 2) <= bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
(bpow (e - 1) - bpow (fexp (e - 1)) < bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
(bpow (e - 1) - bpow (fexp (e - 1)) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:fexp (e - 1) = (e - 1)%Z
F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}

(2 <= bpow 1)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
(bpow 1 * bpow (e - 2) <= bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
(bpow (e - 1) - bpow (fexp (e - 1)) < bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
(bpow (e - 1) - bpow (fexp (e - 1)) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:fexp (e - 1) = (e - 1)%Z
F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}

(2 <= IZR beta)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
IZR beta = bpow 1
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
(bpow 1 * bpow (e - 2) <= bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
(bpow (e - 1) - bpow (fexp (e - 1)) < bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
(bpow (e - 1) - bpow (fexp (e - 1)) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:fexp (e - 1) = (e - 1)%Z
F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}

(2 <= beta)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
IZR beta = bpow 1
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
(bpow 1 * bpow (e - 2) <= bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
(bpow (e - 1) - bpow (fexp (e - 1)) < bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
(bpow (e - 1) - bpow (fexp (e - 1)) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:fexp (e - 1) = (e - 1)%Z
F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}

(2 <=? beta)%Z = true
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
IZR beta = bpow 1
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
(bpow 1 * bpow (e - 2) <= bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
(bpow (e - 1) - bpow (fexp (e - 1)) < bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
(bpow (e - 1) - bpow (fexp (e - 1)) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:fexp (e - 1) = (e - 1)%Z
F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}

IZR beta = bpow 1
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
(bpow 1 * bpow (e - 2) <= bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
(bpow (e - 1) - bpow (fexp (e - 1)) < bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
(bpow (e - 1) - bpow (fexp (e - 1)) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:fexp (e - 1) = (e - 1)%Z
F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}

IZR beta = IZR (Z.pow_pos beta 1)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
(bpow 1 * bpow (e - 2) <= bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
(bpow (e - 1) - bpow (fexp (e - 1)) < bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
(bpow (e - 1) - bpow (fexp (e - 1)) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:fexp (e - 1) = (e - 1)%Z
F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}

IZR beta = IZR (beta * 1)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
(bpow 1 * bpow (e - 2) <= bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
(bpow (e - 1) - bpow (fexp (e - 1)) < bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
(bpow (e - 1) - bpow (fexp (e - 1)) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:fexp (e - 1) = (e - 1)%Z
F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}

(bpow 1 * bpow (e - 2) <= bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
(bpow (e - 1) - bpow (fexp (e - 1)) < bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
(bpow (e - 1) - bpow (fexp (e - 1)) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:fexp (e - 1) = (e - 1)%Z
F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}

(bpow (1 + (e - 2)) <= bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
(bpow (e - 1) - bpow (fexp (e - 1)) < bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
(bpow (e - 1) - bpow (fexp (e - 1)) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:fexp (e - 1) = (e - 1)%Z
F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}

(bpow (e - 1) <= bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
(bpow (e - 1) - bpow (fexp (e - 1)) < bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
(bpow (e - 1) - bpow (fexp (e - 1)) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:fexp (e - 1) = (e - 1)%Z
F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}

(bpow (e - 1) - bpow (fexp (e - 1)) < bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
(bpow (e - 1) - bpow (fexp (e - 1)) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:fexp (e - 1) = (e - 1)%Z
F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}

(bpow (e - 1) - bpow (fexp (e - 1)) < bpow (e - 1) + 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
(bpow (e - 1) - bpow (fexp (e - 1)) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:fexp (e - 1) = (e - 1)%Z
F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}

(- bpow (fexp (e - 1)) < 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
(bpow (e - 1) - bpow (fexp (e - 1)) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:fexp (e - 1) = (e - 1)%Z
F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}

(- bpow (fexp (e - 1)) < - 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
(bpow (e - 1) - bpow (fexp (e - 1)) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:fexp (e - 1) = (e - 1)%Z
F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}

(0 < bpow (fexp (e - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}
(bpow (e - 1) - bpow (fexp (e - 1)) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:fexp (e - 1) = (e - 1)%Z
F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}

(bpow (e - 1) - bpow (fexp (e - 1)) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:fexp (e - 1) = (e - 1)%Z
F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}

(bpow (fexp (e - 1)) <= bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:fexp (e - 1) = (e - 1)%Z
F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:(fexp (e - 1) < e - 1)%Z
H:f = F2R {| Fnum := beta ^ (e - 1 - fexp (e - 1)) - 1; Fexp := fexp (e - 1) |}

(fexp (e - 1) <= e - 1)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:fexp (e - 1) = (e - 1)%Z
F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:fexp (e - 1) = (e - 1)%Z

F f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:fexp (e - 1) = (e - 1)%Z

F 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:fexp (e - 1) = (e - 1)%Z
0%R = f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:fexp (e - 1) = (e - 1)%Z

0%R = f
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:fexp (e - 1) = (e - 1)%Z

0%R = (x - bpow (fexp (e - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hx:x = bpow (e - 1)
f:=(x - bpow (fexp (e - 1)))%R:R
He:fexp (e - 1) = (e - 1)%Z

0%R = (bpow (e - 1) - bpow (e - 1))%R
ring. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, (0 < x)%R -> F x -> F (x + ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, (0 < x)%R -> F x -> F (x + ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x

F (x + ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R

F (x + ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R

F (x + ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex, Ex':(bpow (ex - 1) <= Rabs x < bpow ex)%R

F (x + ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R

F (x + ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex, Ex':(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R

F (x + ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex
F (x + ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex, Ex':(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R

(x + ulp x)%R = F2R {| Fnum := Ztrunc ((x + ulp x) * bpow (- fexp (mag beta (x + ulp x)))); Fexp := fexp (mag beta (x + ulp x)) |}
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex
F (x + ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex, Ex':(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R

(x + ulp x)%R = F2R {| Fnum := Ztrunc ((x + ulp x) * bpow (- fexp ex)); Fexp := fexp ex |}
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R
(bpow (ex - 1) <= Rabs (x + ulp x) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex
F (x + ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex, Ex':(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R

(F2R {| Fnum := Ztrunc (scaled_mantissa beta fexp x); Fexp := cexp beta fexp x |} + ulp x)%R = F2R {| Fnum := Ztrunc ((F2R {| Fnum := Ztrunc (scaled_mantissa beta fexp x); Fexp := cexp beta fexp x |} + ulp x) * bpow (- fexp ex)); Fexp := fexp ex |}
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R
(bpow (ex - 1) <= Rabs (x + ulp x) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex
F (x + ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex, Ex':(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R

(F2R {| Fnum := Ztrunc (scaled_mantissa beta fexp x); Fexp := cexp beta fexp x |} + bpow (cexp beta fexp x))%R = F2R {| Fnum := Ztrunc ((F2R {| Fnum := Ztrunc (scaled_mantissa beta fexp x); Fexp := cexp beta fexp x |} + bpow (cexp beta fexp x)) * bpow (- fexp ex)); Fexp := fexp ex |}
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R
(bpow (ex - 1) <= Rabs (x + ulp x) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex
F (x + ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex, Ex':(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R

(F2R {| Fnum := Ztrunc (x * bpow (- cexp beta fexp x)); Fexp := cexp beta fexp x |} + bpow (cexp beta fexp x))%R = F2R {| Fnum := Ztrunc ((F2R {| Fnum := Ztrunc (x * bpow (- cexp beta fexp x)); Fexp := cexp beta fexp x |} + bpow (cexp beta fexp x)) * bpow (- fexp ex)); Fexp := fexp ex |}
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R
(bpow (ex - 1) <= Rabs (x + ulp x) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex
F (x + ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex, Ex':(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R

(F2R {| Fnum := Ztrunc (x * bpow (- fexp ex)); Fexp := fexp ex |} + bpow (fexp ex))%R = F2R {| Fnum := Ztrunc ((F2R {| Fnum := Ztrunc (x * bpow (- fexp ex)); Fexp := fexp ex |} + bpow (fexp ex)) * bpow (- fexp ex)); Fexp := fexp ex |}
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R
(bpow (ex - 1) <= Rabs (x + ulp x) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex
F (x + ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex, Ex':(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R

(IZR (Fnum {| Fnum := Ztrunc (x * bpow (- fexp ex)); Fexp := fexp ex |}) * bpow (Fexp {| Fnum := Ztrunc (x * bpow (- fexp ex)); Fexp := fexp ex |}) + bpow (fexp ex))%R = (IZR (Fnum {| Fnum := Ztrunc ((IZR (Fnum {| Fnum := Ztrunc (x * bpow (- fexp ex)); Fexp := fexp ex |}) * bpow (Fexp {| Fnum := Ztrunc (x * bpow (- fexp ex)); Fexp := fexp ex |}) + bpow (fexp ex)) * bpow (- fexp ex)); Fexp := fexp ex |}) * bpow (Fexp {| Fnum := Ztrunc ((IZR (Fnum {| Fnum := Ztrunc (x * bpow (- fexp ex)); Fexp := fexp ex |}) * bpow (Fexp {| Fnum := Ztrunc (x * bpow (- fexp ex)); Fexp := fexp ex |}) + bpow (fexp ex)) * bpow (- fexp ex)); Fexp := fexp ex |}))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R
(bpow (ex - 1) <= Rabs (x + ulp x) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex
F (x + ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex, Ex':(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R

(IZR (Ztrunc (x * bpow (- fexp ex))) * bpow (fexp ex) + bpow (fexp ex))%R = (IZR (Ztrunc ((IZR (Ztrunc (x * bpow (- fexp ex))) * bpow (fexp ex) + bpow (fexp ex)) * bpow (- fexp ex))) * bpow (fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R
(bpow (ex - 1) <= Rabs (x + ulp x) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex
F (x + ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex, Ex':(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R

(IZR (Ztrunc (x * bpow (- fexp ex))) * bpow (fexp ex) + bpow (fexp ex))%R = (IZR (Ztrunc (IZR (Ztrunc (x * bpow (- fexp ex))) * bpow (fexp ex) * bpow (- fexp ex) + bpow (fexp ex) * bpow (- fexp ex))) * bpow (fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R
(bpow (ex - 1) <= Rabs (x + ulp x) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex
F (x + ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex, Ex':(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R

(IZR (Ztrunc (x * bpow (- fexp ex))) * bpow (fexp ex) + bpow (fexp ex))%R = (IZR (Ztrunc (IZR (Ztrunc (x * bpow (- fexp ex))) * (bpow (fexp ex) * bpow (- fexp ex)) + bpow (fexp ex) * bpow (- fexp ex))) * bpow (fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R
(bpow (ex - 1) <= Rabs (x + ulp x) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex
F (x + ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex, Ex':(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R

(IZR (Ztrunc (x * bpow (- fexp ex))) * bpow (fexp ex) + bpow (fexp ex))%R = (IZR (Ztrunc (IZR (Ztrunc (x * bpow (- fexp ex))) + bpow 0)) * bpow (fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R
(bpow (ex - 1) <= Rabs (x + ulp x) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex
F (x + ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex, Ex':(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R

(IZR (Ztrunc (x * bpow (- fexp ex))) * bpow (fexp ex) + bpow (fexp ex))%R = (IZR (Ztrunc (IZR (Ztrunc (x * bpow (- fexp ex))) + 1)) * bpow (fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R
(bpow (ex - 1) <= Rabs (x + ulp x) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex
F (x + ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex, Ex':(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R

(IZR (Ztrunc (x * bpow (- fexp ex))) * bpow (fexp ex) + bpow (fexp ex))%R = (IZR (Ztrunc (IZR (Ztrunc (x * bpow (- fexp ex)) + 1))) * bpow (fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R
(bpow (ex - 1) <= Rabs (x + ulp x) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex
F (x + ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex, Ex':(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R

(IZR (Ztrunc (x * bpow (- fexp ex))) * bpow (fexp ex) + bpow (fexp ex))%R = (IZR (Ztrunc (x * bpow (- fexp ex)) + 1) * bpow (fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R
(bpow (ex - 1) <= Rabs (x + ulp x) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex
F (x + ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex, Ex':(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R

(IZR (Ztrunc (x * bpow (- fexp ex))) * bpow (fexp ex) + bpow (fexp ex))%R = ((IZR (Ztrunc (x * bpow (- fexp ex))) + 1) * bpow (fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R
(bpow (ex - 1) <= Rabs (x + ulp x) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex
F (x + ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex, Ex':(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R

(IZR (Ztrunc (x * bpow (- fexp ex))) * bpow (fexp ex) + bpow (fexp ex))%R = (IZR (Ztrunc (x * bpow (- fexp ex))) * bpow (fexp ex) + 1 * bpow (fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R
(bpow (ex - 1) <= Rabs (x + ulp x) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex
F (x + ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex, Ex':(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R

x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R
(bpow (ex - 1) <= Rabs (x + ulp x) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex
F (x + ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex, Ex':(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R

(bpow (ex - 1) <= Rabs (x + ulp x) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex
F (x + ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex, Ex':(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R

(bpow (ex - 1) <= x + ulp x < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R
(0 <= x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex
F (x + ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex, Ex':(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R

(bpow (ex - 1) <= x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R
(x + ulp x < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R
(0 <= x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex
F (x + ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex, Ex':(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R

(x <= x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R
(x + ulp x < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R
(0 <= x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex
F (x + ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex, Ex':(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R

(x + 0 <= x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R
(x + ulp x < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R
(0 <= x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex
F (x + ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex, Ex':(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R

(0 <= ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R
(x + ulp x < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R
(0 <= x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex
F (x + ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex, Ex':(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R

(x + ulp x < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R
(0 <= x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex
F (x + ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex, Ex':(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R

(0 <= x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex
F (x + ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex, Ex':(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R

(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R
(0 <= ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex
F (x + ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex, Ex':(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x < bpow ex)%R

(0 <= ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex
F (x + ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex, Ex':(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex

F (x + ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex, Ex':(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex

F (bpow ex)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex, Ex':(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex

(fexp (ex + 1) <= ex)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex, Ex':(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex

(fexp ex < ex)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex, Ex':(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex
H0:(ex <= fexp ex)%Z

(fexp ex < ex)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex, Ex':(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex
H0:(ex <= fexp ex)%Z

(x <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex, Ex':(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex
H0:(ex <= fexp ex)%Z

(F2R {| Fnum := Ztrunc (scaled_mantissa beta fexp x); Fexp := cexp beta fexp x |} <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex, Ex':(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex
H0:(ex <= fexp ex)%Z

(F2R {| Fnum := 0; Fexp := cexp beta fexp x |} <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex
H0:(ex <= fexp ex)%Z
0%Z = Ztrunc (scaled_mantissa beta fexp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex, Ex':(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex
H0:(ex <= fexp ex)%Z

(0 <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex
H0:(ex <= fexp ex)%Z
0%Z = Ztrunc (scaled_mantissa beta fexp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex, Ex':(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex
H0:(ex <= fexp ex)%Z

0%Z = Ztrunc (scaled_mantissa beta fexp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex, Ex':(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex
H0:(ex <= fexp ex)%Z

0%Z = Ztrunc (x * bpow (- cexp beta fexp x))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex, Ex':(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex
H0:(ex <= fexp ex)%Z

0%Z = Ztrunc (x * bpow (- fexp ex))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex, Ex':(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex
H0:(ex <= fexp ex)%Z
H1:(0 < x * bpow (- fexp ex))%R
H2:(x * bpow (- fexp ex) < 1)%R

0%Z = Ztrunc (x * bpow (- fexp ex))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex, Ex':(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex
H0:(ex <= fexp ex)%Z
H1:(0 < x * bpow (- fexp ex))%R
H2:(x * bpow (- fexp ex) < 1)%R

0%Z = Zfloor (x * bpow (- fexp ex))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex
H0:(ex <= fexp ex)%Z
H1:(0 < x * bpow (- fexp ex))%R
H2:(x * bpow (- fexp ex) < 1)%R
(0 <= x * bpow (- fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex, Ex':(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex
H0:(ex <= fexp ex)%Z
H1:(0 < x * bpow (- fexp ex))%R
H2:(x * bpow (- fexp ex) < 1)%R

Zfloor (x * bpow (- fexp ex)) = 0%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex
H0:(ex <= fexp ex)%Z
H1:(0 < x * bpow (- fexp ex))%R
H2:(x * bpow (- fexp ex) < 1)%R
(0 <= x * bpow (- fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex, Ex':(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex
H0:(ex <= fexp ex)%Z
H1:(0 < x * bpow (- fexp ex))%R
H2:(x * bpow (- fexp ex) < 1)%R

(0 <= x * bpow (- fexp ex) < IZR (0 + 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex
H0:(ex <= fexp ex)%Z
H1:(0 < x * bpow (- fexp ex))%R
H2:(x * bpow (- fexp ex) < 1)%R
(0 <= x * bpow (- fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex, Ex':(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex
H0:(ex <= fexp ex)%Z
H1:(0 < x * bpow (- fexp ex))%R
H2:(x * bpow (- fexp ex) < 1)%R

(0 <= x * bpow (- fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex
H0:(ex <= fexp ex)%Z
H1:(0 < x * bpow (- fexp ex))%R
H2:(x * bpow (- fexp ex) < 1)%R
(x * bpow (- fexp ex) < IZR (0 + 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex
H0:(ex <= fexp ex)%Z
H1:(0 < x * bpow (- fexp ex))%R
H2:(x * bpow (- fexp ex) < 1)%R
(0 <= x * bpow (- fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex, Ex':(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex
H0:(ex <= fexp ex)%Z
H1:(0 < x * bpow (- fexp ex))%R
H2:(x * bpow (- fexp ex) < 1)%R

(x * bpow (- fexp ex) < IZR (0 + 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex
H0:(ex <= fexp ex)%Z
H1:(0 < x * bpow (- fexp ex))%R
H2:(x * bpow (- fexp ex) < 1)%R
(0 <= x * bpow (- fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex, Ex':(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Ex':(bpow (ex - 1) <= x < bpow ex)%R
H:(x + ulp x)%R = bpow ex
H0:(ex <= fexp ex)%Z
H1:(0 < x * bpow (- fexp ex))%R
H2:(x * bpow (- fexp ex) < 1)%R

(0 <= x * bpow (- fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex, Ex':(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex, Ex':(bpow (ex - 1) <= Rabs x < bpow ex)%R

(0 <= x)%R
now apply Rlt_le. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, F x -> (0 < x)%R -> F (pred_pos x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, F x -> (0 < x)%R -> F (pred_pos x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Zx:(0 < x)%R

F (pred_pos x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Zx:(0 < x)%R
H:x = bpow (mag beta x - 1)

F (x - bpow (fexp (mag beta x - 1)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Zx:(0 < x)%R
H:x <> bpow (mag beta x - 1)
F (x - ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Zx:(0 < x)%R
H:x <> bpow (mag beta x - 1)

F (x - ulp x)
now apply generic_format_pred_aux1. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, F x -> F (succ x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, F x -> F (succ x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x

F (succ x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Zx:(0 <= x)%R

F (x + ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Zx:(x < 0)%R
F (- pred_pos (- x))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Zx:(0 < x)%R

F (x + ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Zx:0%R = x
F (x + ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Zx:(x < 0)%R
F (- pred_pos (- x))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Zx:0%R = x

F (x + ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Zx:(x < 0)%R
F (- pred_pos (- x))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Zx:0%R = x

F (ulp 0)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Zx:(x < 0)%R
F (- pred_pos (- x))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Zx:(x < 0)%R

F (- pred_pos (- x))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Zx:(x < 0)%R

F (pred_pos (- x))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Zx:(x < 0)%R

F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Zx:(x < 0)%R
(0 < - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Zx:(x < 0)%R

(0 < - x)%R
now apply Ropp_0_gt_lt_contravar. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, F x -> F (pred x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, F x -> F (pred x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x

F (pred x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x

F (- succ (- x))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x

F (succ (- x))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x

F (- x)
now apply generic_format_opp. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, x <> 0%R -> (pred_pos x < x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, x <> 0%R -> (pred_pos x < x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:x <> 0%R

(pred_pos x < x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:x <> 0%R

((if Req_bool x (bpow (mag beta x - 1)) then x - bpow (fexp (mag beta x - 1)) else x - ulp x) < x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:x <> 0%R
H:x = bpow (mag beta x - 1)

(x - bpow (fexp (mag beta x - 1)) < x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:x <> 0%R
H:x <> bpow (mag beta x - 1)
(x - ulp x < x)%R
(* *)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:x <> 0%R
H:x = bpow (mag beta x - 1)

(x - bpow (fexp (mag beta x - 1)) < x + 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:x <> 0%R
H:x <> bpow (mag beta x - 1)
(x - ulp x < x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:x <> 0%R
H:x = bpow (mag beta x - 1)

(- bpow (fexp (mag beta x - 1)) < 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:x <> 0%R
H:x <> bpow (mag beta x - 1)
(x - ulp x < x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:x <> 0%R
H:x = bpow (mag beta x - 1)

(- bpow (fexp (mag beta x - 1)) < - 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:x <> 0%R
H:x <> bpow (mag beta x - 1)
(x - ulp x < x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:x <> 0%R
H:x = bpow (mag beta x - 1)

(0 < bpow (fexp (mag beta x - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:x <> 0%R
H:x <> bpow (mag beta x - 1)
(x - ulp x < x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:x <> 0%R
H:x <> bpow (mag beta x - 1)

(x - ulp x < x)%R
(* *)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:x <> 0%R
H:x <> bpow (mag beta x - 1)

(x - ulp x < x + 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:x <> 0%R
H:x <> bpow (mag beta x - 1)

(- ulp x < 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:x <> 0%R
H:x <> bpow (mag beta x - 1)

(- ulp x < - 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:x <> 0%R
H:x <> bpow (mag beta x - 1)

(0 < ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:x <> 0%R
H:x <> bpow (mag beta x - 1)

(0 < bpow (cexp beta fexp x))%R
apply bpow_gt_0. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, x <> 0%R -> (x < succ x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, x <> 0%R -> (x < succ x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:x <> 0%R

(x < (if Rle_bool 0 x then x + ulp x else - pred_pos (- x)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:x <> 0%R
Hx:(0 <= x)%R

(x < x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:x <> 0%R
Hx:(x < 0)%R
(x < - pred_pos (- x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:x <> 0%R
Hx:(0 <= x)%R

(x + 0 < x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:x <> 0%R
Hx:(x < 0)%R
(x < - pred_pos (- x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:x <> 0%R
Hx:(0 <= x)%R

(0 < ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:x <> 0%R
Hx:(x < 0)%R
(x < - pred_pos (- x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:x <> 0%R
Hx:(0 <= x)%R

(0 < bpow (cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:x <> 0%R
Hx:(x < 0)%R
(x < - pred_pos (- x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:x <> 0%R
Hx:(x < 0)%R

(x < - pred_pos (- x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:x <> 0%R
Hx:(x < 0)%R

(- - x < - pred_pos (- x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:x <> 0%R
Hx:(x < 0)%R

(pred_pos (- x) < - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:x <> 0%R
Hx:(x < 0)%R

(- x)%R <> 0%R
auto with real. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, x <> 0%R -> (pred x < x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, x <> 0%R -> (pred x < x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:x <> 0%R

(- succ (- x) < x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:x <> 0%R

(- succ (- x) < - - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:x <> 0%R

(- x < succ (- x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:x <> 0%R

(- x)%R <> 0%R
auto with real. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, (x <= succ x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, (x <= succ x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R

x = 0%R -> (x <= succ x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
x <> 0%R -> (x <= succ x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
V:x = 0%R

(0 <= succ 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
x <> 0%R -> (x <= succ x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
V:x = 0%R

(0 <= 0 + ulp 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
x <> 0%R -> (x <= succ x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R

x <> 0%R -> (x <= succ x)%R
intros; left; now apply succ_gt_id. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, (pred x <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, (pred x <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R

(- succ (- x) <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R

(- succ (- x) <= - - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R

(- x <= succ (- x))%R
apply succ_ge_id. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, (0 < x)%R -> F x -> (0 <= pred_pos x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, (0 < x)%R -> F x -> (0 <= pred_pos x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x

(0 <= pred_pos x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x

(0 <= (if Req_bool x (bpow (mag beta x - 1)) then x - bpow (fexp (mag beta x - 1)) else x - ulp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
H:x = bpow (mag beta x - 1)

(0 <= x - bpow (fexp (mag beta x - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
H:x <> bpow (mag beta x - 1)
(0 <= x - ulp x)%R
(* *)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
H:x = bpow (mag beta x - 1)

(bpow (fexp (mag beta x - 1)) <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
H:x <> bpow (mag beta x - 1)
(0 <= x - ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
H:x = bpow (mag beta x - 1)

(bpow (fexp (mag beta (bpow (mag beta x - 1)) - 1)) <= bpow (mag beta x - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
H:x <> bpow (mag beta x - 1)
(0 <= x - ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
H:x = bpow (mag beta x - 1)

(fexp (mag beta (bpow (mag beta x - 1)) - 1) <= mag beta x - 1)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
H:x <> bpow (mag beta x - 1)
(0 <= x - ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
H:x = bpow (Build_mag_prop beta x ex Ex - 1)

(fexp (mag beta (bpow (ex - 1)) - 1) <= ex - 1)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
H:x <> bpow (mag beta x - 1)
(0 <= x - ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
H:x = bpow (Build_mag_prop beta x ex Ex - 1)

(fexp (ex - 1 + 1 - 1) <= ex - 1)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
H:x <> bpow (mag beta x - 1)
(0 <= x - ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
H:x = bpow (Build_mag_prop beta x ex Ex - 1)

(fexp (ex - 1) <= ex - 1)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
H:x <> bpow (mag beta x - 1)
(0 <= x - ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
H:x = bpow (Build_mag_prop beta x ex Ex - 1)

F (bpow (ex - 1))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
H:x <> bpow (mag beta x - 1)
(0 <= x - ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Ex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
H:x = bpow (ex - 1)

F (bpow (ex - 1))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
H:x <> bpow (mag beta x - 1)
(0 <= x - ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
H:x <> bpow (mag beta x - 1)

(0 <= x - ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
H:x <> bpow (mag beta x - 1)

(ulp x <= x)%R
now apply ulp_le_id. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, (0 < x)%R -> F x -> (0 <= pred x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, (0 < x)%R -> F x -> (0 <= pred x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x

(0 <= pred x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x

(0 <= pred_pos x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x

(0 <= x)%R
now left. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, (0 < x)%R -> F x -> x <> bpow (mag beta x - 1) -> (x - ulp x + ulp (x - ulp x))%R = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, (0 < x)%R -> F x -> x <> bpow (mag beta x - 1) -> (x - ulp x + ulp (x - ulp x))%R = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
Hx:x <> bpow (mag beta x - 1)

(x - ulp x + ulp (x - ulp x))%R = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
Hx:x <> bpow (mag beta x - 1)

(x - ulp x + ulp x)%R = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
Hx:x <> bpow (mag beta x - 1)
ulp x = ulp (x - ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
Hx:x <> bpow (mag beta x - 1)

ulp x = ulp (x - ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
Hx:x <> bpow (mag beta x - 1)
H:x <> 0%R

ulp x = ulp (x - ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
Hx:x <> bpow (mag beta x - 1)
H:x <> 0%R

x <> bpow (cexp beta fexp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
Hx:x <> bpow (mag beta x - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
ulp x = ulp (x - ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
Hx:x <> bpow (mag beta x - 1)
H:x <> 0%R
M:x = bpow (fexp (mag beta x))

False
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
Hx:x <> bpow (mag beta x - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
ulp x = ulp (x - ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
Hx:x <> bpow (mag beta x - 1)
H:x <> 0%R
M:x = bpow (fexp (mag beta x))
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
T:mag beta x = Build_mag_prop beta x ex Hex

False
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
Hx:x <> bpow (mag beta x - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
ulp x = ulp (x - ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
Hx:x <> bpow (mag beta x - 1)
H:x <> 0%R
M:x = bpow (fexp (mag beta x))
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
T:mag beta x = Build_mag_prop beta x ex Hex

mag beta x = ex
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
Hx:x <> bpow (mag beta x - 1)
H:x <> 0%R
M:x = bpow (fexp (mag beta x))
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
T:mag beta x = Build_mag_prop beta x ex Hex
Lex:mag beta x = ex
False
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
Hx:x <> bpow (mag beta x - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
ulp x = ulp (x - ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
Hx:x <> bpow (mag beta x - 1)
H:x <> 0%R
M:x = bpow (fexp (mag beta x))
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
T:mag beta x = Build_mag_prop beta x ex Hex
Lex:mag beta x = ex

False
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
Hx:x <> bpow (mag beta x - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
ulp x = ulp (x - ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
M:x = bpow (fexp ex)
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
T:mag beta x = Build_mag_prop beta x ex Hex
Lex:mag beta x = ex

False
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
Hx:x <> bpow (mag beta x - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
ulp x = ulp (x - ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
M:x = bpow (fexp ex)
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Lex:mag beta x = ex

False
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
Hx:x <> bpow (mag beta x - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
ulp x = ulp (x - ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
M:x = bpow (fexp ex)
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Lex:mag beta x = ex

False
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
Hx:x <> bpow (mag beta x - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
ulp x = ulp (x - ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
M:x = bpow (fexp ex)
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Lex:mag beta x = ex

(ex - 1 < fexp ex < ex)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
M:x = bpow (fexp ex)
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Lex:mag beta x = ex
H0:(ex - 1 < fexp ex < ex)%Z
False
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
Hx:x <> bpow (mag beta x - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
ulp x = ulp (x - ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
M:x = bpow (fexp ex)
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Lex:mag beta x = ex

(bpow (ex - 1) < x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
M:x = bpow (fexp ex)
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Lex:mag beta x = ex
(x < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
M:x = bpow (fexp ex)
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Lex:mag beta x = ex
H0:(ex - 1 < fexp ex < ex)%Z
False
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
Hx:x <> bpow (mag beta x - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
ulp x = ulp (x - ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
M:x = bpow (fexp ex)
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Lex:mag beta x = ex

(x < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
M:x = bpow (fexp ex)
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Lex:mag beta x = ex
H0:(ex - 1 < fexp ex < ex)%Z
False
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
Hx:x <> bpow (mag beta x - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
ulp x = ulp (x - ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
M:x = bpow (fexp ex)
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Lex:mag beta x = ex
H0:(ex - 1 < fexp ex < ex)%Z

False
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
Hx:x <> bpow (mag beta x - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
ulp x = ulp (x - ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
Hx:x <> bpow (mag beta x - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)

ulp x = ulp (x - ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
Hx:x <> bpow (mag beta x - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)

bpow (cexp beta fexp x) = bpow (cexp beta fexp (x - bpow (cexp beta fexp x)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
Hx:x <> bpow (mag beta x - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)

cexp beta fexp x = cexp beta fexp (x - bpow (cexp beta fexp x))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
Hx:x <> bpow (mag beta x - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)

mag beta x = mag beta (x - bpow (fexp (mag beta x)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
Hx:x <> bpow (mag beta x - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
T:mag beta x = Build_mag_prop beta x ex Hex

Build_mag_prop beta x ex Hex = mag beta (x - bpow (fexp (Build_mag_prop beta x ex Hex)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
Hx:x <> bpow (mag beta x - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
T:mag beta x = Build_mag_prop beta x ex Hex

mag beta x = ex
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
Hx:x <> bpow (mag beta x - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
T:mag beta x = Build_mag_prop beta x ex Hex
Lex:mag beta x = ex
Build_mag_prop beta x ex Hex = mag beta (x - bpow (fexp (Build_mag_prop beta x ex Hex)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
Hx:x <> bpow (mag beta x - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
T:mag beta x = Build_mag_prop beta x ex Hex
Lex:mag beta x = ex

Build_mag_prop beta x ex Hex = mag beta (x - bpow (fexp (Build_mag_prop beta x ex Hex)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
Lex:mag beta x = ex

ex = mag beta (x - bpow (fexp ex))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Lex:mag beta x = ex

ex = mag beta (x - bpow (fexp ex))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Lex:mag beta x = ex

(bpow (ex - 1) <= Rabs (x - bpow (fexp ex)) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Lex:mag beta x = ex

(bpow (ex - 1) <= x - bpow (fexp ex) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Lex:mag beta x = ex
(x - bpow (fexp ex) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Lex:mag beta x = ex

(bpow (ex - 1) <= x - bpow (fexp ex) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Lex:mag beta x = ex
(x >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Lex:mag beta x = ex
(x - bpow (fexp ex) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Lex:mag beta x = ex

(bpow (ex - 1) <= x - bpow (fexp ex) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Lex:mag beta x = ex
(x - bpow (fexp ex) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Lex:mag beta x = ex

(bpow (ex - 1) <= x - bpow (fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Lex:mag beta x = ex
(x - bpow (fexp ex) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Lex:mag beta x = ex
(x - bpow (fexp ex) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
H1:(bpow (ex - 1) < x)%R
H2:(x < bpow ex)%R
Lex:mag beta x = ex

(bpow (ex - 1) <= x - bpow (fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
H1:bpow (ex - 1) = x
H2:(x < bpow ex)%R
Lex:mag beta x = ex
(bpow (ex - 1) <= x - bpow (fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Lex:mag beta x = ex
(x - bpow (fexp ex) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Lex:mag beta x = ex
(x - bpow (fexp ex) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
H1:(bpow (ex - 1) < x)%R
H2:(x < bpow ex)%R
Lex:mag beta x = ex

(bpow (ex - 1) <= x - ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
H1:(bpow (ex - 1) < x)%R
H2:(x < bpow ex)%R
Lex:mag beta x = ex
(x - ulp x <= x - bpow (fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
H1:bpow (ex - 1) = x
H2:(x < bpow ex)%R
Lex:mag beta x = ex
(bpow (ex - 1) <= x - bpow (fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Lex:mag beta x = ex
(x - bpow (fexp ex) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Lex:mag beta x = ex
(x - bpow (fexp ex) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
H1:(bpow (ex - 1) < x)%R
H2:(x < bpow ex)%R
Lex:mag beta x = ex

x <> ulp x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
H1:(bpow (ex - 1) < x)%R
H2:(x < bpow ex)%R
Lex:mag beta x = ex
(x - ulp x <= x - bpow (fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
H1:bpow (ex - 1) = x
H2:(x < bpow ex)%R
Lex:mag beta x = ex
(bpow (ex - 1) <= x - bpow (fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Lex:mag beta x = ex
(x - bpow (fexp ex) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Lex:mag beta x = ex
(x - bpow (fexp ex) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
H1:(bpow (ex - 1) < x)%R
H2:(x < bpow ex)%R
Lex:mag beta x = ex

(x - ulp x <= x - bpow (fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
H1:bpow (ex - 1) = x
H2:(x < bpow ex)%R
Lex:mag beta x = ex
(bpow (ex - 1) <= x - bpow (fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Lex:mag beta x = ex
(x - bpow (fexp ex) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Lex:mag beta x = ex
(x - bpow (fexp ex) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
H1:(bpow (ex - 1) < x)%R
H2:(x < bpow ex)%R
Lex:mag beta x = ex

(x - bpow (cexp beta fexp x) <= x - bpow (fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
H1:bpow (ex - 1) = x
H2:(x < bpow ex)%R
Lex:mag beta x = ex
(bpow (ex - 1) <= x - bpow (fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Lex:mag beta x = ex
(x - bpow (fexp ex) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Lex:mag beta x = ex
(x - bpow (fexp ex) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
H1:bpow (ex - 1) = x
H2:(x < bpow ex)%R
Lex:mag beta x = ex

(bpow (ex - 1) <= x - bpow (fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Lex:mag beta x = ex
(x - bpow (fexp ex) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Lex:mag beta x = ex
(x - bpow (fexp ex) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Lex:mag beta x = ex

(x - bpow (fexp ex) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Lex:mag beta x = ex
(x - bpow (fexp ex) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Lex:mag beta x = ex

(x - bpow (fexp ex) <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Lex:mag beta x = ex
(x - bpow (fexp ex) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Lex:mag beta x = ex

(x - bpow (fexp ex) <= x + 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Lex:mag beta x = ex
(x - bpow (fexp ex) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Lex:mag beta x = ex

(- bpow (fexp ex) <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Lex:mag beta x = ex
(x - bpow (fexp ex) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Lex:mag beta x = ex

(- bpow (fexp ex) <= - 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Lex:mag beta x = ex
(x - bpow (fexp ex) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Lex:mag beta x = ex

(0 <= bpow (fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Lex:mag beta x = ex
(x - bpow (fexp ex) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Lex:mag beta x = ex

(x - bpow (fexp ex) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Lex:mag beta x = ex

(0 <= x - bpow (fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Lex:mag beta x = ex

(bpow (fexp ex) <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Lex:mag beta x = ex

(bpow (fexp ex) <= F2R {| Fnum := Ztrunc (scaled_mantissa beta fexp x); Fexp := cexp beta fexp x |})%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Lex:mag beta x = ex

(bpow (fexp ex) <= IZR (Ztrunc (scaled_mantissa beta fexp x)) * bpow (fexp (mag beta x)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Lex:mag beta x = ex

(bpow (fexp ex) <= IZR (Ztrunc (scaled_mantissa beta fexp x)) * bpow (fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Lex:mag beta x = ex

(1 * bpow (fexp ex) <= IZR (Ztrunc (scaled_mantissa beta fexp x)) * bpow (fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Lex:mag beta x = ex

(0 <= bpow (fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Lex:mag beta x = ex
(1 <= IZR (Ztrunc (scaled_mantissa beta fexp x)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Lex:mag beta x = ex

(1 <= IZR (Ztrunc (scaled_mantissa beta fexp x)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Lex:mag beta x = ex

(0 < Ztrunc (scaled_mantissa beta fexp x))%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
ex:Z
Hx:x <> bpow (ex - 1)
H:x <> 0%R
H':x <> bpow (cexp beta fexp x)
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Lex:mag beta x = ex

(0 < F2R {| Fnum := Ztrunc (scaled_mantissa beta fexp x); Fexp := cexp beta fexp x |})%R
now rewrite <- Fx. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, (0 < x)%R -> F x -> let e := mag_val beta x (mag beta x) in x = bpow (e - 1) -> (x - bpow (fexp (e - 1)))%R <> 0%R -> (x - bpow (fexp (e - 1)) + ulp (x - bpow (fexp (e - 1))))%R = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, (0 < x)%R -> F x -> let e := mag_val beta x (mag beta x) in x = bpow (e - 1) -> (x - bpow (fexp (e - 1)))%R <> 0%R -> (x - bpow (fexp (e - 1)) + ulp (x - bpow (fexp (e - 1))))%R = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R

(x - bpow (fexp (e - 1)) + ulp (x - bpow (fexp (e - 1))))%R = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R

(x - bpow (fexp (e - 1)) + bpow (fexp (e - 1)))%R = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
bpow (fexp (e - 1)) = ulp (x - bpow (fexp (e - 1)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R

bpow (fexp (e - 1)) = ulp (x - bpow (fexp (e - 1)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R

(fexp (e - 1) <= e - 1)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) <= e - 1)%Z
bpow (fexp (e - 1)) = ulp (x - bpow (fexp (e - 1)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R

F (bpow (e - 1))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) <= e - 1)%Z
bpow (fexp (e - 1)) = ulp (x - bpow (fexp (e - 1)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) <= e - 1)%Z

bpow (fexp (e - 1)) = ulp (x - bpow (fexp (e - 1)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z

bpow (fexp (e - 1)) = ulp (x - bpow (fexp (e - 1)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:fexp (e - 1) = (e - 1)%Z
bpow (fexp (e - 1)) = ulp (x - bpow (fexp (e - 1)))
(* *)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z

bpow (fexp (e - 1)) = bpow (cexp beta fexp (x - bpow (fexp (e - 1))))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:fexp (e - 1) = (e - 1)%Z
bpow (fexp (e - 1)) = ulp (x - bpow (fexp (e - 1)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z

fexp (e - 1) = cexp beta fexp (x - bpow (fexp (e - 1)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:fexp (e - 1) = (e - 1)%Z
bpow (fexp (e - 1)) = ulp (x - bpow (fexp (e - 1)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z

(e - 1)%Z = mag beta (x - bpow (fexp (e - 1)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:fexp (e - 1) = (e - 1)%Z
bpow (fexp (e - 1)) = ulp (x - bpow (fexp (e - 1)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z

mag beta (x - bpow (fexp (e - 1))) = (e - 1)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:fexp (e - 1) = (e - 1)%Z
bpow (fexp (e - 1)) = ulp (x - bpow (fexp (e - 1)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z

(bpow (e - 1 - 1) <= Rabs (x - bpow (fexp (e - 1))) < bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:fexp (e - 1) = (e - 1)%Z
bpow (fexp (e - 1)) = ulp (x - bpow (fexp (e - 1)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z

(bpow (e - 1 - 1) <= x - bpow (fexp (e - 1)) < bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
(x - bpow (fexp (e - 1)) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:fexp (e - 1) = (e - 1)%Z
bpow (fexp (e - 1)) = ulp (x - bpow (fexp (e - 1)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z

(bpow (e - 1 - 1) <= x - bpow (fexp (e - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
(x - bpow (fexp (e - 1)) < bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
(x - bpow (fexp (e - 1)) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:fexp (e - 1) = (e - 1)%Z
bpow (fexp (e - 1)) = ulp (x - bpow (fexp (e - 1)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z

(bpow (fexp (e - 1)) + bpow (e - 1 - 1) <= bpow (fexp (e - 1)) + (x - bpow (fexp (e - 1))))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
(x - bpow (fexp (e - 1)) < bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
(x - bpow (fexp (e - 1)) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:fexp (e - 1) = (e - 1)%Z
bpow (fexp (e - 1)) = ulp (x - bpow (fexp (e - 1)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z

(bpow (fexp (e - 1)) + bpow (e - 1 - 1) <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
(x - bpow (fexp (e - 1)) < bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
(x - bpow (fexp (e - 1)) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:fexp (e - 1) = (e - 1)%Z
bpow (fexp (e - 1)) = ulp (x - bpow (fexp (e - 1)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z

(bpow (fexp (e - 1)) + bpow (e - 1 - 1) <= bpow (e - 2) + bpow (e - 2))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
(bpow (e - 2) + bpow (e - 2) <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
(x - bpow (fexp (e - 1)) < bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
(x - bpow (fexp (e - 1)) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:fexp (e - 1) = (e - 1)%Z
bpow (fexp (e - 1)) = ulp (x - bpow (fexp (e - 1)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z

(bpow (e - 2) + bpow (e - 2) <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
(x - bpow (fexp (e - 1)) < bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
(x - bpow (fexp (e - 1)) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:fexp (e - 1) = (e - 1)%Z
bpow (fexp (e - 1)) = ulp (x - bpow (fexp (e - 1)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z

(2 * bpow (e - 2) <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
(x - bpow (fexp (e - 1)) < bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
(x - bpow (fexp (e - 1)) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:fexp (e - 1) = (e - 1)%Z
bpow (fexp (e - 1)) = ulp (x - bpow (fexp (e - 1)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z

(2 * bpow (e - 2) <= bpow 1 * bpow (e - 2))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
(bpow 1 * bpow (e - 2) <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
(x - bpow (fexp (e - 1)) < bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
(x - bpow (fexp (e - 1)) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:fexp (e - 1) = (e - 1)%Z
bpow (fexp (e - 1)) = ulp (x - bpow (fexp (e - 1)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z

(0 <= bpow (e - 2))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
(2 <= bpow 1)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
(bpow 1 * bpow (e - 2) <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
(x - bpow (fexp (e - 1)) < bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
(x - bpow (fexp (e - 1)) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:fexp (e - 1) = (e - 1)%Z
bpow (fexp (e - 1)) = ulp (x - bpow (fexp (e - 1)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z

(2 <= bpow 1)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
(bpow 1 * bpow (e - 2) <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
(x - bpow (fexp (e - 1)) < bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
(x - bpow (fexp (e - 1)) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:fexp (e - 1) = (e - 1)%Z
bpow (fexp (e - 1)) = ulp (x - bpow (fexp (e - 1)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z

(2 <= IZR beta)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
IZR beta = bpow 1
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
(bpow 1 * bpow (e - 2) <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
(x - bpow (fexp (e - 1)) < bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
(x - bpow (fexp (e - 1)) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:fexp (e - 1) = (e - 1)%Z
bpow (fexp (e - 1)) = ulp (x - bpow (fexp (e - 1)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z

(2 <= beta)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
IZR beta = bpow 1
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
(bpow 1 * bpow (e - 2) <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
(x - bpow (fexp (e - 1)) < bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
(x - bpow (fexp (e - 1)) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:fexp (e - 1) = (e - 1)%Z
bpow (fexp (e - 1)) = ulp (x - bpow (fexp (e - 1)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z

(2 <=? beta)%Z = true
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
IZR beta = bpow 1
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
(bpow 1 * bpow (e - 2) <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
(x - bpow (fexp (e - 1)) < bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
(x - bpow (fexp (e - 1)) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:fexp (e - 1) = (e - 1)%Z
bpow (fexp (e - 1)) = ulp (x - bpow (fexp (e - 1)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z

IZR beta = bpow 1
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
(bpow 1 * bpow (e - 2) <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
(x - bpow (fexp (e - 1)) < bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
(x - bpow (fexp (e - 1)) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:fexp (e - 1) = (e - 1)%Z
bpow (fexp (e - 1)) = ulp (x - bpow (fexp (e - 1)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z

IZR beta = IZR (Z.pow_pos beta 1)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
(bpow 1 * bpow (e - 2) <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
(x - bpow (fexp (e - 1)) < bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
(x - bpow (fexp (e - 1)) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:fexp (e - 1) = (e - 1)%Z
bpow (fexp (e - 1)) = ulp (x - bpow (fexp (e - 1)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z

IZR beta = IZR (beta * 1)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
(bpow 1 * bpow (e - 2) <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
(x - bpow (fexp (e - 1)) < bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
(x - bpow (fexp (e - 1)) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:fexp (e - 1) = (e - 1)%Z
bpow (fexp (e - 1)) = ulp (x - bpow (fexp (e - 1)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z

(bpow 1 * bpow (e - 2) <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
(x - bpow (fexp (e - 1)) < bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
(x - bpow (fexp (e - 1)) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:fexp (e - 1) = (e - 1)%Z
bpow (fexp (e - 1)) = ulp (x - bpow (fexp (e - 1)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z

(bpow (1 + (e - 2)) <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
(x - bpow (fexp (e - 1)) < bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
(x - bpow (fexp (e - 1)) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:fexp (e - 1) = (e - 1)%Z
bpow (fexp (e - 1)) = ulp (x - bpow (fexp (e - 1)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z

(bpow (e - 1) <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
(x - bpow (fexp (e - 1)) < bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
(x - bpow (fexp (e - 1)) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:fexp (e - 1) = (e - 1)%Z
bpow (fexp (e - 1)) = ulp (x - bpow (fexp (e - 1)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z

(x - bpow (fexp (e - 1)) < bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
(x - bpow (fexp (e - 1)) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:fexp (e - 1) = (e - 1)%Z
bpow (fexp (e - 1)) = ulp (x - bpow (fexp (e - 1)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z

(bpow (e - 1) - bpow (fexp (e - 1)) < bpow (e - 1) + 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
(x - bpow (fexp (e - 1)) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:fexp (e - 1) = (e - 1)%Z
bpow (fexp (e - 1)) = ulp (x - bpow (fexp (e - 1)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z

(- bpow (fexp (e - 1)) < 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
(x - bpow (fexp (e - 1)) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:fexp (e - 1) = (e - 1)%Z
bpow (fexp (e - 1)) = ulp (x - bpow (fexp (e - 1)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z

(- bpow (fexp (e - 1)) < - 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
(x - bpow (fexp (e - 1)) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:fexp (e - 1) = (e - 1)%Z
bpow (fexp (e - 1)) = ulp (x - bpow (fexp (e - 1)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z

(0 < bpow (fexp (e - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z
(x - bpow (fexp (e - 1)) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:fexp (e - 1) = (e - 1)%Z
bpow (fexp (e - 1)) = ulp (x - bpow (fexp (e - 1)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z

(x - bpow (fexp (e - 1)) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:fexp (e - 1) = (e - 1)%Z
bpow (fexp (e - 1)) = ulp (x - bpow (fexp (e - 1)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z

(bpow (fexp (e - 1)) <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:fexp (e - 1) = (e - 1)%Z
bpow (fexp (e - 1)) = ulp (x - bpow (fexp (e - 1)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z

(bpow (fexp (e - 1)) <= bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:fexp (e - 1) = (e - 1)%Z
bpow (fexp (e - 1)) = ulp (x - bpow (fexp (e - 1)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:(fexp (e - 1) < e - 1)%Z

(fexp (e - 1) <= e - 1)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:fexp (e - 1) = (e - 1)%Z
bpow (fexp (e - 1)) = ulp (x - bpow (fexp (e - 1)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
Zp:(x - bpow (fexp (e - 1)))%R <> 0%R
He:fexp (e - 1) = (e - 1)%Z

bpow (fexp (e - 1)) = ulp (x - bpow (fexp (e - 1)))
(* *)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
Hxe:x = bpow (e - 1)
He:fexp (e - 1) = (e - 1)%Z

(x - bpow (fexp (e - 1)))%R = 0%R
rewrite Hxe, He; ring. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, (0 < x)%R -> F x -> let e := mag_val beta x (mag beta x) in x = bpow (e - 1) -> (x - bpow (fexp (e - 1)))%R = 0%R -> ulp 0 = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, (0 < x)%R -> F x -> let e := mag_val beta x (mag beta x) in x = bpow (e - 1) -> (x - bpow (fexp (e - 1)))%R = 0%R -> ulp 0 = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
H1:x = bpow (e - 1)
H2:(x - bpow (fexp (e - 1)))%R = 0%R

ulp 0 = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
H1:x = bpow (e - 1)
H2:(x - bpow (fexp (e - 1)))%R = 0%R

x = bpow (fexp (e - 1))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
H1:x = bpow (e - 1)
H2:(x - bpow (fexp (e - 1)))%R = 0%R
H3:x = bpow (fexp (e - 1))
ulp 0 = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
H1:x = bpow (e - 1)
H2:(x - bpow (fexp (e - 1)))%R = 0%R
H3:x = bpow (fexp (e - 1))

ulp 0 = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
H1:x = bpow (e - 1)
H2:(x - bpow (fexp (e - 1)))%R = 0%R
H3:x = bpow (fexp (e - 1))

fexp (e - 1) = (e - 1)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
H1:x = bpow (e - 1)
H2:(x - bpow (fexp (e - 1)))%R = 0%R
H3:x = bpow (fexp (e - 1))
H4:fexp (e - 1) = (e - 1)%Z
ulp 0 = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
H1:x = bpow (e - 1)
H2:(x - bpow (fexp (e - 1)))%R = 0%R
H3:x = bpow (fexp (e - 1))

bpow (fexp (e - 1)) = bpow (e - 1)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
H1:x = bpow (e - 1)
H2:(x - bpow (fexp (e - 1)))%R = 0%R
H3:x = bpow (fexp (e - 1))
H4:fexp (e - 1) = (e - 1)%Z
ulp 0 = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
H1:x = bpow (e - 1)
H2:(x - bpow (fexp (e - 1)))%R = 0%R
H3:x = bpow (fexp (e - 1))
H4:fexp (e - 1) = (e - 1)%Z

ulp 0 = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
H1:x = bpow (e - 1)
H2:(x - bpow (fexp (e - 1)))%R = 0%R
H3:x = bpow (fexp (e - 1))
H4:fexp (e - 1) = (e - 1)%Z

match negligible_exp with | Some n => bpow (fexp n) | None => 0%R end = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
H1:x = bpow (e - 1)
H2:(x - bpow (fexp (e - 1)))%R = 0%R
H3:x = bpow (fexp (e - 1))
H4:fexp (e - 1) = (e - 1)%Z

(forall n : Z, (fexp n < n)%Z) -> 0%R = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
H1:x = bpow (e - 1)
H2:(x - bpow (fexp (e - 1)))%R = 0%R
H3:x = bpow (fexp (e - 1))
H4:fexp (e - 1) = (e - 1)%Z
forall n : Z, (n <= fexp n)%Z -> bpow (fexp n) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
H1:x = bpow (e - 1)
H2:(x - bpow (fexp (e - 1)))%R = 0%R
H3:x = bpow (fexp (e - 1))
H4:fexp (e - 1) = (e - 1)%Z
K:forall n : Z, (fexp n < n)%Z

0%R = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
H1:x = bpow (e - 1)
H2:(x - bpow (fexp (e - 1)))%R = 0%R
H3:x = bpow (fexp (e - 1))
H4:fexp (e - 1) = (e - 1)%Z
forall n : Z, (n <= fexp n)%Z -> bpow (fexp n) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
H1:x = bpow (e - 1)
H2:(x - bpow (fexp (e - 1)))%R = 0%R
H3:x = bpow (fexp (e - 1))
H4:fexp (e - 1) = (e - 1)%Z
K:(fexp (e - 1) < e - 1)%Z

0%R = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
H1:x = bpow (e - 1)
H2:(x - bpow (fexp (e - 1)))%R = 0%R
H3:x = bpow (fexp (e - 1))
H4:fexp (e - 1) = (e - 1)%Z
forall n : Z, (n <= fexp n)%Z -> bpow (fexp n) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
H1:x = bpow (e - 1)
H2:(x - bpow (fexp (e - 1)))%R = 0%R
H3:x = bpow (fexp (e - 1))
H4:fexp (e - 1) = (e - 1)%Z

forall n : Z, (n <= fexp n)%Z -> bpow (fexp n) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
H1:x = bpow (e - 1)
H2:(x - bpow (fexp (e - 1)))%R = 0%R
H3:x = bpow (fexp (e - 1))
H4:fexp (e - 1) = (e - 1)%Z
n:Z
Hn:(n <= fexp n)%Z

bpow (fexp n) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
H1:x = bpow (e - 1)
H2:(x - bpow (fexp (e - 1)))%R = 0%R
H3:x = bpow (fexp (e - 1))
H4:fexp (e - 1) = (e - 1)%Z
n:Z
Hn:(n <= fexp n)%Z

fexp n = fexp (e - 1)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
H1:x = bpow (e - 1)
H2:(x - bpow (fexp (e - 1)))%R = 0%R
H3:x = bpow (fexp (e - 1))
H4:fexp (e - 1) = (e - 1)%Z
n:Z
Hn:(n <= fexp n)%Z
H6:(n <= e - 1)%Z

fexp n = fexp (e - 1)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
H1:x = bpow (e - 1)
H2:(x - bpow (fexp (e - 1)))%R = 0%R
H3:x = bpow (fexp (e - 1))
H4:fexp (e - 1) = (e - 1)%Z
n:Z
Hn:(n <= fexp n)%Z
H6:(e - 1 < n)%Z
fexp n = fexp (e - 1)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Fx:F x
e:=mag_val beta x (mag beta x):Z
H1:x = bpow (e - 1)
H2:(x - bpow (fexp (e - 1)))%R = 0%R
H3:x = bpow (fexp (e - 1))
H4:fexp (e - 1) = (e - 1)%Z
n:Z
Hn:(n <= fexp n)%Z
H6:(e - 1 < n)%Z

fexp n = fexp (e - 1)
apply sym_eq, valid_exp; omega. Qed.
The following one is false for x = 0 in FTZ
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, (0 < x)%R -> F x -> (pred_pos x + ulp (pred_pos x))%R = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, (0 < x)%R -> F x -> (pred_pos x + ulp (pred_pos x))%R = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x

(pred_pos x + ulp (pred_pos x))%R = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x

((if Req_bool x (bpow (mag beta x - 1)) then x - bpow (fexp (mag beta x - 1)) else x - ulp x) + ulp (if Req_bool x (bpow (mag beta x - 1)) then x - bpow (fexp (mag beta x - 1)) else x - ulp x))%R = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
H:x = bpow (mag beta x - 1)

(x - bpow (fexp (mag beta x - 1)) + ulp (x - bpow (fexp (mag beta x - 1))))%R = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
H:x <> bpow (mag beta x - 1)
(x - ulp x + ulp (x - ulp x))%R = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
H:x = bpow (mag beta x - 1)
H1:(x - bpow (fexp (mag beta x - 1)))%R = 0%R

(x - bpow (fexp (mag beta x - 1)) + ulp (x - bpow (fexp (mag beta x - 1))))%R = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
H:x = bpow (mag beta x - 1)
H1:(x - bpow (fexp (mag beta x - 1)))%R <> 0%R
(x - bpow (fexp (mag beta x - 1)) + ulp (x - bpow (fexp (mag beta x - 1))))%R = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
H:x <> bpow (mag beta x - 1)
(x - ulp x + ulp (x - ulp x))%R = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
H:x = bpow (mag beta x - 1)
H1:(x - bpow (fexp (mag beta x - 1)))%R = 0%R

ulp 0 = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
H:x = bpow (mag beta x - 1)
H1:(x - bpow (fexp (mag beta x - 1)))%R <> 0%R
(x - bpow (fexp (mag beta x - 1)) + ulp (x - bpow (fexp (mag beta x - 1))))%R = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
H:x <> bpow (mag beta x - 1)
(x - ulp x + ulp (x - ulp x))%R = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
H:x = bpow (mag beta x - 1)
H1:(x - bpow (fexp (mag beta x - 1)))%R <> 0%R

(x - bpow (fexp (mag beta x - 1)) + ulp (x - bpow (fexp (mag beta x - 1))))%R = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
H:x <> bpow (mag beta x - 1)
(x - ulp x + ulp (x - ulp x))%R = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
H:x <> bpow (mag beta x - 1)

(x - ulp x + ulp (x - ulp x))%R = x
now apply pred_pos_plus_ulp_aux1. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, (0 < x)%R -> F x -> (pred x + ulp (pred x))%R = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, (0 < x)%R -> F x -> (pred x + ulp (pred x))%R = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Fx:F x

(pred x + ulp (pred x))%R = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Fx:F x

(pred_pos x + ulp (pred_pos x))%R = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Fx:F x
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Fx:F x

(0 <= x)%R
now apply Rlt_le. Qed.
Rounding x + small epsilon
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, (0 < x)%R -> F x -> forall eps : R, (0 <= eps < ulp x)%R -> mag beta (x + eps) = mag beta x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, (0 < x)%R -> F x -> forall eps : R, (0 <= eps < ulp x)%R -> mag beta (x + eps) = mag beta x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R

mag beta (x + eps) = mag beta x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R

mag beta (x + eps) = Build_mag_prop beta x ex He
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R

mag beta (x + eps) = ex
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= Rabs x < bpow ex)%R

mag beta (x + eps) = ex
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= Rabs x < bpow ex)%R

(bpow (ex - 1) <= Rabs (x + eps) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= Rabs x < bpow ex)%R

(bpow (ex - 1) <= x + eps < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x + eps)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R

(bpow (ex - 1) <= x + eps < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x + eps)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R

(bpow (ex - 1) <= x + eps)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R
(x + eps < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x + eps)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R

(x <= x + eps)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R
(x + eps < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x + eps)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R

(x + 0 <= x + eps)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R
(x + eps < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x + eps)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R

(x + eps < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x + eps)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R

(x + eps < x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R
(x + ulp x <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x + eps)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R

(x + ulp x <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x + eps)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R

(F2R {| Fnum := Ztrunc (scaled_mantissa beta fexp x); Fexp := cexp beta fexp x |} + ulp x <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x + eps)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R

(F2R {| Fnum := Ztrunc (scaled_mantissa beta fexp x); Fexp := cexp beta fexp x |} + bpow (cexp beta fexp x) <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x + eps)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R

(IZR (Fnum {| Fnum := Ztrunc (scaled_mantissa beta fexp x); Fexp := cexp beta fexp x |}) * bpow (Fexp {| Fnum := Ztrunc (scaled_mantissa beta fexp x); Fexp := cexp beta fexp x |}) + bpow (cexp beta fexp x) <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x + eps)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R

(IZR (Ztrunc (scaled_mantissa beta fexp x)) * bpow (cexp beta fexp x) + bpow (cexp beta fexp x) <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x + eps)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R

(IZR (Ztrunc (scaled_mantissa beta fexp x)) * bpow (cexp beta fexp x) + 1 * bpow (cexp beta fexp x) <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x + eps)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R

((IZR (Ztrunc (scaled_mantissa beta fexp x)) + 1) * bpow (cexp beta fexp x) <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x + eps)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R

(IZR (Ztrunc (scaled_mantissa beta fexp x) + 1) * bpow (cexp beta fexp x) <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x + eps)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R

(0 < Ztrunc (scaled_mantissa beta fexp x))%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R
(F2R {| Fnum := Ztrunc (scaled_mantissa beta fexp x); Fexp := fexp (mag beta x) |} < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x + eps)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R

(0 < F2R {| Fnum := Ztrunc (scaled_mantissa beta fexp x); Fexp := cexp beta fexp x |})%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R
(F2R {| Fnum := Ztrunc (scaled_mantissa beta fexp x); Fexp := fexp (mag beta x) |} < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x + eps)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R

(F2R {| Fnum := Ztrunc (scaled_mantissa beta fexp x); Fexp := fexp (mag beta x) |} < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x + eps)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R

x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x + eps)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= Rabs x < bpow ex)%R

(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x + eps)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= Rabs x < bpow ex)%R

(0 <= x + eps)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= Rabs x < bpow ex)%R

(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= eps)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
ex:Z
He:(bpow (ex - 1) <= Rabs x < bpow ex)%R

(0 <= eps)%R
apply Heps. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, (0 <= x)%R -> F x -> forall eps : R, (0 <= eps < ulp x)%R -> round beta fexp Zfloor (x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, (0 <= x)%R -> F x -> forall eps : R, (0 <= eps < ulp x)%R -> round beta fexp Zfloor (x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R

round beta fexp Zfloor (x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R

round beta fexp Zfloor (x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
round beta fexp Zfloor (x + eps) = x
(* . 0 < x *)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R

round beta fexp Zfloor (x + eps) = F2R {| Fnum := Ztrunc (scaled_mantissa beta fexp x); Fexp := cexp beta fexp x |}
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
round beta fexp Zfloor (x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R

F2R {| Fnum := Zfloor (scaled_mantissa beta fexp (x + eps)); Fexp := cexp beta fexp (x + eps) |} = F2R {| Fnum := Ztrunc (scaled_mantissa beta fexp x); Fexp := cexp beta fexp x |}
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
round beta fexp Zfloor (x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R

F2R {| Fnum := Zfloor ((x + eps) * bpow (- cexp beta fexp (x + eps))); Fexp := cexp beta fexp (x + eps) |} = F2R {| Fnum := Ztrunc (x * bpow (- cexp beta fexp x)); Fexp := cexp beta fexp x |}
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
round beta fexp Zfloor (x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R

F2R {| Fnum := Zfloor ((x + eps) * bpow (- cexp beta fexp (x + eps))); Fexp := cexp beta fexp (x + eps) |} = F2R {| Fnum := Ztrunc (x * bpow (- cexp beta fexp x)); Fexp := cexp beta fexp x |}
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
round beta fexp Zfloor (x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R

F2R {| Fnum := Zfloor ((x + eps) * bpow (- fexp (mag beta (x + eps)))); Fexp := fexp (mag beta (x + eps)) |} = F2R {| Fnum := Ztrunc (x * bpow (- cexp beta fexp x)); Fexp := cexp beta fexp x |}
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
round beta fexp Zfloor (x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R

F2R {| Fnum := Zfloor ((x + eps) * bpow (- fexp (mag beta x))); Fexp := fexp (mag beta x) |} = F2R {| Fnum := Ztrunc (x * bpow (- cexp beta fexp x)); Fexp := cexp beta fexp x |}
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
round beta fexp Zfloor (x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R

Zfloor ((x + eps) * bpow (- fexp (mag beta x))) = Ztrunc (x * bpow (- cexp beta fexp x))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
round beta fexp Zfloor (x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R

Zfloor ((x + eps) * bpow (- fexp (mag beta x))) = Zfloor (x * bpow (- cexp beta fexp x))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
(0 <= x * bpow (- cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
round beta fexp Zfloor (x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R

(IZR (Zfloor (x * bpow (- cexp beta fexp x))) <= (x + eps) * bpow (- fexp (mag beta x)) < IZR (Zfloor (x * bpow (- cexp beta fexp x)) + 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
(0 <= x * bpow (- cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
round beta fexp Zfloor (x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R

(IZR (Zfloor (x * bpow (- cexp beta fexp x))) <= (x + eps) * bpow (- fexp (mag beta x)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
((x + eps) * bpow (- fexp (mag beta x)) < IZR (Zfloor (x * bpow (- cexp beta fexp x)) + 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
(0 <= x * bpow (- cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
round beta fexp Zfloor (x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R

(x * bpow (- cexp beta fexp x) <= (x + eps) * bpow (- fexp (mag beta x)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
((x + eps) * bpow (- fexp (mag beta x)) < IZR (Zfloor (x * bpow (- cexp beta fexp x)) + 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
(0 <= x * bpow (- cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
round beta fexp Zfloor (x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R

(0 <= bpow (- fexp (mag beta x)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
(x <= x + eps)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
((x + eps) * bpow (- fexp (mag beta x)) < IZR (Zfloor (x * bpow (- cexp beta fexp x)) + 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
(0 <= x * bpow (- cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
round beta fexp Zfloor (x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R

(x <= x + eps)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
((x + eps) * bpow (- fexp (mag beta x)) < IZR (Zfloor (x * bpow (- cexp beta fexp x)) + 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
(0 <= x * bpow (- cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
round beta fexp Zfloor (x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R

(x + 0 <= x + eps)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
((x + eps) * bpow (- fexp (mag beta x)) < IZR (Zfloor (x * bpow (- cexp beta fexp x)) + 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
(0 <= x * bpow (- cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
round beta fexp Zfloor (x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R

((x + eps) * bpow (- fexp (mag beta x)) < IZR (Zfloor (x * bpow (- cexp beta fexp x)) + 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
(0 <= x * bpow (- cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
round beta fexp Zfloor (x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R

((x + eps) * bpow (- fexp (mag beta x)) < (x + ulp x) * bpow (- cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
((x + ulp x) * bpow (- cexp beta fexp x) <= IZR (Zfloor (x * bpow (- cexp beta fexp x)) + 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
(0 <= x * bpow (- cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
round beta fexp Zfloor (x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R

(0 < bpow (- cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
(x + eps < x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
((x + ulp x) * bpow (- cexp beta fexp x) <= IZR (Zfloor (x * bpow (- cexp beta fexp x)) + 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
(0 <= x * bpow (- cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
round beta fexp Zfloor (x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R

(x + eps < x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
((x + ulp x) * bpow (- cexp beta fexp x) <= IZR (Zfloor (x * bpow (- cexp beta fexp x)) + 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
(0 <= x * bpow (- cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
round beta fexp Zfloor (x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R

((x + ulp x) * bpow (- cexp beta fexp x) <= IZR (Zfloor (x * bpow (- cexp beta fexp x)) + 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
(0 <= x * bpow (- cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
round beta fexp Zfloor (x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R

(x * bpow (- cexp beta fexp x) + ulp x * bpow (- cexp beta fexp x) <= IZR (Zfloor (x * bpow (- cexp beta fexp x)) + 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
(0 <= x * bpow (- cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
round beta fexp Zfloor (x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R

(x * bpow (- cexp beta fexp x) + ulp x * bpow (- cexp beta fexp x) <= IZR (Zfloor (x * bpow (- cexp beta fexp x))) + 1)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
(0 <= x * bpow (- cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
round beta fexp Zfloor (x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R

(x * bpow (- cexp beta fexp x) <= IZR (Zfloor (x * bpow (- cexp beta fexp x))))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
(ulp x * bpow (- cexp beta fexp x) <= 1)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
(0 <= x * bpow (- cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
round beta fexp Zfloor (x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R

(F2R {| Fnum := Ztrunc (scaled_mantissa beta fexp x); Fexp := cexp beta fexp x |} * bpow (- cexp beta fexp x) <= IZR (Zfloor (F2R {| Fnum := Ztrunc (scaled_mantissa beta fexp x); Fexp := cexp beta fexp x |} * bpow (- cexp beta fexp x))))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
(ulp x * bpow (- cexp beta fexp x) <= 1)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
(0 <= x * bpow (- cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
round beta fexp Zfloor (x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R

(IZR (Fnum {| Fnum := Ztrunc (scaled_mantissa beta fexp x); Fexp := cexp beta fexp x |}) * bpow (Fexp {| Fnum := Ztrunc (scaled_mantissa beta fexp x); Fexp := cexp beta fexp x |}) * bpow (- cexp beta fexp x) <= IZR (Zfloor (IZR (Fnum {| Fnum := Ztrunc (scaled_mantissa beta fexp x); Fexp := cexp beta fexp x |}) * bpow (Fexp {| Fnum := Ztrunc (scaled_mantissa beta fexp x); Fexp := cexp beta fexp x |}) * bpow (- cexp beta fexp x))))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
(ulp x * bpow (- cexp beta fexp x) <= 1)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
(0 <= x * bpow (- cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
round beta fexp Zfloor (x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R

(IZR (Ztrunc (scaled_mantissa beta fexp x)) * bpow (cexp beta fexp x) * bpow (- cexp beta fexp x) <= IZR (Zfloor (IZR (Ztrunc (scaled_mantissa beta fexp x)) * bpow (cexp beta fexp x) * bpow (- cexp beta fexp x))))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
(ulp x * bpow (- cexp beta fexp x) <= 1)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
(0 <= x * bpow (- cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
round beta fexp Zfloor (x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R

(IZR (Ztrunc (scaled_mantissa beta fexp x)) * (bpow (cexp beta fexp x) * bpow (- cexp beta fexp x)) <= IZR (Zfloor (IZR (Ztrunc (scaled_mantissa beta fexp x)) * (bpow (cexp beta fexp x) * bpow (- cexp beta fexp x)))))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
(ulp x * bpow (- cexp beta fexp x) <= 1)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
(0 <= x * bpow (- cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
round beta fexp Zfloor (x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R

(IZR (Ztrunc (scaled_mantissa beta fexp x)) * bpow (cexp beta fexp x + - cexp beta fexp x) <= IZR (Zfloor (IZR (Ztrunc (scaled_mantissa beta fexp x)) * bpow (cexp beta fexp x + - cexp beta fexp x))))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
(ulp x * bpow (- cexp beta fexp x) <= 1)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
(0 <= x * bpow (- cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
round beta fexp Zfloor (x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R

(IZR (Ztrunc (scaled_mantissa beta fexp x)) * bpow 0 <= IZR (Zfloor (IZR (Ztrunc (scaled_mantissa beta fexp x)) * bpow 0)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
(ulp x * bpow (- cexp beta fexp x) <= 1)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
(0 <= x * bpow (- cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
round beta fexp Zfloor (x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R

(IZR (Ztrunc (scaled_mantissa beta fexp x)) <= IZR (Zfloor (IZR (Ztrunc (scaled_mantissa beta fexp x)))))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
(ulp x * bpow (- cexp beta fexp x) <= 1)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
(0 <= x * bpow (- cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
round beta fexp Zfloor (x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R

(IZR (Ztrunc (scaled_mantissa beta fexp x)) <= IZR (Ztrunc (scaled_mantissa beta fexp x)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
(ulp x * bpow (- cexp beta fexp x) <= 1)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
(0 <= x * bpow (- cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
round beta fexp Zfloor (x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R

(ulp x * bpow (- cexp beta fexp x) <= 1)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
(0 <= x * bpow (- cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
round beta fexp Zfloor (x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R

(bpow (cexp beta fexp x) * bpow (- cexp beta fexp x) <= 1)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
(0 <= x * bpow (- cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
round beta fexp Zfloor (x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R

(bpow (cexp beta fexp x) * bpow (- cexp beta fexp x) <= 1)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
(0 <= x * bpow (- cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
round beta fexp Zfloor (x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R

(bpow (cexp beta fexp x + - cexp beta fexp x) <= 1)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
(0 <= x * bpow (- cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
round beta fexp Zfloor (x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R

(bpow 0 <= 1)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
(0 <= x * bpow (- cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
round beta fexp Zfloor (x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R

(0 <= x * bpow (- cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
round beta fexp Zfloor (x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R

(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
(0 <= bpow (- cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
round beta fexp Zfloor (x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R

(0 <= bpow (- cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
round beta fexp Zfloor (x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R

round beta fexp Zfloor (x + eps) = x
(* . x=0 *)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp 0)%R

round beta fexp Zfloor eps = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp 0)%R
P:(0 < eps)%R

round beta fexp Zfloor eps = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp 0)%R
P:0%R = eps
round beta fexp Zfloor eps = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp 0)%R
P:(0 < eps)%R

F2R {| Fnum := Zfloor (eps * bpow (- fexp (mag beta eps))); Fexp := fexp (mag beta eps) |} = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp 0)%R
P:0%R = eps
round beta fexp Zfloor eps = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
P:(0 < eps)%R

(0 <= eps < (if Req_bool 0 0 then match negligible_exp with | Some n => bpow (fexp n) | None => 0 end else bpow (cexp beta fexp 0)))%R -> F2R {| Fnum := Zfloor (eps * bpow (- fexp (mag beta eps))); Fexp := fexp (mag beta eps) |} = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp 0)%R
P:0%R = eps
round beta fexp Zfloor eps = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
P:(0 < eps)%R

(0 <= eps < match negligible_exp with | Some n => bpow (fexp n) | None => 0 end)%R -> F2R {| Fnum := Zfloor (eps * bpow (- fexp (mag beta eps))); Fexp := fexp (mag beta eps) |} = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp 0)%R
P:0%R = eps
round beta fexp Zfloor eps = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
P:(0 < eps)%R

(forall n : Z, (fexp n < n)%Z) -> (0 <= eps < 0)%R -> F2R {| Fnum := Zfloor (eps * bpow (- fexp (mag beta eps))); Fexp := fexp (mag beta eps) |} = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
P:(0 < eps)%R
forall n : Z, (n <= fexp n)%Z -> (0 <= eps < bpow (fexp n))%R -> F2R {| Fnum := Zfloor (eps * bpow (- fexp (mag beta eps))); Fexp := fexp (mag beta eps) |} = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp 0)%R
P:0%R = eps
round beta fexp Zfloor eps = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
P:(0 < eps)%R
H1:(0 <= eps)%R
H2:(eps < 0)%R

F2R {| Fnum := Zfloor (eps * bpow (- fexp (mag beta eps))); Fexp := fexp (mag beta eps) |} = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
P:(0 < eps)%R
forall n : Z, (n <= fexp n)%Z -> (0 <= eps < bpow (fexp n))%R -> F2R {| Fnum := Zfloor (eps * bpow (- fexp (mag beta eps))); Fexp := fexp (mag beta eps) |} = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp 0)%R
P:0%R = eps
round beta fexp Zfloor eps = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
P:(0 < eps)%R

forall n : Z, (n <= fexp n)%Z -> (0 <= eps < bpow (fexp n))%R -> F2R {| Fnum := Zfloor (eps * bpow (- fexp (mag beta eps))); Fexp := fexp (mag beta eps) |} = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp 0)%R
P:0%R = eps
round beta fexp Zfloor eps = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
P:(0 < eps)%R
n:Z
Hn:(n <= fexp n)%Z
H:(0 <= eps < bpow (fexp n))%R

F2R {| Fnum := Zfloor (eps * bpow (- fexp (mag beta eps))); Fexp := fexp (mag beta eps) |} = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp 0)%R
P:0%R = eps
round beta fexp Zfloor eps = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
P:(0 < eps)%R
n:Z
Hn:(n <= fexp n)%Z
H:(0 <= eps < bpow (fexp n))%R

fexp (mag beta eps) = fexp n
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
P:(0 < eps)%R
n:Z
Hn:(n <= fexp n)%Z
H:(0 <= eps < bpow (fexp n))%R
H0:fexp (mag beta eps) = fexp n
F2R {| Fnum := Zfloor (eps * bpow (- fexp (mag beta eps))); Fexp := fexp (mag beta eps) |} = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp 0)%R
P:0%R = eps
round beta fexp Zfloor eps = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
P:(0 < eps)%R
n:Z
Hn:(n <= fexp n)%Z
H:(0 <= eps < bpow (fexp n))%R

(mag beta eps <= fexp n)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
P:(0 < eps)%R
n:Z
Hn:(n <= fexp n)%Z
H:(0 <= eps < bpow (fexp n))%R
H0:fexp (mag beta eps) = fexp n
F2R {| Fnum := Zfloor (eps * bpow (- fexp (mag beta eps))); Fexp := fexp (mag beta eps) |} = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp 0)%R
P:0%R = eps
round beta fexp Zfloor eps = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
P:(0 < eps)%R
n:Z
Hn:(n <= fexp n)%Z
H:(0 <= eps < bpow (fexp n))%R

(mag beta eps - 1 < fexp n)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
P:(0 < eps)%R
n:Z
Hn:(n <= fexp n)%Z
H:(0 <= eps < bpow (fexp n))%R
H0:fexp (mag beta eps) = fexp n
F2R {| Fnum := Zfloor (eps * bpow (- fexp (mag beta eps))); Fexp := fexp (mag beta eps) |} = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp 0)%R
P:0%R = eps
round beta fexp Zfloor eps = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
P:(0 < eps)%R
n:Z
Hn:(n <= fexp n)%Z
H:(0 <= eps < bpow (fexp n))%R

(bpow (mag beta eps - 1) < bpow (fexp n))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
P:(0 < eps)%R
n:Z
Hn:(n <= fexp n)%Z
H:(0 <= eps < bpow (fexp n))%R
H0:fexp (mag beta eps) = fexp n
F2R {| Fnum := Zfloor (eps * bpow (- fexp (mag beta eps))); Fexp := fexp (mag beta eps) |} = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp 0)%R
P:0%R = eps
round beta fexp Zfloor eps = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
P:(0 < eps)%R
n:Z
Hn:(n <= fexp n)%Z
H:(0 <= eps < bpow (fexp n))%R

(bpow (mag beta eps - 1) <= eps)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
P:(0 < eps)%R
n:Z
Hn:(n <= fexp n)%Z
H:(0 <= eps < bpow (fexp n))%R
H0:fexp (mag beta eps) = fexp n
F2R {| Fnum := Zfloor (eps * bpow (- fexp (mag beta eps))); Fexp := fexp (mag beta eps) |} = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp 0)%R
P:0%R = eps
round beta fexp Zfloor eps = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
P:(0 < eps)%R
n:Z
Hn:(n <= fexp n)%Z
H:(0 <= eps < bpow (fexp n))%R
e:Z
He:eps <> 0%R -> (bpow (e - 1) <= Rabs eps < bpow e)%R

(bpow (Build_mag_prop beta eps e He - 1) <= eps)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
P:(0 < eps)%R
n:Z
Hn:(n <= fexp n)%Z
H:(0 <= eps < bpow (fexp n))%R
H0:fexp (mag beta eps) = fexp n
F2R {| Fnum := Zfloor (eps * bpow (- fexp (mag beta eps))); Fexp := fexp (mag beta eps) |} = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp 0)%R
P:0%R = eps
round beta fexp Zfloor eps = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
P:(0 < eps)%R
n:Z
Hn:(n <= fexp n)%Z
H:(0 <= eps < bpow (fexp n))%R
e:Z
He:eps <> 0%R -> (bpow (e - 1) <= eps < bpow e)%R

(bpow (e - 1) <= eps)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
P:(0 < eps)%R
n:Z
Hn:(n <= fexp n)%Z
H:(0 <= eps < bpow (fexp n))%R
e:Z
He:eps <> 0%R -> (bpow (e - 1) <= Rabs eps < bpow e)%R
(0 <= eps)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
P:(0 < eps)%R
n:Z
Hn:(n <= fexp n)%Z
H:(0 <= eps < bpow (fexp n))%R
H0:fexp (mag beta eps) = fexp n
F2R {| Fnum := Zfloor (eps * bpow (- fexp (mag beta eps))); Fexp := fexp (mag beta eps) |} = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp 0)%R
P:0%R = eps
round beta fexp Zfloor eps = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
P:(0 < eps)%R
n:Z
Hn:(n <= fexp n)%Z
H:(0 <= eps < bpow (fexp n))%R
e:Z
He:eps <> 0%R -> (bpow (e - 1) <= Rabs eps < bpow e)%R

(0 <= eps)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
P:(0 < eps)%R
n:Z
Hn:(n <= fexp n)%Z
H:(0 <= eps < bpow (fexp n))%R
H0:fexp (mag beta eps) = fexp n
F2R {| Fnum := Zfloor (eps * bpow (- fexp (mag beta eps))); Fexp := fexp (mag beta eps) |} = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp 0)%R
P:0%R = eps
round beta fexp Zfloor eps = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
P:(0 < eps)%R
n:Z
Hn:(n <= fexp n)%Z
H:(0 <= eps < bpow (fexp n))%R
H0:fexp (mag beta eps) = fexp n

F2R {| Fnum := Zfloor (eps * bpow (- fexp (mag beta eps))); Fexp := fexp (mag beta eps) |} = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp 0)%R
P:0%R = eps
round beta fexp Zfloor eps = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
P:(0 < eps)%R
n:Z
Hn:(n <= fexp n)%Z
H:(0 <= eps < bpow (fexp n))%R
H0:fexp (mag beta eps) = fexp n

F2R {| Fnum := 0; Fexp := fexp (mag beta eps) |} = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
P:(0 < eps)%R
n:Z
Hn:(n <= fexp n)%Z
H:(0 <= eps < bpow (fexp n))%R
H0:fexp (mag beta eps) = fexp n
0%Z = Zfloor (eps * bpow (- fexp (mag beta eps)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp 0)%R
P:0%R = eps
round beta fexp Zfloor eps = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
P:(0 < eps)%R
n:Z
Hn:(n <= fexp n)%Z
H:(0 <= eps < bpow (fexp n))%R
H0:fexp (mag beta eps) = fexp n

0%Z = Zfloor (eps * bpow (- fexp (mag beta eps)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp 0)%R
P:0%R = eps
round beta fexp Zfloor eps = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
P:(0 < eps)%R
n:Z
Hn:(n <= fexp n)%Z
H:(0 <= eps < bpow (fexp n))%R
H0:fexp (mag beta eps) = fexp n

(0 <= eps * bpow (- fexp (mag beta eps)) < IZR (0 + 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp 0)%R
P:0%R = eps
round beta fexp Zfloor eps = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
P:(0 < eps)%R
n:Z
Hn:(n <= fexp n)%Z
H:(0 <= eps < bpow (fexp n))%R
H0:fexp (mag beta eps) = fexp n

(0 <= eps * bpow (- fexp (mag beta eps)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
P:(0 < eps)%R
n:Z
Hn:(n <= fexp n)%Z
H:(0 <= eps < bpow (fexp n))%R
H0:fexp (mag beta eps) = fexp n
(eps * bpow (- fexp (mag beta eps)) < IZR (0 + 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp 0)%R
P:0%R = eps
round beta fexp Zfloor eps = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
P:(0 < eps)%R
n:Z
Hn:(n <= fexp n)%Z
H:(0 <= eps < bpow (fexp n))%R
H0:fexp (mag beta eps) = fexp n

(0 <= eps)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
P:(0 < eps)%R
n:Z
Hn:(n <= fexp n)%Z
H:(0 <= eps < bpow (fexp n))%R
H0:fexp (mag beta eps) = fexp n
(0 <= bpow (- fexp (mag beta eps)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
P:(0 < eps)%R
n:Z
Hn:(n <= fexp n)%Z
H:(0 <= eps < bpow (fexp n))%R
H0:fexp (mag beta eps) = fexp n
(eps * bpow (- fexp (mag beta eps)) < IZR (0 + 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp 0)%R
P:0%R = eps
round beta fexp Zfloor eps = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
P:(0 < eps)%R
n:Z
Hn:(n <= fexp n)%Z
H:(0 <= eps < bpow (fexp n))%R
H0:fexp (mag beta eps) = fexp n

(0 <= bpow (- fexp (mag beta eps)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
P:(0 < eps)%R
n:Z
Hn:(n <= fexp n)%Z
H:(0 <= eps < bpow (fexp n))%R
H0:fexp (mag beta eps) = fexp n
(eps * bpow (- fexp (mag beta eps)) < IZR (0 + 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp 0)%R
P:0%R = eps
round beta fexp Zfloor eps = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
P:(0 < eps)%R
n:Z
Hn:(n <= fexp n)%Z
H:(0 <= eps < bpow (fexp n))%R
H0:fexp (mag beta eps) = fexp n

(eps * bpow (- fexp (mag beta eps)) < IZR (0 + 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp 0)%R
P:0%R = eps
round beta fexp Zfloor eps = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
P:(0 < eps)%R
n:Z
Hn:(n <= fexp n)%Z
H:(0 <= eps < bpow (fexp n))%R
H0:fexp (mag beta eps) = fexp n

(0 < bpow (fexp n))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
P:(0 < eps)%R
n:Z
Hn:(n <= fexp n)%Z
H:(0 <= eps < bpow (fexp n))%R
H0:fexp (mag beta eps) = fexp n
(eps * bpow (- fexp (mag beta eps)) * bpow (fexp n) < IZR (0 + 1) * bpow (fexp n))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp 0)%R
P:0%R = eps
round beta fexp Zfloor eps = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
P:(0 < eps)%R
n:Z
Hn:(n <= fexp n)%Z
H:(0 <= eps < bpow (fexp n))%R
H0:fexp (mag beta eps) = fexp n

(eps * bpow (- fexp (mag beta eps)) * bpow (fexp n) < IZR (0 + 1) * bpow (fexp n))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp 0)%R
P:0%R = eps
round beta fexp Zfloor eps = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
P:(0 < eps)%R
n:Z
Hn:(n <= fexp n)%Z
H:(0 <= eps < bpow (fexp n))%R
H0:fexp (mag beta eps) = fexp n

(eps * bpow (- fexp (mag beta eps) + fexp n) < IZR (0 + 1) * bpow (fexp n))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp 0)%R
P:0%R = eps
round beta fexp Zfloor eps = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
P:(0 < eps)%R
n:Z
Hn:(n <= fexp n)%Z
H:(0 <= eps < bpow (fexp n))%R
H0:fexp (mag beta eps) = fexp n

(eps * bpow 0 < IZR (0 + 1) * bpow (fexp n))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp 0)%R
P:0%R = eps
round beta fexp Zfloor eps = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
P:(0 < eps)%R
n:Z
Hn:(n <= fexp n)%Z
H:(0 <= eps < bpow (fexp n))%R
H0:fexp (mag beta eps) = fexp n

(eps < bpow (fexp n))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp 0)%R
P:0%R = eps
round beta fexp Zfloor eps = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp 0)%R
P:0%R = eps

round beta fexp Zfloor eps = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:0%R = x
Fx:F x
eps:R
Heps:(0 <= eps < ulp 0)%R
P:0%R = eps

Valid_rnd Zfloor
apply valid_rnd_DN. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, (0 <= x)%R -> F x -> forall eps : R, (0 < eps <= ulp x)%R -> round beta fexp Zceil (x + eps) = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, (0 <= x)%R -> F x -> forall eps : R, (0 < eps <= ulp x)%R -> round beta fexp Zceil (x + eps) = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R

(0 < eps <= ulp x)%R -> round beta fexp Zceil (x + eps) = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:(0 < x)%R

(0 < eps <= ulp x)%R -> round beta fexp Zceil (x + eps) = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
(0 < eps <= ulp x)%R -> round beta fexp Zceil (x + eps) = (x + ulp x)%R
(* . 0 < x *)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:(0 < x)%R
Heps1:(0 < eps)%R
Heps2:(eps < ulp x)%R

round beta fexp Zceil (x + eps) = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:(0 < x)%R
Heps1:(0 < eps)%R
Heps2:eps = ulp x
round beta fexp Zceil (x + eps) = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
(0 < eps <= ulp x)%R -> round beta fexp Zceil (x + eps) = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:(0 < x)%R
Heps1:(0 < eps)%R
Heps2:(eps < ulp x)%R

(0 <= eps < ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:(0 < x)%R
Heps1:(0 < eps)%R
Heps2:(eps < ulp x)%R
Heps:(0 <= eps < ulp x)%R
round beta fexp Zceil (x + eps) = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:(0 < x)%R
Heps1:(0 < eps)%R
Heps2:eps = ulp x
round beta fexp Zceil (x + eps) = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
(0 < eps <= ulp x)%R -> round beta fexp Zceil (x + eps) = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:(0 < x)%R
Heps1:(0 < eps)%R
Heps2:(eps < ulp x)%R

(0 <= eps)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:(0 < x)%R
Heps1:(0 < eps)%R
Heps2:(eps < ulp x)%R
(eps < ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:(0 < x)%R
Heps1:(0 < eps)%R
Heps2:(eps < ulp x)%R
Heps:(0 <= eps < ulp x)%R
round beta fexp Zceil (x + eps) = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:(0 < x)%R
Heps1:(0 < eps)%R
Heps2:eps = ulp x
round beta fexp Zceil (x + eps) = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
(0 < eps <= ulp x)%R -> round beta fexp Zceil (x + eps) = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:(0 < x)%R
Heps1:(0 < eps)%R
Heps2:(eps < ulp x)%R

(eps < ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:(0 < x)%R
Heps1:(0 < eps)%R
Heps2:(eps < ulp x)%R
Heps:(0 <= eps < ulp x)%R
round beta fexp Zceil (x + eps) = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:(0 < x)%R
Heps1:(0 < eps)%R
Heps2:eps = ulp x
round beta fexp Zceil (x + eps) = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
(0 < eps <= ulp x)%R -> round beta fexp Zceil (x + eps) = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:(0 < x)%R
Heps1:(0 < eps)%R
Heps2:(eps < ulp x)%R
Heps:(0 <= eps < ulp x)%R

round beta fexp Zceil (x + eps) = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:(0 < x)%R
Heps1:(0 < eps)%R
Heps2:eps = ulp x
round beta fexp Zceil (x + eps) = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
(0 < eps <= ulp x)%R -> round beta fexp Zceil (x + eps) = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:(0 < x)%R
Heps1:(0 < eps)%R
Heps2:(eps < ulp x)%R
Heps:(0 <= eps < ulp x)%R
Hd:round beta fexp Zfloor (x + eps) = x

round beta fexp Zceil (x + eps) = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:(0 < x)%R
Heps1:(0 < eps)%R
Heps2:eps = ulp x
round beta fexp Zceil (x + eps) = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
(0 < eps <= ulp x)%R -> round beta fexp Zceil (x + eps) = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:(0 < x)%R
Heps1:(0 < eps)%R
Heps2:(eps < ulp x)%R
Heps:(0 <= eps < ulp x)%R
Hd:round beta fexp Zfloor (x + eps) = x

(round beta fexp Zfloor (x + eps) + ulp (x + eps))%R = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:(0 < x)%R
Heps1:(0 < eps)%R
Heps2:(eps < ulp x)%R
Heps:(0 <= eps < ulp x)%R
Hd:round beta fexp Zfloor (x + eps) = x
~ F (x + eps)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:(0 < x)%R
Heps1:(0 < eps)%R
Heps2:eps = ulp x
round beta fexp Zceil (x + eps) = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
(0 < eps <= ulp x)%R -> round beta fexp Zceil (x + eps) = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:(0 < x)%R
Heps1:(0 < eps)%R
Heps2:(eps < ulp x)%R
Heps:(0 <= eps < ulp x)%R
Hd:round beta fexp Zfloor (x + eps) = x

(x + ulp (x + eps))%R = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:(0 < x)%R
Heps1:(0 < eps)%R
Heps2:(eps < ulp x)%R
Heps:(0 <= eps < ulp x)%R
Hd:round beta fexp Zfloor (x + eps) = x
~ F (x + eps)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:(0 < x)%R
Heps1:(0 < eps)%R
Heps2:eps = ulp x
round beta fexp Zceil (x + eps) = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
(0 < eps <= ulp x)%R -> round beta fexp Zceil (x + eps) = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:(0 < x)%R
Heps1:(0 < eps)%R
Heps2:(eps < ulp x)%R
Heps:(0 <= eps < ulp x)%R
Hd:round beta fexp Zfloor (x + eps) = x

(x + bpow (cexp beta fexp (x + eps)))%R = (x + bpow (cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:(0 < x)%R
Heps1:(0 < eps)%R
Heps2:(eps < ulp x)%R
Heps:(0 <= eps < ulp x)%R
Hd:round beta fexp Zfloor (x + eps) = x
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:(0 < x)%R
Heps1:(0 < eps)%R
Heps2:(eps < ulp x)%R
Heps:(0 <= eps < ulp x)%R
Hd:round beta fexp Zfloor (x + eps) = x
(x + eps)%R <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:(0 < x)%R
Heps1:(0 < eps)%R
Heps2:(eps < ulp x)%R
Heps:(0 <= eps < ulp x)%R
Hd:round beta fexp Zfloor (x + eps) = x
~ F (x + eps)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:(0 < x)%R
Heps1:(0 < eps)%R
Heps2:eps = ulp x
round beta fexp Zceil (x + eps) = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
(0 < eps <= ulp x)%R -> round beta fexp Zceil (x + eps) = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:(0 < x)%R
Heps1:(0 < eps)%R
Heps2:(eps < ulp x)%R
Heps:(0 <= eps < ulp x)%R
Hd:round beta fexp Zfloor (x + eps) = x

(x + bpow (fexp (mag beta (x + eps))))%R = (x + bpow (fexp (mag beta x)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:(0 < x)%R
Heps1:(0 < eps)%R
Heps2:(eps < ulp x)%R
Heps:(0 <= eps < ulp x)%R
Hd:round beta fexp Zfloor (x + eps) = x
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:(0 < x)%R
Heps1:(0 < eps)%R
Heps2:(eps < ulp x)%R
Heps:(0 <= eps < ulp x)%R
Hd:round beta fexp Zfloor (x + eps) = x
(x + eps)%R <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:(0 < x)%R
Heps1:(0 < eps)%R
Heps2:(eps < ulp x)%R
Heps:(0 <= eps < ulp x)%R
Hd:round beta fexp Zfloor (x + eps) = x
~ F (x + eps)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:(0 < x)%R
Heps1:(0 < eps)%R
Heps2:eps = ulp x
round beta fexp Zceil (x + eps) = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
(0 < eps <= ulp x)%R -> round beta fexp Zceil (x + eps) = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:(0 < x)%R
Heps1:(0 < eps)%R
Heps2:(eps < ulp x)%R
Heps:(0 <= eps < ulp x)%R
Hd:round beta fexp Zfloor (x + eps) = x

x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:(0 < x)%R
Heps1:(0 < eps)%R
Heps2:(eps < ulp x)%R
Heps:(0 <= eps < ulp x)%R
Hd:round beta fexp Zfloor (x + eps) = x
(x + eps)%R <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:(0 < x)%R
Heps1:(0 < eps)%R
Heps2:(eps < ulp x)%R
Heps:(0 <= eps < ulp x)%R
Hd:round beta fexp Zfloor (x + eps) = x
~ F (x + eps)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:(0 < x)%R
Heps1:(0 < eps)%R
Heps2:eps = ulp x
round beta fexp Zceil (x + eps) = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
(0 < eps <= ulp x)%R -> round beta fexp Zceil (x + eps) = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:(0 < x)%R
Heps1:(0 < eps)%R
Heps2:(eps < ulp x)%R
Heps:(0 <= eps < ulp x)%R
Hd:round beta fexp Zfloor (x + eps) = x

(x + eps)%R <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:(0 < x)%R
Heps1:(0 < eps)%R
Heps2:(eps < ulp x)%R
Heps:(0 <= eps < ulp x)%R
Hd:round beta fexp Zfloor (x + eps) = x
~ F (x + eps)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:(0 < x)%R
Heps1:(0 < eps)%R
Heps2:eps = ulp x
round beta fexp Zceil (x + eps) = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
(0 < eps <= ulp x)%R -> round beta fexp Zceil (x + eps) = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:(0 < x)%R
Heps1:(0 < eps)%R
Heps2:(eps < ulp x)%R
Heps:(0 <= eps < ulp x)%R
Hd:round beta fexp Zfloor (x + eps) = x

~ F (x + eps)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:(0 < x)%R
Heps1:(0 < eps)%R
Heps2:eps = ulp x
round beta fexp Zceil (x + eps) = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
(0 < eps <= ulp x)%R -> round beta fexp Zceil (x + eps) = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:(0 < x)%R
Heps1:(0 < eps)%R
Heps2:(eps < ulp x)%R
Heps:(0 <= eps < ulp x)%R
Hd:round beta fexp Zfloor (x + eps) = x
Fs:F (x + eps)

False
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:(0 < x)%R
Heps1:(0 < eps)%R
Heps2:eps = ulp x
round beta fexp Zceil (x + eps) = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
(0 < eps <= ulp x)%R -> round beta fexp Zceil (x + eps) = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:(0 < x)%R
Heps1:(0 < eps)%R
Heps2:(eps < ulp x)%R
Heps:(0 <= eps < ulp x)%R
Hd:(x + eps)%R = x
Fs:F (x + eps)

False
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:(0 < x)%R
Heps1:(0 < eps)%R
Heps2:eps = ulp x
round beta fexp Zceil (x + eps) = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
(0 < eps <= ulp x)%R -> round beta fexp Zceil (x + eps) = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:(0 < x)%R
Heps1:(0 < eps)%R
Heps2:(eps < ulp x)%R
Heps:(0 <= eps < ulp x)%R
Hd:(x + eps)%R = x
Fs:F (x + eps)

(x + eps > x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:(0 < x)%R
Heps1:(0 < eps)%R
Heps2:eps = ulp x
round beta fexp Zceil (x + eps) = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
(0 < eps <= ulp x)%R -> round beta fexp Zceil (x + eps) = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:(0 < x)%R
Heps1:(0 < eps)%R
Heps2:(eps < ulp x)%R
Heps:(0 <= eps < ulp x)%R
Hd:(x + eps)%R = x
Fs:F (x + eps)

(x + eps > x + 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:(0 < x)%R
Heps1:(0 < eps)%R
Heps2:eps = ulp x
round beta fexp Zceil (x + eps) = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
(0 < eps <= ulp x)%R -> round beta fexp Zceil (x + eps) = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:(0 < x)%R
Heps1:(0 < eps)%R
Heps2:eps = ulp x

round beta fexp Zceil (x + eps) = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
(0 < eps <= ulp x)%R -> round beta fexp Zceil (x + eps) = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:(0 < x)%R
Heps1:(0 < eps)%R
Heps2:eps = ulp x

round beta fexp Zceil (x + ulp x) = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
(0 < eps <= ulp x)%R -> round beta fexp Zceil (x + eps) = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:(0 < x)%R
Heps1:(0 < eps)%R
Heps2:eps = ulp x

F (x + ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
(0 < eps <= ulp x)%R -> round beta fexp Zceil (x + eps) = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x

(0 < eps <= ulp x)%R -> round beta fexp Zceil (x + eps) = (x + ulp x)%R
(* . x=0 *)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x

(0 < eps <= ulp 0)%R -> round beta fexp Zceil eps = ulp 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R

round beta fexp Zceil eps = ulp 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R

(eps < ulp 0)%R -> round beta fexp Zceil eps = ulp 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
eps = ulp 0 -> round beta fexp Zceil eps = ulp 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R

(eps < ulp 0)%R -> F2R {| Fnum := Zceil (eps * bpow (- fexp (mag beta eps))); Fexp := fexp (mag beta eps) |} = ulp 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
eps = ulp 0 -> round beta fexp Zceil eps = ulp 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R

(eps < (if Req_bool 0 0 then match negligible_exp with | Some n => bpow (fexp n) | None => 0 end else bpow (cexp beta fexp 0)))%R -> F2R {| Fnum := Zceil (eps * bpow (- fexp (mag beta eps))); Fexp := fexp (mag beta eps) |} = (if Req_bool 0 0 then match negligible_exp with | Some n => bpow (fexp n) | None => 0%R end else bpow (cexp beta fexp 0))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
eps = ulp 0 -> round beta fexp Zceil eps = ulp 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R

(eps < match negligible_exp with | Some n => bpow (fexp n) | None => 0 end)%R -> F2R {| Fnum := Zceil (eps * bpow (- fexp (mag beta eps))); Fexp := fexp (mag beta eps) |} = match negligible_exp with | Some n => bpow (fexp n) | None => 0%R end
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
eps = ulp 0 -> round beta fexp Zceil eps = ulp 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R

(forall n : Z, (fexp n < n)%Z) -> (eps < 0)%R -> F2R {| Fnum := Zceil (eps * bpow (- fexp (mag beta eps))); Fexp := fexp (mag beta eps) |} = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
forall n : Z, (n <= fexp n)%Z -> (eps < bpow (fexp n))%R -> F2R {| Fnum := Zceil (eps * bpow (- fexp (mag beta eps))); Fexp := fexp (mag beta eps) |} = bpow (fexp n)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
eps = ulp 0 -> round beta fexp Zceil eps = ulp 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R

forall n : Z, (n <= fexp n)%Z -> (eps < bpow (fexp n))%R -> F2R {| Fnum := Zceil (eps * bpow (- fexp (mag beta eps))); Fexp := fexp (mag beta eps) |} = bpow (fexp n)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
eps = ulp 0 -> round beta fexp Zceil eps = ulp 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
n:Z
Hn:(n <= fexp n)%Z
H:(eps < bpow (fexp n))%R

F2R {| Fnum := Zceil (eps * bpow (- fexp (mag beta eps))); Fexp := fexp (mag beta eps) |} = bpow (fexp n)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
eps = ulp 0 -> round beta fexp Zceil eps = ulp 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
n:Z
Hn:(n <= fexp n)%Z
H:(eps < bpow (fexp n))%R

fexp (mag beta eps) = fexp n
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
n:Z
Hn:(n <= fexp n)%Z
H:(eps < bpow (fexp n))%R
H0:fexp (mag beta eps) = fexp n
F2R {| Fnum := Zceil (eps * bpow (- fexp (mag beta eps))); Fexp := fexp (mag beta eps) |} = bpow (fexp n)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
eps = ulp 0 -> round beta fexp Zceil eps = ulp 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
n:Z
Hn:(n <= fexp n)%Z
H:(eps < bpow (fexp n))%R

(mag beta eps <= fexp n)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
n:Z
Hn:(n <= fexp n)%Z
H:(eps < bpow (fexp n))%R
H0:fexp (mag beta eps) = fexp n
F2R {| Fnum := Zceil (eps * bpow (- fexp (mag beta eps))); Fexp := fexp (mag beta eps) |} = bpow (fexp n)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
eps = ulp 0 -> round beta fexp Zceil eps = ulp 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
n:Z
Hn:(n <= fexp n)%Z
H:(eps < bpow (fexp n))%R

(mag beta eps - 1 < fexp n)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
n:Z
Hn:(n <= fexp n)%Z
H:(eps < bpow (fexp n))%R
H0:fexp (mag beta eps) = fexp n
F2R {| Fnum := Zceil (eps * bpow (- fexp (mag beta eps))); Fexp := fexp (mag beta eps) |} = bpow (fexp n)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
eps = ulp 0 -> round beta fexp Zceil eps = ulp 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
n:Z
Hn:(n <= fexp n)%Z
H:(eps < bpow (fexp n))%R

(bpow (mag beta eps - 1) < bpow (fexp n))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
n:Z
Hn:(n <= fexp n)%Z
H:(eps < bpow (fexp n))%R
H0:fexp (mag beta eps) = fexp n
F2R {| Fnum := Zceil (eps * bpow (- fexp (mag beta eps))); Fexp := fexp (mag beta eps) |} = bpow (fexp n)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
eps = ulp 0 -> round beta fexp Zceil eps = ulp 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
n:Z
Hn:(n <= fexp n)%Z
H:(eps < bpow (fexp n))%R

(bpow (mag beta eps - 1) <= eps)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
n:Z
Hn:(n <= fexp n)%Z
H:(eps < bpow (fexp n))%R
H0:fexp (mag beta eps) = fexp n
F2R {| Fnum := Zceil (eps * bpow (- fexp (mag beta eps))); Fexp := fexp (mag beta eps) |} = bpow (fexp n)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
eps = ulp 0 -> round beta fexp Zceil eps = ulp 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
n:Z
Hn:(n <= fexp n)%Z
H:(eps < bpow (fexp n))%R
e:Z
He:eps <> 0%R -> (bpow (e - 1) <= Rabs eps < bpow e)%R

(bpow (Build_mag_prop beta eps e He - 1) <= eps)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
n:Z
Hn:(n <= fexp n)%Z
H:(eps < bpow (fexp n))%R
H0:fexp (mag beta eps) = fexp n
F2R {| Fnum := Zceil (eps * bpow (- fexp (mag beta eps))); Fexp := fexp (mag beta eps) |} = bpow (fexp n)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
eps = ulp 0 -> round beta fexp Zceil eps = ulp 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
n:Z
Hn:(n <= fexp n)%Z
H:(eps < bpow (fexp n))%R
e:Z
He:eps <> 0%R -> (bpow (e - 1) <= eps < bpow e)%R

(bpow (e - 1) <= eps)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
n:Z
Hn:(n <= fexp n)%Z
H:(eps < bpow (fexp n))%R
e:Z
He:eps <> 0%R -> (bpow (e - 1) <= Rabs eps < bpow e)%R
(0 <= eps)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
n:Z
Hn:(n <= fexp n)%Z
H:(eps < bpow (fexp n))%R
H0:fexp (mag beta eps) = fexp n
F2R {| Fnum := Zceil (eps * bpow (- fexp (mag beta eps))); Fexp := fexp (mag beta eps) |} = bpow (fexp n)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
eps = ulp 0 -> round beta fexp Zceil eps = ulp 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
n:Z
Hn:(n <= fexp n)%Z
H:(eps < bpow (fexp n))%R
e:Z
He:eps <> 0%R -> (bpow (e - 1) <= Rabs eps < bpow e)%R

(0 <= eps)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
n:Z
Hn:(n <= fexp n)%Z
H:(eps < bpow (fexp n))%R
H0:fexp (mag beta eps) = fexp n
F2R {| Fnum := Zceil (eps * bpow (- fexp (mag beta eps))); Fexp := fexp (mag beta eps) |} = bpow (fexp n)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
eps = ulp 0 -> round beta fexp Zceil eps = ulp 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
n:Z
Hn:(n <= fexp n)%Z
H:(eps < bpow (fexp n))%R
H0:fexp (mag beta eps) = fexp n

F2R {| Fnum := Zceil (eps * bpow (- fexp (mag beta eps))); Fexp := fexp (mag beta eps) |} = bpow (fexp n)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
eps = ulp 0 -> round beta fexp Zceil eps = ulp 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
n:Z
Hn:(n <= fexp n)%Z
H:(eps < bpow (fexp n))%R
H0:fexp (mag beta eps) = fexp n

F2R {| Fnum := 1; Fexp := fexp (mag beta eps) |} = bpow (fexp n)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
n:Z
Hn:(n <= fexp n)%Z
H:(eps < bpow (fexp n))%R
H0:fexp (mag beta eps) = fexp n
1%Z = Zceil (eps * bpow (- fexp (mag beta eps)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
eps = ulp 0 -> round beta fexp Zceil eps = ulp 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
n:Z
Hn:(n <= fexp n)%Z
H:(eps < bpow (fexp n))%R
H0:fexp (mag beta eps) = fexp n

1%Z = Zceil (eps * bpow (- fexp (mag beta eps)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
eps = ulp 0 -> round beta fexp Zceil eps = ulp 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
n:Z
Hn:(n <= fexp n)%Z
H:(eps < bpow (fexp n))%R
H0:fexp (mag beta eps) = fexp n

(IZR (1 - 1) < eps * bpow (- fexp (mag beta eps)) <= 1)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
eps = ulp 0 -> round beta fexp Zceil eps = ulp 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
n:Z
Hn:(n <= fexp n)%Z
H:(eps < bpow (fexp n))%R
H0:fexp (mag beta eps) = fexp n

(IZR (1 - 1) < eps * bpow (- fexp (mag beta eps)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
n:Z
Hn:(n <= fexp n)%Z
H:(eps < bpow (fexp n))%R
H0:fexp (mag beta eps) = fexp n
(eps * bpow (- fexp (mag beta eps)) <= 1)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
eps = ulp 0 -> round beta fexp Zceil eps = ulp 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
n:Z
Hn:(n <= fexp n)%Z
H:(eps < bpow (fexp n))%R
H0:fexp (mag beta eps) = fexp n

(0 < eps)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
n:Z
Hn:(n <= fexp n)%Z
H:(eps < bpow (fexp n))%R
H0:fexp (mag beta eps) = fexp n
(0 < bpow (- fexp (mag beta eps)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
n:Z
Hn:(n <= fexp n)%Z
H:(eps < bpow (fexp n))%R
H0:fexp (mag beta eps) = fexp n
(eps * bpow (- fexp (mag beta eps)) <= 1)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
eps = ulp 0 -> round beta fexp Zceil eps = ulp 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
n:Z
Hn:(n <= fexp n)%Z
H:(eps < bpow (fexp n))%R
H0:fexp (mag beta eps) = fexp n

(0 < bpow (- fexp (mag beta eps)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
n:Z
Hn:(n <= fexp n)%Z
H:(eps < bpow (fexp n))%R
H0:fexp (mag beta eps) = fexp n
(eps * bpow (- fexp (mag beta eps)) <= 1)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
eps = ulp 0 -> round beta fexp Zceil eps = ulp 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
n:Z
Hn:(n <= fexp n)%Z
H:(eps < bpow (fexp n))%R
H0:fexp (mag beta eps) = fexp n

(eps * bpow (- fexp (mag beta eps)) <= 1)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
eps = ulp 0 -> round beta fexp Zceil eps = ulp 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
n:Z
Hn:(n <= fexp n)%Z
H:(eps < bpow (fexp n))%R
H0:fexp (mag beta eps) = fexp n

(0 < bpow (fexp n))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
n:Z
Hn:(n <= fexp n)%Z
H:(eps < bpow (fexp n))%R
H0:fexp (mag beta eps) = fexp n
(eps * bpow (- fexp (mag beta eps)) * bpow (fexp n) <= 1 * bpow (fexp n))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
eps = ulp 0 -> round beta fexp Zceil eps = ulp 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
n:Z
Hn:(n <= fexp n)%Z
H:(eps < bpow (fexp n))%R
H0:fexp (mag beta eps) = fexp n

(eps * bpow (- fexp (mag beta eps)) * bpow (fexp n) <= 1 * bpow (fexp n))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
eps = ulp 0 -> round beta fexp Zceil eps = ulp 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
n:Z
Hn:(n <= fexp n)%Z
H:(eps < bpow (fexp n))%R
H0:fexp (mag beta eps) = fexp n

(eps * bpow (- fexp (mag beta eps) + fexp n) <= 1 * bpow (fexp n))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
eps = ulp 0 -> round beta fexp Zceil eps = ulp 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
n:Z
Hn:(n <= fexp n)%Z
H:(eps < bpow (fexp n))%R
H0:fexp (mag beta eps) = fexp n

(eps * bpow 0 <= 1 * bpow (fexp n))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
eps = ulp 0 -> round beta fexp Zceil eps = ulp 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
n:Z
Hn:(n <= fexp n)%Z
H:(eps < bpow (fexp n))%R
H0:fexp (mag beta eps) = fexp n

(eps <= bpow (fexp n))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
eps = ulp 0 -> round beta fexp Zceil eps = ulp 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R

eps = ulp 0 -> round beta fexp Zceil eps = ulp 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
P:eps = ulp 0

round beta fexp Zceil (ulp 0) = ulp 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Zx:(0 <= x)%R
Fx:F x
eps:R
Zx1:0%R = x
Heps:(0 < eps <= ulp 0)%R
P:eps = ulp 0

F (ulp 0)
apply generic_format_ulp_0. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, (0 < x)%R -> F x -> forall eps : R, (0 < eps <= ulp (pred x))%R -> round beta fexp Zceil (pred x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, (0 < x)%R -> F x -> forall eps : R, (0 < eps <= ulp (pred x))%R -> round beta fexp Zceil (pred x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred x))%R

round beta fexp Zceil (pred x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred x))%R

(pred x + ulp (pred x))%R = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred x))%R
(0 <= pred x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred x))%R
F (pred x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred x))%R

(pred_pos x + ulp (pred_pos x))%R = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred x))%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred x))%R
(0 <= pred x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred x))%R
F (pred x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred x))%R

(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred x))%R
(0 <= pred x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred x))%R
F (pred x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred x))%R

(0 <= pred x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred x))%R
F (pred x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred x))%R

F (pred x)
apply generic_format_pred; trivial. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, (0 < x)%R -> F x -> forall eps : R, (0 < eps <= ulp (pred x))%R -> round beta fexp Zfloor (x - eps) = pred x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, (0 < x)%R -> F x -> forall eps : R, (0 < eps <= ulp (pred x))%R -> round beta fexp Zfloor (x - eps) = pred x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hpx:(0 < x)%R
Fx:F x
eps:R

(0 < eps <= ulp (pred x))%R -> round beta fexp Zfloor (x - eps) = pred x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hpx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred_pos x))%R

round beta fexp Zfloor (x - eps) = pred_pos x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hpx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred_pos x))%R

round beta fexp Zfloor (pred_pos x + (ulp (pred_pos x) - eps)) = pred_pos x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hpx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred_pos x))%R
(pred_pos x + (ulp (pred_pos x) - eps))%R = (x - eps)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hpx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred_pos x))%R

round beta fexp Zfloor (pred_pos x + (ulp (pred_pos x) - eps)) = pred_pos x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hpx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred_pos x))%R
(pred_pos x + (ulp (pred_pos x) - eps))%R = (pred_pos x + ulp (pred_pos x) - eps)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hpx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred_pos x))%R

round beta fexp Zfloor (pred_pos x + (ulp (pred_pos x) - eps)) = pred_pos x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hpx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred_pos x))%R

(0 <= pred_pos x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hpx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred_pos x))%R
F (pred_pos x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hpx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred_pos x))%R
(0 <= ulp (pred_pos x) - eps < ulp (pred_pos x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hpx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred_pos x))%R

F (pred_pos x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hpx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred_pos x))%R
(0 <= ulp (pred_pos x) - eps < ulp (pred_pos x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hpx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred_pos x))%R

(0 <= ulp (pred_pos x) - eps < ulp (pred_pos x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hpx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred_pos x))%R

(0 <= ulp (pred_pos x) - eps)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hpx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred_pos x))%R
(ulp (pred_pos x) - eps < ulp (pred_pos x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hpx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred_pos x))%R

(eps <= ulp (pred_pos x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hpx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred_pos x))%R
(ulp (pred_pos x) - eps < ulp (pred_pos x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hpx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred_pos x))%R

(ulp (pred_pos x) - eps < ulp (pred_pos x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hpx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred_pos x))%R

(ulp (pred_pos x) - eps < ulp (pred_pos x) + 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hpx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred_pos x))%R

(- eps < 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hpx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred_pos x))%R

(- eps < - 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hpx:(0 < x)%R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred_pos x))%R

(0 < eps)%R
now apply Heps. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, F x -> forall eps : R, (0 <= eps < (if Rle_bool 0 x then ulp x else ulp (pred (- x))))%R -> round beta fexp Zfloor (x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, F x -> forall eps : R, (0 <= eps < (if Rle_bool 0 x then ulp x else ulp (pred (- x))))%R -> round beta fexp Zfloor (x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 <= eps < (if Rle_bool 0 x then ulp x else ulp (pred (- x))))%R

round beta fexp Zfloor (x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 <= eps < (if Rle_bool 0 x then ulp x else ulp (pred (- x))))%R
Zx:(0 <= x)%R

round beta fexp Zfloor (x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 <= eps < (if Rle_bool 0 x then ulp x else ulp (pred (- x))))%R
Zx:(x < 0)%R
round beta fexp Zfloor (x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 <= eps < (if Rle_bool 0 x then ulp x else ulp (pred (- x))))%R
Zx:(0 <= x)%R

(0 <= eps < ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 <= eps < (if Rle_bool 0 x then ulp x else ulp (pred (- x))))%R
Zx:(x < 0)%R
round beta fexp Zfloor (x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 <= eps < (if Rle_bool 0 x then ulp x else ulp (pred (- x))))%R
Zx:(0 <= x)%R

(eps < ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 <= eps < (if Rle_bool 0 x then ulp x else ulp (pred (- x))))%R
Zx:(x < 0)%R
round beta fexp Zfloor (x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 <= eps < ulp x)%R
Zx:(0 <= x)%R

(eps < ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 <= eps < (if Rle_bool 0 x then ulp x else ulp (pred (- x))))%R
Zx:(x < 0)%R
round beta fexp Zfloor (x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 <= eps < (if Rle_bool 0 x then ulp x else ulp (pred (- x))))%R
Zx:(x < 0)%R

round beta fexp Zfloor (x + eps) = x
(* *)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 <= eps < ulp (pred (- x)))%R
Zx:(x < 0)%R

round beta fexp Zfloor (x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 <= eps < ulp (pred (- x)))%R
Zx:(x < 0)%R

round beta fexp Zfloor (- - (x + eps)) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 <= eps < ulp (pred (- x)))%R
Zx:(x < 0)%R

round beta fexp Zfloor (- - (x + eps)) = (- - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 <= eps < ulp (pred (- x)))%R
Zx:(x < 0)%R

(- round beta fexp Zceil (- (x + eps)))%R = (- - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 <= eps < ulp (pred (- x)))%R
Zx:(x < 0)%R

round beta fexp Zceil (- (x + eps)) = (- x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 <= eps < ulp (pred (- x)))%R
Zx:(x < 0)%R

round beta fexp Zceil (pred (- x) + (ulp (pred (- x)) - eps)) = (- x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 <= eps < ulp (pred (- x)))%R
Zx:(x < 0)%R
(pred (- x) + (ulp (pred (- x)) - eps))%R = (- (x + eps))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 <= eps < ulp (pred (- x)))%R
Zx:(x < 0)%R

(0 < - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 <= eps < ulp (pred (- x)))%R
Zx:(x < 0)%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 <= eps < ulp (pred (- x)))%R
Zx:(x < 0)%R
(0 < ulp (pred (- x)) - eps <= ulp (pred (- x)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 <= eps < ulp (pred (- x)))%R
Zx:(x < 0)%R
(pred (- x) + (ulp (pred (- x)) - eps))%R = (- (x + eps))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 <= eps < ulp (pred (- x)))%R
Zx:(x < 0)%R

F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 <= eps < ulp (pred (- x)))%R
Zx:(x < 0)%R
(0 < ulp (pred (- x)) - eps <= ulp (pred (- x)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 <= eps < ulp (pred (- x)))%R
Zx:(x < 0)%R
(pred (- x) + (ulp (pred (- x)) - eps))%R = (- (x + eps))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 <= eps < ulp (pred (- x)))%R
Zx:(x < 0)%R

(0 < ulp (pred (- x)) - eps <= ulp (pred (- x)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 <= eps < ulp (pred (- x)))%R
Zx:(x < 0)%R
(pred (- x) + (ulp (pred (- x)) - eps))%R = (- (x + eps))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 <= eps < ulp (pred (- x)))%R
Zx:(x < 0)%R

(0 < ulp (pred (- x)) - eps)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 <= eps < ulp (pred (- x)))%R
Zx:(x < 0)%R
(ulp (pred (- x)) - eps <= ulp (pred (- x)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 <= eps < ulp (pred (- x)))%R
Zx:(x < 0)%R
(pred (- x) + (ulp (pred (- x)) - eps))%R = (- (x + eps))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 <= eps < ulp (pred (- x)))%R
Zx:(x < 0)%R

(eps < ulp (pred (- x)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 <= eps < ulp (pred (- x)))%R
Zx:(x < 0)%R
(ulp (pred (- x)) - eps <= ulp (pred (- x)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 <= eps < ulp (pred (- x)))%R
Zx:(x < 0)%R
(pred (- x) + (ulp (pred (- x)) - eps))%R = (- (x + eps))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 <= eps < ulp (pred (- x)))%R
Zx:(x < 0)%R

(ulp (pred (- x)) - eps <= ulp (pred (- x)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 <= eps < ulp (pred (- x)))%R
Zx:(x < 0)%R
(pred (- x) + (ulp (pred (- x)) - eps))%R = (- (x + eps))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 <= eps < ulp (pred (- x)))%R
Zx:(x < 0)%R

(0 <= eps)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 <= eps < ulp (pred (- x)))%R
Zx:(x < 0)%R
(pred (- x) + (ulp (pred (- x)) - eps))%R = (- (x + eps))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 <= eps < ulp (pred (- x)))%R
Zx:(x < 0)%R

(pred (- x) + (ulp (pred (- x)) - eps))%R = (- (x + eps))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 <= eps < ulp (pred (- x)))%R
Zx:(x < 0)%R

(- succ (- - x) + (ulp (- succ (- - x)) - eps))%R = (- (x + eps))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 <= eps < ulp (pred (- x)))%R
Zx:(x < 0)%R

(- succ x + (ulp (- succ x) - eps))%R = (- (x + eps))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 <= eps < ulp (pred (- x)))%R
Zx:(x < 0)%R

(- - pred_pos (- x) + (ulp (- - pred_pos (- x)) - eps))%R = (- (x + eps))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 <= eps < ulp (pred (- x)))%R
Zx:(x < 0)%R

(pred_pos (- x) + (ulp (pred_pos (- x)) + - eps))%R = (- (x + eps))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 <= eps < ulp (pred (- x)))%R
Zx:(x < 0)%R

(- x + - eps)%R = (- (x + eps))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 <= eps < ulp (pred (- x)))%R
Zx:(x < 0)%R
(0 < - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 <= eps < ulp (pred (- x)))%R
Zx:(x < 0)%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 <= eps < ulp (pred (- x)))%R
Zx:(x < 0)%R

(0 < - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 <= eps < ulp (pred (- x)))%R
Zx:(x < 0)%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 <= eps < ulp (pred (- x)))%R
Zx:(x < 0)%R

F (- x)
now apply generic_format_opp. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, F x -> forall eps : R, (0 < eps <= (if Rle_bool 0 x then ulp x else ulp (pred (- x))))%R -> round beta fexp Zceil (x + eps) = succ x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, F x -> forall eps : R, (0 < eps <= (if Rle_bool 0 x then ulp x else ulp (pred (- x))))%R -> round beta fexp Zceil (x + eps) = succ x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool 0 x then ulp x else ulp (pred (- x))))%R

round beta fexp Zceil (x + eps) = succ x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool 0 x then ulp x else ulp (pred (- x))))%R
Zx:(0 <= x)%R

round beta fexp Zceil (x + eps) = succ x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool 0 x then ulp x else ulp (pred (- x))))%R
Zx:(x < 0)%R
round beta fexp Zceil (x + eps) = succ x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool 0 x then ulp x else ulp (pred (- x))))%R
Zx:(0 <= x)%R

round beta fexp Zceil (x + eps) = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool 0 x then ulp x else ulp (pred (- x))))%R
Zx:(x < 0)%R
round beta fexp Zceil (x + eps) = succ x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp x)%R
Zx:(0 <= x)%R

round beta fexp Zceil (x + eps) = (x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool 0 x then ulp x else ulp (pred (- x))))%R
Zx:(x < 0)%R
round beta fexp Zceil (x + eps) = succ x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool 0 x then ulp x else ulp (pred (- x))))%R
Zx:(x < 0)%R

round beta fexp Zceil (x + eps) = succ x
(* *)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R

round beta fexp Zceil (x + eps) = succ x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R

round beta fexp Zceil (- - (x + eps)) = succ x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R

round beta fexp Zceil (- - (x + eps)) = (- - succ x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R

(- round beta fexp Zfloor (- (x + eps)))%R = (- - succ x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R

round beta fexp Zfloor (- (x + eps)) = (- succ x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R

round beta fexp Zfloor (- succ x + (- eps + ulp (pred (- x)))) = (- succ x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R
(- succ x + (- eps + ulp (pred (- x))))%R = (- (x + eps))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R

(0 <= - succ x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R
F (- succ x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R
(0 <= - eps + ulp (pred (- x)) < ulp (- succ x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R
(- succ x + (- eps + ulp (pred (- x))))%R = (- (x + eps))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R

(0 <= pred (- x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R
F (- succ x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R
(0 <= - eps + ulp (pred (- x)) < ulp (- succ x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R
(- succ x + (- eps + ulp (pred (- x))))%R = (- (x + eps))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R

(0 < - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R
F (- succ x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R
(0 <= - eps + ulp (pred (- x)) < ulp (- succ x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R
(- succ x + (- eps + ulp (pred (- x))))%R = (- (x + eps))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R

F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R
F (- succ x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R
(0 <= - eps + ulp (pred (- x)) < ulp (- succ x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R
(- succ x + (- eps + ulp (pred (- x))))%R = (- (x + eps))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R

F (- succ x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R
(0 <= - eps + ulp (pred (- x)) < ulp (- succ x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R
(- succ x + (- eps + ulp (pred (- x))))%R = (- (x + eps))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R

(0 <= - eps + ulp (pred (- x)) < ulp (- succ x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R
(- succ x + (- eps + ulp (pred (- x))))%R = (- (x + eps))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R

(0 <= - eps + ulp (pred (- x)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R
(- eps + ulp (pred (- x)) < ulp (- succ x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R
(- succ x + (- eps + ulp (pred (- x))))%R = (- (x + eps))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R

(eps <= ulp (pred (- x)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R
(- eps + ulp (pred (- x)) < ulp (- succ x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R
(- succ x + (- eps + ulp (pred (- x))))%R = (- (x + eps))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R

(- eps + ulp (pred (- x)) < ulp (- succ x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R
(- succ x + (- eps + ulp (pred (- x))))%R = (- (x + eps))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R

(- eps + ulp (- succ x) < ulp (- succ x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R
(- succ x + (- eps + ulp (pred (- x))))%R = (- (x + eps))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R

(0 < eps)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R
(- succ x + (- eps + ulp (pred (- x))))%R = (- (x + eps))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R

(- succ x + (- eps + ulp (pred (- x))))%R = (- (x + eps))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R

(- - pred_pos (- x) + (- eps + ulp (pred (- x))))%R = (- (x + eps))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R

(- - pred_pos (- x) + (- eps + ulp (pred (- x))))%R = (- x + - eps)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R

(- - pred_pos (- x) + (- eps + ulp (pred (- x))))%R = (pred_pos (- x) + ulp (pred_pos (- x)) + - eps)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R
(0 < - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R

(- - pred_pos (- x) + (- eps + ulp (pred_pos (- x))))%R = (pred_pos (- x) + ulp (pred_pos (- x)) + - eps)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R
(0 <= - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R
(0 < - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R

(0 <= - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R
(0 < - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R

(0 < - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= ulp (pred (- x)))%R
Zx:(x < 0)%R

F (- x)
now apply generic_format_opp. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x y : R, F x -> F y -> (0 <= x < y)%R -> (x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x y : R, F x -> F y -> (0 <= x < y)%R -> (x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R

(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R

(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R

(0 < y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R

(x < y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R

(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
(* *)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R

(0 < pred y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R

(0 <= pred y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 <= pred y)%R
(0 < pred y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 <= pred y)%R

(0 < pred y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y

(0 < pred y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y

0%R = pred y -> (0 < pred y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y

0%R = pred_pos y -> (0 < pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y

0%R = (if Req_bool y (bpow (mag beta y - 1)) then (y - bpow (fexp (mag beta y - 1)))%R else (y - ulp y)%R) -> (0 < (if Req_bool y (bpow (mag beta y - 1)) then y - bpow (fexp (mag beta y - 1)) else y - ulp y))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R

0%R = (if Req_bool y (bpow (ey - 1)) then (y - bpow (fexp (ey - 1)))%R else (y - ulp y)%R) -> (0 < (if Req_bool y (bpow (ey - 1)) then y - bpow (fexp (ey - 1)) else y - ulp y))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)

0%R = (y - bpow (fexp (ey - 1)))%R -> (0 < y - bpow (fexp (ey - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
0%R = (y - ulp y)%R -> (0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
(* . *)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R

(0 < y - bpow (fexp (ey - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
0%R = (y - ulp y)%R -> (0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R

(ey - 1)%Z = fexp (ey - 1)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
(0 < y - bpow (fexp (ey - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
0%R = (y - ulp y)%R -> (0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R

bpow (ey - 1) = bpow (fexp (ey - 1))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
(0 < y - bpow (fexp (ey - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
0%R = (y - ulp y)%R -> (0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R

y = (y - bpow (fexp (ey - 1)) + bpow (fexp (ey - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
(0 < y - bpow (fexp (ey - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
0%R = (y - ulp y)%R -> (0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)

(0 < y - bpow (fexp (ey - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
0%R = (y - ulp y)%R -> (0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)

x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
(0 < y - bpow (fexp (ey - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
0%R = (y - ulp y)%R -> (0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R

(0 < y - bpow (fexp (ey - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
0%R = (y - ulp y)%R -> (0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R

(0 < y - bpow (fexp (ey - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
0%R = (y - ulp y)%R -> (0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R

(0 < y - bpow (fexp (ey - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
0%R = (y - ulp y)%R -> (0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R

(ex <= ey)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
(0 < y - bpow (fexp (ey - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
0%R = (y - ulp y)%R -> (0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R

(bpow (ex - 1) < bpow ey)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
(0 < y - bpow (fexp (ey - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
0%R = (y - ulp y)%R -> (0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R

(Rabs x < bpow ey)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
(0 < y - bpow (fexp (ey - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
0%R = (y - ulp y)%R -> (0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R

(Rabs x < Rabs y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(Rabs y < bpow ey)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
(0 < y - bpow (fexp (ey - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
0%R = (y - ulp y)%R -> (0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R

(x < y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(y >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(x >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(Rabs y < bpow ey)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
(0 < y - bpow (fexp (ey - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
0%R = (y - ulp y)%R -> (0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R

(y >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(x >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(Rabs y < bpow ey)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
(0 < y - bpow (fexp (ey - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
0%R = (y - ulp y)%R -> (0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R

(x >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(Rabs y < bpow ey)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
(0 < y - bpow (fexp (ey - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
0%R = (y - ulp y)%R -> (0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R

(Rabs y < bpow ey)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
(0 < y - bpow (fexp (ey - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
0%R = (y - ulp y)%R -> (0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R

y <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
(0 < y - bpow (fexp (ey - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
0%R = (y - ulp y)%R -> (0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z

(0 < y - bpow (fexp (ey - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
0%R = (y - ulp y)%R -> (0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:(ex < ey)%Z

(0 < y - bpow (fexp (ey - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:ex = ey
(0 < y - bpow (fexp (ey - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
0%R = (y - ulp y)%R -> (0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:(ex < ey)%Z

fexp ex = fexp (ey - 1)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:(ex < ey)%Z
H3:fexp ex = fexp (ey - 1)
(0 < y - bpow (fexp (ey - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:ex = ey
(0 < y - bpow (fexp (ey - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
0%R = (y - ulp y)%R -> (0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:(ex < ey)%Z

(ey - 1 <= fexp (ey - 1))%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:(ex < ey)%Z
(ex <= fexp (ey - 1))%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:(ex < ey)%Z
H3:fexp ex = fexp (ey - 1)
(0 < y - bpow (fexp (ey - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:ex = ey
(0 < y - bpow (fexp (ey - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
0%R = (y - ulp y)%R -> (0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:(ex < ey)%Z

(ex <= fexp (ey - 1))%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:(ex < ey)%Z
H3:fexp ex = fexp (ey - 1)
(0 < y - bpow (fexp (ey - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:ex = ey
(0 < y - bpow (fexp (ey - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
0%R = (y - ulp y)%R -> (0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:(ex < ey)%Z

(ex <= ey - 1)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:(ex < ey)%Z
H3:fexp ex = fexp (ey - 1)
(0 < y - bpow (fexp (ey - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:ex = ey
(0 < y - bpow (fexp (ey - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
0%R = (y - ulp y)%R -> (0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:(ex < ey)%Z
H3:fexp ex = fexp (ey - 1)

(0 < y - bpow (fexp (ey - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:ex = ey
(0 < y - bpow (fexp (ey - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
0%R = (y - ulp y)%R -> (0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:(ex < ey)%Z
H3:fexp ex = fexp (ey - 1)

~ (0 < Ztrunc (scaled_mantissa beta fexp x) < 1)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:(ex < ey)%Z
H3:fexp ex = fexp (ey - 1)
(0 < Ztrunc (scaled_mantissa beta fexp x) < 1)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:ex = ey
(0 < y - bpow (fexp (ey - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
0%R = (y - ulp y)%R -> (0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:(ex < ey)%Z
H3:fexp ex = fexp (ey - 1)

(0 < Ztrunc (scaled_mantissa beta fexp x) < 1)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:ex = ey
(0 < y - bpow (fexp (ey - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
0%R = (y - ulp y)%R -> (0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:(ex < ey)%Z
H3:fexp ex = fexp (ey - 1)

(0 < Ztrunc (scaled_mantissa beta fexp x))%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:(ex < ey)%Z
H3:fexp ex = fexp (ey - 1)
(Ztrunc (scaled_mantissa beta fexp x) < 1)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:ex = ey
(0 < y - bpow (fexp (ey - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
0%R = (y - ulp y)%R -> (0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:(ex < ey)%Z
H3:fexp ex = fexp (ey - 1)

(0 < F2R {| Fnum := Ztrunc (scaled_mantissa beta fexp x); Fexp := cexp beta fexp x |})%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:(ex < ey)%Z
H3:fexp ex = fexp (ey - 1)
(Ztrunc (scaled_mantissa beta fexp x) < 1)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:ex = ey
(0 < y - bpow (fexp (ey - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
0%R = (y - ulp y)%R -> (0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:(ex < ey)%Z
H3:fexp ex = fexp (ey - 1)

(Ztrunc (scaled_mantissa beta fexp x) < 1)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:ex = ey
(0 < y - bpow (fexp (ey - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
0%R = (y - ulp y)%R -> (0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:(ex < ey)%Z
H3:fexp ex = fexp (ey - 1)

(IZR (Ztrunc (scaled_mantissa beta fexp x)) < 1)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:ex = ey
(0 < y - bpow (fexp (ey - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
0%R = (y - ulp y)%R -> (0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:(ex < ey)%Z
H3:fexp ex = fexp (ey - 1)

(0 < bpow (cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:(ex < ey)%Z
H3:fexp ex = fexp (ey - 1)
(IZR (Ztrunc (scaled_mantissa beta fexp x)) * bpow (cexp beta fexp x) < 1 * bpow (cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:ex = ey
(0 < y - bpow (fexp (ey - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
0%R = (y - ulp y)%R -> (0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:(ex < ey)%Z
H3:fexp ex = fexp (ey - 1)

(IZR (Ztrunc (scaled_mantissa beta fexp x)) * bpow (cexp beta fexp x) < 1 * bpow (cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:ex = ey
(0 < y - bpow (fexp (ey - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
0%R = (y - ulp y)%R -> (0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:(ex < ey)%Z
H3:fexp ex = fexp (ey - 1)

(x < 1 * bpow (cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:ex = ey
(0 < y - bpow (fexp (ey - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
0%R = (y - ulp y)%R -> (0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:(ex < ey)%Z
H3:fexp ex = fexp (ey - 1)

(x < bpow (cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:ex = ey
(0 < y - bpow (fexp (ey - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
0%R = (y - ulp y)%R -> (0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:(ex < ey)%Z
H3:fexp ex = fexp (ey - 1)

(x < bpow (fexp (mag beta x)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:ex = ey
(0 < y - bpow (fexp (ey - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
0%R = (y - ulp y)%R -> (0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:(ex < ey)%Z
H3:fexp ex = fexp (ey - 1)

(x < bpow (fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:(ex < ey)%Z
H3:fexp ex = fexp (ey - 1)
(bpow (ex - 1) <= Rabs x < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:ex = ey
(0 < y - bpow (fexp (ey - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
0%R = (y - ulp y)%R -> (0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:(ex < ey)%Z
H3:fexp ex = fexp (ey - 1)

(x < y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:(ex < ey)%Z
H3:fexp ex = fexp (ey - 1)
(bpow (ex - 1) <= Rabs x < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:ex = ey
(0 < y - bpow (fexp (ey - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
0%R = (y - ulp y)%R -> (0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:(ex < ey)%Z
H3:fexp ex = fexp (ey - 1)

(bpow (ex - 1) <= Rabs x < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:ex = ey
(0 < y - bpow (fexp (ey - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
0%R = (y - ulp y)%R -> (0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:ex = ey

(0 < y - bpow (fexp (ey - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
0%R = (y - ulp y)%R -> (0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:ex = ey

~ (y <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:ex = ey
(y <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
0%R = (y - ulp y)%R -> (0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:ex = ey

(y <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
0%R = (y - ulp y)%R -> (0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:ex = ey

(y <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:ex = ey
(x >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
0%R = (y - ulp y)%R -> (0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:ex = ey

(y <= bpow (ex - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:ex = ey
(x >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
0%R = (y - ulp y)%R -> (0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:ex = ey

(bpow (ey - 1) <= bpow (ey - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:ex = ey
(x >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
0%R = (y - ulp y)%R -> (0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y = bpow (ey - 1)
Hy3:0%R = (y - bpow (fexp (ey - 1)))%R
H1:(ey - 1)%Z = fexp (ey - 1)
Zx:x <> 0%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:(ex <= ey)%Z
Hexy:ex = ey

(x >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
0%R = (y - ulp y)%R -> (0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)

0%R = (y - ulp y)%R -> (0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
(* . *)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
Hy3:0%R = (y - ulp y)%R

(0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
Hy3:0%R = (y - ulp y)%R

y = bpow (fexp ey)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
Hy3:0%R = (y - ulp y)%R
H1:y = bpow (fexp ey)
(0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
Hy3:0%R = (y - ulp y)%R

(y - bpow (fexp ey))%R = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
Hy3:0%R = (y - ulp y)%R
H1:y = bpow (fexp ey)
(0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
Hy3:0%R = (y - ulp y)%R

(y - bpow (fexp ey))%R = (y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
Hy3:0%R = (y - ulp y)%R
H1:y = bpow (fexp ey)
(0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
Hy3:0%R = (y - ulp y)%R

(y - bpow (fexp ey))%R = (y - bpow (cexp beta fexp y))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
Hy3:0%R = (y - ulp y)%R
H1:y = bpow (fexp ey)
(0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
Hy3:0%R = (y - ulp y)%R

(y - bpow (fexp ey))%R = (y - bpow (fexp (mag beta y)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
Hy3:0%R = (y - ulp y)%R
H1:y = bpow (fexp ey)
(0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
Hy3:0%R = (y - ulp y)%R

(bpow (ey - 1) <= Rabs y < bpow ey)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
Hy3:0%R = (y - ulp y)%R
H1:y = bpow (fexp ey)
(0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
Hy3:0%R = (y - ulp y)%R

y <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
Hy3:0%R = (y - ulp y)%R
H1:y = bpow (fexp ey)
(0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy2:y <> bpow (ey - 1)
Hy3:0%R = (y - ulp y)%R
H1:y = bpow (fexp ey)

(0 < y - ulp y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy3:0%R = (y - ulp y)%R
H1:y = bpow (fexp ey)

y = bpow (ey - 1)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy3:0%R = (y - ulp y)%R
H1:y = bpow (fexp ey)

bpow (fexp ey) = bpow (ey - 1)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy3:0%R = (y - ulp y)%R
H1:y = bpow (fexp ey)

fexp ey = (ey - 1)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy3:0%R = (y - ulp y)%R
H1:y = bpow (fexp ey)

(1 + fexp ey)%Z = (1 + (ey - 1))%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy3:0%R = (y - ulp y)%R
H1:y = bpow (fexp ey)

(fexp ey + 1)%Z = ey
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy3:0%R = (y - ulp y)%R
H1:y = bpow (fexp ey)

(fexp ey + 1)%Z = mag beta y
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy3:0%R = (y - ulp y)%R
H1:y = bpow (fexp ey)
mag beta y = ey
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy3:0%R = (y - ulp y)%R
H1:y = bpow (fexp ey)

(bpow (fexp ey + 1 - 1) <= Rabs y < bpow (fexp ey + 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy3:0%R = (y - ulp y)%R
H1:y = bpow (fexp ey)
mag beta y = ey
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy3:0%R = (y - ulp y)%R
H1:y = bpow (fexp ey)

(bpow (fexp ey + 1 - 1) <= bpow (fexp ey) < bpow (fexp ey + 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy3:0%R = (y - ulp y)%R
H1:y = bpow (fexp ey)
(bpow (fexp ey) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy3:0%R = (y - ulp y)%R
H1:y = bpow (fexp ey)
mag beta y = ey
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy3:0%R = (y - ulp y)%R
H1:y = bpow (fexp ey)

(bpow (fexp ey + 1 - 1) <= bpow (fexp ey))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy3:0%R = (y - ulp y)%R
H1:y = bpow (fexp ey)
(bpow (fexp ey) < bpow (fexp ey + 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy3:0%R = (y - ulp y)%R
H1:y = bpow (fexp ey)
(bpow (fexp ey) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy3:0%R = (y - ulp y)%R
H1:y = bpow (fexp ey)
mag beta y = ey
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy3:0%R = (y - ulp y)%R
H1:y = bpow (fexp ey)

(fexp ey + 1 - 1 <= fexp ey)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy3:0%R = (y - ulp y)%R
H1:y = bpow (fexp ey)
(bpow (fexp ey) < bpow (fexp ey + 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy3:0%R = (y - ulp y)%R
H1:y = bpow (fexp ey)
(bpow (fexp ey) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy3:0%R = (y - ulp y)%R
H1:y = bpow (fexp ey)
mag beta y = ey
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy3:0%R = (y - ulp y)%R
H1:y = bpow (fexp ey)

(bpow (fexp ey) < bpow (fexp ey + 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy3:0%R = (y - ulp y)%R
H1:y = bpow (fexp ey)
(bpow (fexp ey) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy3:0%R = (y - ulp y)%R
H1:y = bpow (fexp ey)
mag beta y = ey
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy3:0%R = (y - ulp y)%R
H1:y = bpow (fexp ey)

(fexp ey < fexp ey + 1)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy3:0%R = (y - ulp y)%R
H1:y = bpow (fexp ey)
(bpow (fexp ey) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy3:0%R = (y - ulp y)%R
H1:y = bpow (fexp ey)
mag beta y = ey
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy3:0%R = (y - ulp y)%R
H1:y = bpow (fexp ey)

(bpow (fexp ey) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy3:0%R = (y - ulp y)%R
H1:y = bpow (fexp ey)
mag beta y = ey
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy3:0%R = (y - ulp y)%R
H1:y = bpow (fexp ey)

mag beta y = ey
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy3:0%R = (y - ulp y)%R
H1:y = bpow (fexp ey)

(bpow (ey - 1) <= Rabs y < bpow ey)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
H0:0%R = pred y
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy3:0%R = (y - ulp y)%R
H1:y = bpow (fexp ey)

y <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R

(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
(* *)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
H1:(ulp (pred_pos y) <= y - x)%R

(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
H1:(y - x < ulp (pred_pos y))%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
(* . *)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
H1:(ulp (pred_pos y) <= y - x)%R

(x + (- x + ulp (pred_pos y)) <= pred_pos y + (- x + ulp (pred_pos y)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
H1:(y - x < ulp (pred_pos y))%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
H1:(ulp (pred_pos y) <= y - x)%R

(ulp (pred_pos y) <= pred_pos y + (- x + ulp (pred_pos y)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
H1:(y - x < ulp (pred_pos y))%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
H1:(ulp (pred_pos y) <= y - x)%R

(y - x <= pred_pos y + (- x + ulp (pred_pos y)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
H1:(y - x < ulp (pred_pos y))%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
H1:(ulp (pred_pos y) <= y - x)%R

(pred_pos y + ulp (pred_pos y) - x <= pred_pos y + (- x + ulp (pred_pos y)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
H1:(y - x < ulp (pred_pos y))%R
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
H1:(y - x < ulp (pred_pos y))%R

(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
(* . *)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
H1:(y - x < ulp (pred_pos y))%R

(y - (y - x) <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
H1:(y - x < ulp (pred_pos y))%R

(y - (y - x) <= pred y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
H1:(y - x < ulp (pred_pos y))%R

(y - (y - x) <= round beta fexp Zfloor (y - (y - x)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
H1:(y - x < ulp (pred_pos y))%R
(0 < y - x <= ulp (pred y))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
H1:(y - x < ulp (pred_pos y))%R

(x <= round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
H1:(y - x < ulp (pred_pos y))%R
(0 < y - x <= ulp (pred y))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
H1:(y - x < ulp (pred_pos y))%R

x = round beta fexp Zfloor x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
H1:(y - x < ulp (pred_pos y))%R
(0 < y - x <= ulp (pred y))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
H1:(y - x < ulp (pred_pos y))%R

round beta fexp Zfloor x = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
H1:(y - x < ulp (pred_pos y))%R
(0 < y - x <= ulp (pred y))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
H1:(y - x < ulp (pred_pos y))%R

(0 < y - x <= ulp (pred y))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
H1:(y - x < ulp (pred_pos y))%R

(0 < y - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
H1:(y - x < ulp (pred_pos y))%R
(y - x <= ulp (pred y))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
H1:(y - x < ulp (pred_pos y))%R

(y - x <= ulp (pred y))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:(0 < x)%R
Zy:(0 < y)%R
Zp:(0 < pred y)%R
H1:(y - x < ulp (pred_pos y))%R

(y - x <= ulp (pred_pos y))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x
(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x

(x <= pred_pos y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(0 <= x < y)%R
V:0%R = x

(0 < y)%R
apply Rle_lt_trans with (1:=proj1 H); apply H. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x y : R, F x -> F y -> (0 <= x)%R -> (x < y)%R -> (succ x <= y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x y : R, F x -> F y -> (0 <= x)%R -> (x < y)%R -> (succ x <= y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Hx:F x
Hy:F y
Zx:(0 <= x)%R
H:(x < y)%R

(succ x <= y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Hx:F x
Hy:F y
Zx:(0 <= x)%R
H:(x < y)%R

(x + ulp x <= y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Hx:F x
Hy:F y
Zx:(0 <= x)%R
H:(x < y)%R
H1:(ulp x <= y - x)%R

(x + ulp x <= y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Hx:F x
Hy:F y
Zx:(0 <= x)%R
H:(x < y)%R
H1:(y - x < ulp x)%R
(x + ulp x <= y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Hx:F x
Hy:F y
Zx:(0 <= x)%R
H:(x < y)%R
H1:(ulp x <= y - x)%R

(x + ulp x + - x <= y + - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Hx:F x
Hy:F y
Zx:(0 <= x)%R
H:(x < y)%R
H1:(y - x < ulp x)%R
(x + ulp x <= y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Hx:F x
Hy:F y
Zx:(0 <= x)%R
H:(x < y)%R
H1:(y - x < ulp x)%R

(x + ulp x <= y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Hx:F x
Hy:F y
Zx:(0 <= x)%R
H:(x < y)%R
H1:(y - x < ulp x)%R

(x + ulp x <= x + (y - x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Hx:F x
Hy:F y
Zx:(0 <= x)%R
H:(x < y)%R
H1:(y - x < ulp x)%R

~ (x < y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Hx:F x
Hy:F y
Zx:(0 <= x)%R
H:(x < y)%R
H1:(y - x < ulp x)%R
(x < y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Hx:F x
Hy:F y
Zx:(0 <= x)%R
H:(x < y)%R
H1:(y - x < ulp x)%R

~ (x < y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Hx:F x
Hy:F y
Zx:(0 <= x)%R
H:(x < y)%R
H1:(y - x < ulp x)%R

y = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Hx:F x
Hy:F y
Zx:(0 <= x)%R
H:(x < y)%R
H1:(y - x < ulp x)%R

y = round beta fexp Zfloor (x + (y - x))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Hx:F x
Hy:F y
Zx:(0 <= x)%R
H:(x < y)%R
H1:(y - x < ulp x)%R
(0 <= y - x < ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Hx:F x
Hy:F y
Zx:(0 <= x)%R
H:(x < y)%R
H1:(y - x < ulp x)%R

y = round beta fexp Zfloor y
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Hx:F x
Hy:F y
Zx:(0 <= x)%R
H:(x < y)%R
H1:(y - x < ulp x)%R
(0 <= y - x < ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Hx:F x
Hy:F y
Zx:(0 <= x)%R
H:(x < y)%R
H1:(y - x < ulp x)%R

round beta fexp Zfloor y = y
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Hx:F x
Hy:F y
Zx:(0 <= x)%R
H:(x < y)%R
H1:(y - x < ulp x)%R
(0 <= y - x < ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Hx:F x
Hy:F y
Zx:(0 <= x)%R
H:(x < y)%R
H1:(y - x < ulp x)%R

(0 <= y - x < ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Hx:F x
Hy:F y
Zx:(0 <= x)%R
H:(x < y)%R
H1:(y - x < ulp x)%R

(0 <= y - x)%R
apply Rlt_le; now apply Rlt_Rminus. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x y : R, F x -> F y -> (x < y)%R -> (succ x <= y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x y : R, F x -> F y -> (x < y)%R -> (succ x <= y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(x < y)%R

(succ x <= y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(x < y)%R
Hx:(0 <= x)%R

(succ x <= y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(x < y)%R
Hx:(x < 0)%R
(succ x <= y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(x < y)%R
Hx:(x < 0)%R

(succ x <= y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(x < y)%R
Hx:(x < 0)%R

(- pred_pos (- x) <= y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(x < y)%R
Hx:(x < 0)%R
Hy:(y <= 0)%R

(- pred_pos (- x) <= y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(x < y)%R
Hx:(x < 0)%R
Hy:(0 < y)%R
(- pred_pos (- x) <= y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(x < y)%R
Hx:(x < 0)%R
Hy:(y <= 0)%R

(- pred_pos (- x) <= - - y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(x < y)%R
Hx:(x < 0)%R
Hy:(0 < y)%R
(- pred_pos (- x) <= y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(x < y)%R
Hx:(x < 0)%R
Hy:(y <= 0)%R

(- y <= pred_pos (- x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(x < y)%R
Hx:(x < 0)%R
Hy:(0 < y)%R
(- pred_pos (- x) <= y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(x < y)%R
Hx:(x < 0)%R
Hy:(y <= 0)%R

F (- y)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(x < y)%R
Hx:(x < 0)%R
Hy:(y <= 0)%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(x < y)%R
Hx:(x < 0)%R
Hy:(y <= 0)%R
(0 <= - y < - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(x < y)%R
Hx:(x < 0)%R
Hy:(0 < y)%R
(- pred_pos (- x) <= y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(x < y)%R
Hx:(x < 0)%R
Hy:(y <= 0)%R

F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(x < y)%R
Hx:(x < 0)%R
Hy:(y <= 0)%R
(0 <= - y < - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(x < y)%R
Hx:(x < 0)%R
Hy:(0 < y)%R
(- pred_pos (- x) <= y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(x < y)%R
Hx:(x < 0)%R
Hy:(y <= 0)%R

(0 <= - y < - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(x < y)%R
Hx:(x < 0)%R
Hy:(0 < y)%R
(- pred_pos (- x) <= y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(x < y)%R
Hx:(x < 0)%R
Hy:(y <= 0)%R

(0 <= - y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(x < y)%R
Hx:(x < 0)%R
Hy:(y <= 0)%R
(- y < - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(x < y)%R
Hx:(x < 0)%R
Hy:(0 < y)%R
(- pred_pos (- x) <= y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(x < y)%R
Hx:(x < 0)%R
Hy:(y <= 0)%R

(- y < - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(x < y)%R
Hx:(x < 0)%R
Hy:(0 < y)%R
(- pred_pos (- x) <= y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(x < y)%R
Hx:(x < 0)%R
Hy:(0 < y)%R

(- pred_pos (- x) <= y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(x < y)%R
Hx:(x < 0)%R
Hy:(0 < y)%R

(- pred_pos (- x) <= - 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(x < y)%R
Hx:(x < 0)%R
Hy:(0 < y)%R
(- 0 <= y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(x < y)%R
Hx:(x < 0)%R
Hy:(0 < y)%R

(0 <= pred_pos (- x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(x < y)%R
Hx:(x < 0)%R
Hy:(0 < y)%R
(- 0 <= y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(x < y)%R
Hx:(x < 0)%R
Hy:(0 < y)%R

(0 < - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(x < y)%R
Hx:(x < 0)%R
Hy:(0 < y)%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(x < y)%R
Hx:(x < 0)%R
Hy:(0 < y)%R
(- 0 <= y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(x < y)%R
Hx:(x < 0)%R
Hy:(0 < y)%R

F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(x < y)%R
Hx:(x < 0)%R
Hy:(0 < y)%R
(- 0 <= y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
H:(x < y)%R
Hx:(x < 0)%R
Hy:(0 < y)%R

(- 0 <= y)%R
rewrite Ropp_0; now left. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x y : R, F x -> F y -> (x < y)%R -> (x <= pred y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x y : R, F x -> F y -> (x < y)%R -> (x <= pred y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
Hxy:(x < y)%R

(x <= pred y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
Hxy:(x < y)%R

(- - x <= pred y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
Hxy:(x < y)%R

(succ (- y) <= - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
Hxy:(x < y)%R

F (- y)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
Hxy:(x < y)%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
Hxy:(x < y)%R
(- y < - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
Hxy:(x < y)%R

F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
Hxy:(x < y)%R
(- y < - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
Hxy:(x < y)%R

(- y < - x)%R
now apply Ropp_lt_contravar. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x y : R, y <> 0%R -> (x <= y)%R -> (x < succ y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x y : R, y <> 0%R -> (x <= y)%R -> (x < succ y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Zy:y <> 0%R
Hxy:(x <= y)%R

(x < succ y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Zy:y <> 0%R
Hxy:(x <= y)%R

(y < succ y)%R
now apply succ_gt_id. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x y : R, x <> 0%R -> (x <= y)%R -> (pred x < y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x y : R, x <> 0%R -> (x <= y)%R -> (pred x < y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Zy:x <> 0%R
Hxy:(x <= y)%R

(pred x < y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Zy:x <> 0%R
Hxy:(x <= y)%R

(pred x < x)%R
now apply pred_lt_id. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, F x -> (0 < x)%R -> succ (pred x) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, F x -> (0 < x)%R -> succ (pred x) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R

succ (pred x) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R

succ (pred_pos x) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R

(pred_pos x + ulp (pred_pos x))%R = x
now apply pred_pos_plus_ulp. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

pred (ulp 0) = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

pred (ulp 0) = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

pred_pos (ulp 0) = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
(0 <= ulp 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

pred_pos (ulp 0) = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

pred_pos match negligible_exp with | Some n => bpow (fexp n) | None => 0 end = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

negligible_exp = None /\ (forall n : Z, (fexp n < n)%Z) -> pred_pos match negligible_exp with | Some n => bpow (fexp n) | None => 0 end = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
(exists n : Z, negligible_exp = Some n /\ (n <= fexp n)%Z) -> pred_pos match negligible_exp with | Some n => bpow (fexp n) | None => 0 end = 0%R
(* *)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H1:negligible_exp = None

pred_pos 0 = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
(exists n : Z, negligible_exp = Some n /\ (n <= fexp n)%Z) -> pred_pos match negligible_exp with | Some n => bpow (fexp n) | None => 0 end = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H1:negligible_exp = None

(0 - ulp 0)%R = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H1:negligible_exp = None
0%R <> bpow (mag beta 0 - 1)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
(exists n : Z, negligible_exp = Some n /\ (n <= fexp n)%Z) -> pred_pos match negligible_exp with | Some n => bpow (fexp n) | None => 0 end = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H1:negligible_exp = None

(0 - ulp 0)%R = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
(exists n : Z, negligible_exp = Some n /\ (n <= fexp n)%Z) -> pred_pos match negligible_exp with | Some n => bpow (fexp n) | None => 0 end = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H1:negligible_exp = None

(0 - match negligible_exp with | Some n => bpow (fexp n) | None => 0 end)%R = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
(exists n : Z, negligible_exp = Some n /\ (n <= fexp n)%Z) -> pred_pos match negligible_exp with | Some n => bpow (fexp n) | None => 0 end = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

(exists n : Z, negligible_exp = Some n /\ (n <= fexp n)%Z) -> pred_pos match negligible_exp with | Some n => bpow (fexp n) | None => 0 end = 0%R
(* *)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
n:Z
H1:negligible_exp = Some n
H2:(n <= fexp n)%Z

pred_pos (bpow (fexp n)) = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
n:Z
H1:negligible_exp = Some n
H2:(n <= fexp n)%Z

(if Req_bool (bpow (fexp n)) (bpow (mag beta (bpow (fexp n)) - 1)) then (bpow (fexp n) - bpow (fexp (mag beta (bpow (fexp n)) - 1)))%R else (bpow (fexp n) - ulp (bpow (fexp n)))%R) = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
n:Z
H1:negligible_exp = Some n
H2:(n <= fexp n)%Z

(if Req_bool (bpow (fexp n)) (bpow (fexp n + 1 - 1)) then (bpow (fexp n) - bpow (fexp (fexp n + 1 - 1)))%R else (bpow (fexp n) - ulp (bpow (fexp n)))%R) = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
n:Z
H1:negligible_exp = Some n
H2:(n <= fexp n)%Z

(if Req_bool (bpow (fexp n)) (bpow (fexp n)) then (bpow (fexp n) - bpow (fexp (fexp n)))%R else (bpow (fexp n) - ulp (bpow (fexp n)))%R) = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
n:Z
H1:negligible_exp = Some n
H2:(n <= fexp n)%Z

(bpow (fexp n) - bpow (fexp (fexp n)))%R = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
n:Z
H1:negligible_exp = Some n
H2:(n <= fexp n)%Z

fexp n = fexp (fexp n)
apply sym_eq, valid_exp; omega. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

succ 0 = ulp 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

succ 0 = ulp 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

(if Rle_bool 0 0 then (0 + ulp 0)%R else (- pred_pos (- 0))%R) = ulp 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

(0 + ulp 0)%R = ulp 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
(0 <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

(0 <= 0)%R
apply Rle_refl. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

pred 0 = (- ulp 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

pred 0 = (- ulp 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

pred 0 = (- succ 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

pred (- 0) = (- succ 0)%R
apply pred_opp. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, F x -> (0 < x)%R -> pred (succ x) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, F x -> (0 < x)%R -> pred (succ x) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R

pred (succ x) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R

(pred (succ x) <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R
(x <= pred (succ x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R

(pred (succ x) <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R

~ (x < pred (succ x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R
H:(x < pred (succ x))%R

False
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R
H:(succ x <= pred (succ x))%R

False
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R
H:(x < pred (succ x))%R
F (pred (succ x))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R

(succ x <= pred (succ x))%R -> False
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R
H:(x < pred (succ x))%R
F (pred (succ x))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R

(pred (succ x) < succ x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R
H:(x < pred (succ x))%R
F (pred (succ x))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R

succ x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R
H:(x < pred (succ x))%R
F (pred (succ x))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R

(succ x > 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R
H:(x < pred (succ x))%R
F (pred (succ x))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R

(x <= succ x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R
H:(x < pred (succ x))%R
F (pred (succ x))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R
H:(x < pred (succ x))%R

F (pred (succ x))
now apply generic_format_pred, generic_format_succ.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R

(x <= pred (succ x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R

F (succ x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R
(x < succ x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R

(x < succ x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R

x <> 0%R
now apply Rgt_not_eq. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, F x -> succ (pred x) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, F x -> succ (pred x) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x

succ (pred x) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R

succ (pred x) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:0%R = x
succ (pred x) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(x < 0)%R
succ (pred x) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:0%R = x

succ (pred x) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(x < 0)%R
succ (pred x) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:0%R = x

succ (pred 0) = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(x < 0)%R
succ (pred x) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:0%R = x

(- 0)%R = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(x < 0)%R
succ (pred x) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(x < 0)%R

succ (pred x) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(x < 0)%R

succ (- succ (- x)) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(x < 0)%R

(- - x)%R = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(x < 0)%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(x < 0)%R
(0 < - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(x < 0)%R

F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(x < 0)%R
(0 < - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(x < 0)%R

(0 < - x)%R
now apply Ropp_0_gt_lt_contravar. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, F x -> pred (succ x) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, F x -> pred (succ x) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x

pred (succ x) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x

pred (succ (- - x)) = (- - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x

(- succ (pred (- x)))%R = (- - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x

F (- x)
now apply generic_format_opp. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, F x -> forall eps : R, (0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R -> round beta fexp Zceil (pred x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, F x -> forall eps : R, (0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R -> round beta fexp Zceil (pred x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R

round beta fexp Zceil (pred x + eps) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R

succ (pred x) = x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
F (pred x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
(0 < eps <= (if Rle_bool 0 (pred x) then ulp (pred x) else ulp (pred (- pred x))))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R

F (pred x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
(0 < eps <= (if Rle_bool 0 (pred x) then ulp (pred x) else ulp (pred (- pred x))))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R

(0 < eps <= (if Rle_bool 0 (pred x) then ulp (pred x) else ulp (pred (- pred x))))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R

(0 < eps <= (if Rle_bool 0 (pred x) then ulp (pred x) else ulp (pred (- - succ (- x)))))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R

(0 < eps <= (if Rle_bool 0 (pred x) then ulp (pred x) else ulp (- x)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R

(0 < eps <= (if Rle_bool 0 (pred x) then ulp (pred x) else ulp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(x <= 0)%R
H2:(0 < eps <= ulp x)%R

(0 < eps <= (if Rle_bool 0 (pred x) then ulp (pred x) else ulp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(0 < x)%R
H2:(0 < eps <= ulp (pred x))%R
(0 < eps <= (if Rle_bool 0 (pred x) then ulp (pred x) else ulp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(x <= 0)%R
H2:(0 < eps <= ulp x)%R

(pred x < 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(0 < x)%R
H2:(0 < eps <= ulp (pred x))%R
(0 < eps <= (if Rle_bool 0 (pred x) then ulp (pred x) else ulp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(x <= 0)%R
H2:(0 < eps <= ulp x)%R
H1':(x < 0)%R

(pred x < 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(x <= 0)%R
H2:(0 < eps <= ulp x)%R
H1':x = 0%R
(pred x < 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(0 < x)%R
H2:(0 < eps <= ulp (pred x))%R
(0 < eps <= (if Rle_bool 0 (pred x) then ulp (pred x) else ulp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(x <= 0)%R
H2:(0 < eps <= ulp x)%R
H1':(x < 0)%R

(pred x < x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(x <= 0)%R
H2:(0 < eps <= ulp x)%R
H1':x = 0%R
(pred x < 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(0 < x)%R
H2:(0 < eps <= ulp (pred x))%R
(0 < eps <= (if Rle_bool 0 (pred x) then ulp (pred x) else ulp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(x <= 0)%R
H2:(0 < eps <= ulp x)%R
H1':(x < 0)%R

x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(x <= 0)%R
H2:(0 < eps <= ulp x)%R
H1':x = 0%R
(pred x < 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(0 < x)%R
H2:(0 < eps <= ulp (pred x))%R
(0 < eps <= (if Rle_bool 0 (pred x) then ulp (pred x) else ulp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(x <= 0)%R
H2:(0 < eps <= ulp x)%R
H1':x = 0%R

(pred x < 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(0 < x)%R
H2:(0 < eps <= ulp (pred x))%R
(0 < eps <= (if Rle_bool 0 (pred x) then ulp (pred x) else ulp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(x <= 0)%R
H2:(0 < eps <= ulp x)%R
H1':x = 0%R

(- (if Rle_bool 0 (- 0) then - 0 + ulp (- 0) else - pred_pos (- - 0)) < 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(0 < x)%R
H2:(0 < eps <= ulp (pred x))%R
(0 < eps <= (if Rle_bool 0 (pred x) then ulp (pred x) else ulp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(x <= 0)%R
H2:(0 < eps <= ulp x)%R
H1':x = 0%R

(- (0 + ulp 0) < 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(0 < x)%R
H2:(0 < eps <= ulp (pred x))%R
(0 < eps <= (if Rle_bool 0 (pred x) then ulp (pred x) else ulp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(x <= 0)%R
H2:(0 < eps <= ulp x)%R
H1':x = 0%R

(- ulp 0 < 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(0 < x)%R
H2:(0 < eps <= ulp (pred x))%R
(0 < eps <= (if Rle_bool 0 (pred x) then ulp (pred x) else ulp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(x <= 0)%R
H2:(0 < eps <= ulp x)%R
H1':x = 0%R

(0 < ulp (- 0))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(0 < x)%R
H2:(0 < eps <= ulp (pred x))%R
(0 < eps <= (if Rle_bool 0 (pred x) then ulp (pred x) else ulp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(x <= 0)%R
H2:(0 < eps <= ulp x)%R
H1':x = 0%R

(eps <= ulp (- 0))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(0 < x)%R
H2:(0 < eps <= ulp (pred x))%R
(0 < eps <= (if Rle_bool 0 (pred x) then ulp (pred x) else ulp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(x <= 0)%R
H2:(0 < eps <= ulp x)%R
H1':x = 0%R

(ulp x <= ulp (- 0))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(0 < x)%R
H2:(0 < eps <= ulp (pred x))%R
(0 < eps <= (if Rle_bool 0 (pred x) then ulp (pred x) else ulp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(x <= 0)%R
H2:(0 < eps <= ulp x)%R
H1':x = 0%R

(ulp 0 <= ulp 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(0 < x)%R
H2:(0 < eps <= ulp (pred x))%R
(0 < eps <= (if Rle_bool 0 (pred x) then ulp (pred x) else ulp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(0 < x)%R
H2:(0 < eps <= ulp (pred x))%R

(0 < eps <= (if Rle_bool 0 (pred x) then ulp (pred x) else ulp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(0 < x)%R
H2:(0 < eps <= ulp (pred x))%R

(0 <= pred x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R

F (- x)
now apply generic_format_opp. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, F x -> forall eps : R, (0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R -> round beta fexp Zfloor (x - eps) = pred x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, F x -> forall eps : R, (0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R -> round beta fexp Zfloor (x - eps) = pred x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R

round beta fexp Zfloor (x - eps) = pred x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R

round beta fexp Zfloor (- (- x + eps)) = pred x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R

(- round beta fexp Zceil (- x + eps))%R = pred x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R

round beta fexp Zceil (- x + eps) = succ (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R

round beta fexp Zceil (pred (succ (- x)) + eps) = succ (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R

F (succ (- x))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
(0 < eps <= (if Rle_bool (succ (- x)) 0 then ulp (succ (- x)) else ulp (pred (succ (- x)))))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R

(0 < eps <= (if Rle_bool (succ (- x)) 0 then ulp (succ (- x)) else ulp (pred (succ (- x)))))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R

(0 < eps <= (if Rle_bool (succ (- x)) 0 then ulp (succ (- x)) else ulp (- x)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R

(0 < eps <= (if Rle_bool (succ (- x)) 0 then ulp (succ (- x)) else ulp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(x <= 0)%R
H2:(0 < eps <= ulp x)%R

(0 < eps <= (if Rle_bool (succ (- x)) 0 then ulp (succ (- x)) else ulp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(0 < x)%R
H2:(0 < eps <= ulp (pred x))%R
(0 < eps <= (if Rle_bool (succ (- x)) 0 then ulp (succ (- x)) else ulp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(x <= 0)%R
H2:(0 < eps <= ulp x)%R

(0 < succ (- x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(0 < x)%R
H2:(0 < eps <= ulp (pred x))%R
(0 < eps <= (if Rle_bool (succ (- x)) 0 then ulp (succ (- x)) else ulp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(x <= 0)%R
H2:(0 < eps <= ulp x)%R
H1':(x < 0)%R

(0 < succ (- x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(x <= 0)%R
H2:(0 < eps <= ulp x)%R
H1':x = 0%R
(0 < succ (- x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(0 < x)%R
H2:(0 < eps <= ulp (pred x))%R
(0 < eps <= (if Rle_bool (succ (- x)) 0 then ulp (succ (- x)) else ulp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(x <= 0)%R
H2:(0 < eps <= ulp x)%R
H1':(x < 0)%R

(0 < - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(x <= 0)%R
H2:(0 < eps <= ulp x)%R
H1':(x < 0)%R
(- x <= succ (- x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(x <= 0)%R
H2:(0 < eps <= ulp x)%R
H1':x = 0%R
(0 < succ (- x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(0 < x)%R
H2:(0 < eps <= ulp (pred x))%R
(0 < eps <= (if Rle_bool (succ (- x)) 0 then ulp (succ (- x)) else ulp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(x <= 0)%R
H2:(0 < eps <= ulp x)%R
H1':(x < 0)%R

(- x <= succ (- x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(x <= 0)%R
H2:(0 < eps <= ulp x)%R
H1':x = 0%R
(0 < succ (- x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(0 < x)%R
H2:(0 < eps <= ulp (pred x))%R
(0 < eps <= (if Rle_bool (succ (- x)) 0 then ulp (succ (- x)) else ulp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(x <= 0)%R
H2:(0 < eps <= ulp x)%R
H1':x = 0%R

(0 < succ (- x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(0 < x)%R
H2:(0 < eps <= ulp (pred x))%R
(0 < eps <= (if Rle_bool (succ (- x)) 0 then ulp (succ (- x)) else ulp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(x <= 0)%R
H2:(0 < eps <= ulp x)%R
H1':x = 0%R

(0 < 0 + ulp 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(0 < x)%R
H2:(0 < eps <= ulp (pred x))%R
(0 < eps <= (if Rle_bool (succ (- x)) 0 then ulp (succ (- x)) else ulp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(x <= 0)%R
H2:(0 < eps <= ulp x)%R
H1':x = 0%R

(0 < ulp 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(0 < x)%R
H2:(0 < eps <= ulp (pred x))%R
(0 < eps <= (if Rle_bool (succ (- x)) 0 then ulp (succ (- x)) else ulp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(x <= 0)%R
H2:(0 < eps <= ulp x)%R
H1':x = 0%R

(eps <= ulp 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(0 < x)%R
H2:(0 < eps <= ulp (pred x))%R
(0 < eps <= (if Rle_bool (succ (- x)) 0 then ulp (succ (- x)) else ulp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(0 < x)%R
H2:(0 < eps <= ulp (pred x))%R

(0 < eps <= (if Rle_bool (succ (- x)) 0 then ulp (succ (- x)) else ulp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(0 < x)%R
H2:(0 < eps <= ulp (pred x))%R

(0 < eps <= ulp (succ (- x)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(0 < x)%R
H2:(0 < eps <= ulp (pred x))%R
(succ (- x) <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(0 < x)%R
H2:(0 < eps <= ulp (pred x))%R

(succ (- x) <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(0 < x)%R
H2:(0 < eps <= ulp (pred x))%R

(- pred x <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
H1:(0 < x)%R
H2:(0 < eps <= ulp (pred x))%R

(0 <= pred x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R

F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
eps:R
Heps:(0 < eps <= (if Rle_bool x 0 then ulp x else ulp (pred x)))%R

F (- x)
now apply generic_format_opp. Qed.
Error of a rounding, expressed in number of ulps
false for x=0 in the FLX format
(* was ulp_error *)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall rnd : R -> Z, Valid_rnd rnd -> forall x : R, x <> 0%R -> (Rabs (round beta fexp rnd x - x) < ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall rnd : R -> Z, Valid_rnd rnd -> forall x : R, x <> 0%R -> (Rabs (round beta fexp rnd x - x) < ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R

(Rabs (round beta fexp rnd x - x) < ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:F x

(Rabs (round beta fexp rnd x - x) < ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x
(Rabs (round beta fexp rnd x - x) < ulp x)%R
(* x = rnd x *)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:F x

(Rabs (x - x) < ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x
(Rabs (round beta fexp rnd x - x) < ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:F x

(Rabs (x + - x) < ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x
(Rabs (round beta fexp rnd x - x) < ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:F x

(0 < ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x
(Rabs (round beta fexp rnd x - x) < ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:F x

(0 < bpow (cexp beta fexp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x
(Rabs (round beta fexp rnd x - x) < ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x

(Rabs (round beta fexp rnd x - x) < ulp x)%R
(* x <> rnd x *)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x

(Rabs (round beta fexp Zfloor x - x) < ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x
(Rabs (round beta fexp Zceil x - x) < ulp x)%R
(* . *)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x

(- (round beta fexp Zfloor x - x) < ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x
(round beta fexp Zfloor x - x <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x
(Rabs (round beta fexp Zceil x - x) < ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x

(x - round beta fexp Zfloor x < ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x
(round beta fexp Zfloor x - x <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x
(Rabs (round beta fexp Zceil x - x) < ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x

(round beta fexp Zfloor x + (x - round beta fexp Zfloor x) < round beta fexp Zfloor x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x
(round beta fexp Zfloor x - x <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x
(Rabs (round beta fexp Zceil x - x) < ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x

(round beta fexp Zfloor x + (x - round beta fexp Zfloor x) < round beta fexp Zceil x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x
(round beta fexp Zfloor x - x <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x
(Rabs (round beta fexp Zceil x - x) < ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x

(x < round beta fexp Zceil x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x
(round beta fexp Zfloor x - x <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x
(Rabs (round beta fexp Zceil x - x) < ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x

(x <= round beta fexp Zceil x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x
Hu:(x <= round beta fexp Zceil x)%R
(x < round beta fexp Zceil x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x
(round beta fexp Zfloor x - x <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x
(Rabs (round beta fexp Zceil x - x) < ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x
Hu:(x <= round beta fexp Zceil x)%R

(x < round beta fexp Zceil x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x
(round beta fexp Zfloor x - x <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x
(Rabs (round beta fexp Zceil x - x) < ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x
Hu:(x < round beta fexp Zceil x)%R

(x < round beta fexp Zceil x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x
Hu:x = round beta fexp Zceil x
(x < round beta fexp Zceil x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x
(round beta fexp Zfloor x - x <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x
(Rabs (round beta fexp Zceil x - x) < ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x
Hu:x = round beta fexp Zceil x

(x < round beta fexp Zceil x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x
(round beta fexp Zfloor x - x <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x
(Rabs (round beta fexp Zceil x - x) < ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x
Hu:x = round beta fexp Zceil x

F x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x
(round beta fexp Zfloor x - x <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x
(Rabs (round beta fexp Zceil x - x) < ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x
Hu:x = round beta fexp Zceil x

F (round beta fexp Zceil x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x
(round beta fexp Zfloor x - x <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x
(Rabs (round beta fexp Zceil x - x) < ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x

(round beta fexp Zfloor x - x <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x
(Rabs (round beta fexp Zceil x - x) < ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x

(round beta fexp Zfloor x <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x
(Rabs (round beta fexp Zceil x - x) < ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x

(Rabs (round beta fexp Zceil x - x) < ulp x)%R
(* . *)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x

(round beta fexp Zceil x - x < ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x
(0 <= round beta fexp Zceil x - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x

(round beta fexp Zfloor x + ulp x - x < ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x
(0 <= round beta fexp Zceil x - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x

(round beta fexp Zfloor x + ulp x - x + (x - ulp x) < ulp x + (x - ulp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x
(0 <= round beta fexp Zceil x - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x

(round beta fexp Zfloor x < x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x
(0 <= round beta fexp Zceil x - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x

(round beta fexp Zfloor x <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x
Hd:(round beta fexp Zfloor x <= x)%R
(round beta fexp Zfloor x < x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x
(0 <= round beta fexp Zceil x - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x
Hd:(round beta fexp Zfloor x <= x)%R

(round beta fexp Zfloor x < x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x
(0 <= round beta fexp Zceil x - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x
Hd:(round beta fexp Zfloor x < x)%R

(round beta fexp Zfloor x < x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x
Hd:round beta fexp Zfloor x = x
(round beta fexp Zfloor x < x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x
(0 <= round beta fexp Zceil x - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x
Hd:round beta fexp Zfloor x = x

(round beta fexp Zfloor x < x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x
(0 <= round beta fexp Zceil x - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x
Hd:round beta fexp Zfloor x = x

F x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x
(0 <= round beta fexp Zceil x - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x
Hd:round beta fexp Zfloor x = x

F (round beta fexp Zfloor x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x
(0 <= round beta fexp Zceil x - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x

(0 <= round beta fexp Zceil x - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
Hx:~ F x

(x <= round beta fexp Zceil x)%R
apply round_UP_pt... Qed. (* was ulp_error_le *)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall rnd : R -> Z, Valid_rnd rnd -> forall x : R, (Rabs (round beta fexp rnd x - x) <= ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall rnd : R -> Z, Valid_rnd rnd -> forall x : R, (Rabs (round beta fexp rnd x - x) <= ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R

(Rabs (round beta fexp rnd x - x) <= ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R

x = 0%R -> (Rabs (round beta fexp rnd x - x) <= ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
x <> 0%R -> (Rabs (round beta fexp rnd x - x) <= ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x = 0%R

(Rabs (0 - 0) <= ulp 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
x <> 0%R -> (Rabs (round beta fexp rnd x - x) <= ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x = 0%R

(0 <= ulp 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
x <> 0%R -> (Rabs (round beta fexp rnd x - x) <= ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R

x <> 0%R -> (Rabs (round beta fexp rnd x - x) <= ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R

(Rabs (round beta fexp rnd x - x) < ulp x)%R
now apply error_lt_ulp. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall (choice : Z -> bool) (x : R), (Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall (choice : Z -> bool) (x : R), (Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R

(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:F x

(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp x)%R
(* x = rnd x *)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:F x

(Rabs (x - x) <= / 2 * ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:F x

(Rabs (x + - x) <= / 2 * ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:F x

(0 <= / 2 * ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:F x

(0 <= / 2)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:F x
(0 <= ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:F x

(0 < / 2)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:F x
(0 <= ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:F x

(0 < 2)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:F x
(0 <= ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:F x

(0 <= ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x

(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp x)%R
(* x <> rnd x *)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R

(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R

(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(x - d <= d + ulp x - x)%R

(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(d + ulp x - x < x - d)%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp x)%R
(* . rnd(x) = rndd(x) *)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(x - d <= d + ulp x - x)%R

(Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (d - x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(x - d <= d + ulp x - x)%R
(Rabs (d - x) <= / 2 * ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(d + ulp x - x < x - d)%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(x - d <= d + ulp x - x)%R

F d
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(x - d <= d + ulp x - x)%R
(Rabs (d - x) <= / 2 * ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(d + ulp x - x < x - d)%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(x - d <= d + ulp x - x)%R

(Rabs (d - x) <= / 2 * ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(d + ulp x - x < x - d)%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(x - d <= d + ulp x - x)%R

(- (d - x) <= / 2 * ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(x - d <= d + ulp x - x)%R
(d - x <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(d + ulp x - x < x - d)%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(x - d <= d + ulp x - x)%R

(x - d <= / 2 * ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(x - d <= d + ulp x - x)%R
(d - x <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(d + ulp x - x < x - d)%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(x - d <= d + ulp x - x)%R

(0 < 2)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(x - d <= d + ulp x - x)%R
((x - d) * 2 <= / 2 * ulp x * 2)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(x - d <= d + ulp x - x)%R
(d - x <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(d + ulp x - x < x - d)%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(x - d <= d + ulp x - x)%R

((x - d) * 2 <= / 2 * ulp x * 2)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(x - d <= d + ulp x - x)%R
(d - x <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(d + ulp x - x < x - d)%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(x - d <= d + ulp x - x)%R

((x - d) * 2 + (d - x) <= / 2 * ulp x * 2 + (d - x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(x - d <= d + ulp x - x)%R
(d - x <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(d + ulp x - x < x - d)%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(x - d <= d + ulp x - x)%R

(x - d <= - x + d + 2 * / 2 * ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(x - d <= d + ulp x - x)%R
(d - x <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(d + ulp x - x < x - d)%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(x - d <= d + ulp x - x)%R

(d + ulp x - x <= - x + d + 2 * / 2 * ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(x - d <= d + ulp x - x)%R
(d - x <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(d + ulp x - x < x - d)%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(x - d <= d + ulp x - x)%R

(d + ulp x - x)%R = (- x + d + 2 * / 2 * ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(x - d <= d + ulp x - x)%R
(d - x <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(d + ulp x - x < x - d)%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(x - d <= d + ulp x - x)%R

(d - x <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(d + ulp x - x < x - d)%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(x - d <= d + ulp x - x)%R

(d <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(d + ulp x - x < x - d)%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(d + ulp x - x < x - d)%R

(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp x)%R
(* . rnd(x) = rndu(x) *)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(d + ulp x - x < x - d)%R

(d + ulp x)%R = round beta fexp Zceil x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(d + ulp x - x < x - d)%R
Hu:(d + ulp x)%R = round beta fexp Zceil x
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(d + ulp x - x < x - d)%R

(round beta fexp Zfloor x + ulp x)%R = round beta fexp Zceil x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(d + ulp x - x < x - d)%R
Hu:(d + ulp x)%R = round beta fexp Zceil x
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(d + ulp x - x < x - d)%R
Hu:(d + ulp x)%R = round beta fexp Zceil x

(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(d + ulp x - x < x - d)%R
Hu:(d + ulp x)%R = round beta fexp Zceil x

(Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (d + ulp x - x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(d + ulp x - x < x - d)%R
Hu:(d + ulp x)%R = round beta fexp Zceil x
(Rabs (d + ulp x - x) <= / 2 * ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(d + ulp x - x < x - d)%R
Hu:(d + ulp x)%R = round beta fexp Zceil x

F (d + ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(d + ulp x - x < x - d)%R
Hu:(d + ulp x)%R = round beta fexp Zceil x
(Rabs (d + ulp x - x) <= / 2 * ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(d + ulp x - x < x - d)%R
Hu:(d + ulp x)%R = round beta fexp Zceil x

F (round beta fexp Zceil x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(d + ulp x - x < x - d)%R
Hu:(d + ulp x)%R = round beta fexp Zceil x
(Rabs (d + ulp x - x) <= / 2 * ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(d + ulp x - x < x - d)%R
Hu:(d + ulp x)%R = round beta fexp Zceil x

(Rabs (d + ulp x - x) <= / 2 * ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(d + ulp x - x < x - d)%R
Hu:(d + ulp x)%R = round beta fexp Zceil x

(d + ulp x - x <= / 2 * ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(d + ulp x - x < x - d)%R
Hu:(d + ulp x)%R = round beta fexp Zceil x
(0 <= d + ulp x - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(d + ulp x - x < x - d)%R
Hu:(d + ulp x)%R = round beta fexp Zceil x

(0 < 2)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(d + ulp x - x < x - d)%R
Hu:(d + ulp x)%R = round beta fexp Zceil x
((d + ulp x - x) * 2 <= / 2 * ulp x * 2)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(d + ulp x - x < x - d)%R
Hu:(d + ulp x)%R = round beta fexp Zceil x
(0 <= d + ulp x - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(d + ulp x - x < x - d)%R
Hu:(d + ulp x)%R = round beta fexp Zceil x

((d + ulp x - x) * 2 <= / 2 * ulp x * 2)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(d + ulp x - x < x - d)%R
Hu:(d + ulp x)%R = round beta fexp Zceil x
(0 <= d + ulp x - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(d + ulp x - x < x - d)%R
Hu:(d + ulp x)%R = round beta fexp Zceil x

((d + ulp x - x) * 2 + - (d + ulp x - x) <= / 2 * ulp x * 2 + - (d + ulp x - x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(d + ulp x - x < x - d)%R
Hu:(d + ulp x)%R = round beta fexp Zceil x
(0 <= d + ulp x - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(d + ulp x - x < x - d)%R
Hu:(d + ulp x)%R = round beta fexp Zceil x

(d + ulp x - x <= - d + 2 * ulp x * / 2 - ulp x + x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(d + ulp x - x < x - d)%R
Hu:(d + ulp x)%R = round beta fexp Zceil x
(0 <= d + ulp x - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(d + ulp x - x < x - d)%R
Hu:(d + ulp x)%R = round beta fexp Zceil x

(d + ulp x - x < - d + 2 * ulp x * / 2 - ulp x + x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(d + ulp x - x < x - d)%R
Hu:(d + ulp x)%R = round beta fexp Zceil x
(0 <= d + ulp x - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(d + ulp x - x < x - d)%R
Hu:(d + ulp x)%R = round beta fexp Zceil x

(x - d <= - d + 2 * ulp x * / 2 - ulp x + x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(d + ulp x - x < x - d)%R
Hu:(d + ulp x)%R = round beta fexp Zceil x
(0 <= d + ulp x - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(d + ulp x - x < x - d)%R
Hu:(d + ulp x)%R = round beta fexp Zceil x

(x - d)%R = (- d + 2 * ulp x * / 2 - ulp x + x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(d + ulp x - x < x - d)%R
Hu:(d + ulp x)%R = round beta fexp Zceil x
(0 <= d + ulp x - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(d + ulp x - x < x - d)%R
Hu:(d + ulp x)%R = round beta fexp Zceil x

(0 <= d + ulp x - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(d + ulp x - x < x - d)%R
Hu:(d + ulp x)%R = round beta fexp Zceil x

(x <= d + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
Hx:~ F x
d:=round beta fexp Zfloor x:R
Hr1:F (round beta fexp (Znearest choice) x)
Hr2:forall g : R, F g -> (Rabs (round beta fexp (Znearest choice) x - x) <= Rabs (g - x))%R
H:(d + ulp x - x < x - d)%R
Hu:(d + ulp x)%R = round beta fexp Zceil x

(x <= round beta fexp Zceil x)%R
apply (round_UP_pt beta fexp x). Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, (0 <= x)%R -> ulp (round beta fexp Zfloor x) = ulp x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, (0 <= x)%R -> ulp (round beta fexp Zfloor x) = ulp x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R

ulp (round beta fexp Zfloor x) = ulp x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:0%R = x
ulp (round beta fexp Zfloor x) = ulp x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R

ulp (round beta fexp Zfloor x) = ulp x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R

ulp (round beta fexp Zfloor x) = bpow (cexp beta fexp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R

F 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Hd:(0 < round beta fexp Zfloor x)%R
ulp (round beta fexp Zfloor x) = bpow (cexp beta fexp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Hd:0%R = round beta fexp Zfloor x
ulp (round beta fexp Zfloor x) = bpow (cexp beta fexp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R

(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Hd:(0 < round beta fexp Zfloor x)%R
ulp (round beta fexp Zfloor x) = bpow (cexp beta fexp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Hd:0%R = round beta fexp Zfloor x
ulp (round beta fexp Zfloor x) = bpow (cexp beta fexp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Hd:(0 < round beta fexp Zfloor x)%R

ulp (round beta fexp Zfloor x) = bpow (cexp beta fexp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Hd:0%R = round beta fexp Zfloor x
ulp (round beta fexp Zfloor x) = bpow (cexp beta fexp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Hd:(0 < round beta fexp Zfloor x)%R

ulp (round beta fexp Zfloor x) = bpow (cexp beta fexp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Hd:(0 < round beta fexp Zfloor x)%R

bpow (cexp beta fexp (round beta fexp Zfloor x)) = bpow (cexp beta fexp x)
now rewrite cexp_DN with (2 := Hd).
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Hd:0%R = round beta fexp Zfloor x

ulp (round beta fexp Zfloor x) = bpow (cexp beta fexp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Hd:0%R = round beta fexp Zfloor x

ulp 0 = bpow (cexp beta fexp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Hd:0%R = round beta fexp Zfloor x

ulp 0 = bpow (fexp (mag beta x))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Hd:0%R = round beta fexp Zfloor x
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R

ulp 0 = bpow (fexp (Build_mag_prop beta x e He))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Hd:0%R = round beta fexp Zfloor x
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R

ulp 0 = bpow (fexp e)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Hd:0%R = round beta fexp Zfloor x
e:Z
He:(bpow (e - 1) <= Rabs x < bpow e)%R

ulp 0 = bpow (fexp e)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Hd:round beta fexp Zfloor x = 0%R
e:Z
He:(bpow (e - 1) <= Rabs x < bpow e)%R

ulp 0 = bpow (fexp e)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Hd:round beta fexp Zfloor x = 0%R
e:Z
He:(bpow (e - 1) <= Rabs x < bpow e)%R
H:(e <= fexp e)%Z

ulp 0 = bpow (fexp e)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Hd:round beta fexp Zfloor x = 0%R
e:Z
He:(bpow (e - 1) <= Rabs x < bpow e)%R
H:(e <= fexp e)%Z

(if Req_bool 0 0 then match negligible_exp with | Some n => bpow (fexp n) | None => 0%R end else bpow (cexp beta fexp 0)) = bpow (fexp e)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Hd:round beta fexp Zfloor x = 0%R
e:Z
He:(bpow (e - 1) <= Rabs x < bpow e)%R
H:(e <= fexp e)%Z

match negligible_exp with | Some n => bpow (fexp n) | None => 0%R end = bpow (fexp e)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Hd:round beta fexp Zfloor x = 0%R
e:Z
He:(bpow (e - 1) <= Rabs x < bpow e)%R
H:(e <= fexp e)%Z
H0:forall n : Z, (fexp n < n)%Z

0%R = bpow (fexp e)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Hd:round beta fexp Zfloor x = 0%R
e:Z
He:(bpow (e - 1) <= Rabs x < bpow e)%R
H:(e <= fexp e)%Z
k:Z
Hk:(k <= fexp k)%Z
bpow (fexp k) = bpow (fexp e)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:(0 < x)%R
Hd:round beta fexp Zfloor x = 0%R
e:Z
He:(bpow (e - 1) <= Rabs x < bpow e)%R
H:(e <= fexp e)%Z
k:Z
Hk:(k <= fexp k)%Z

bpow (fexp k) = bpow (fexp e)
now apply f_equal, fexp_negligible_exp_eq.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Hx:0%R = x

ulp (round beta fexp Zfloor x) = ulp x
rewrite <- Hx, round_0... Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

negligible_exp = None -> forall rnd : R -> Z, Valid_rnd rnd -> forall x : R, x <> 0%R -> round beta fexp rnd x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

negligible_exp = None -> forall rnd : R -> Z, Valid_rnd rnd -> forall x : R, x <> 0%R -> round beta fexp rnd x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:negligible_exp = None
rndn:R -> Z
Hrnd:Valid_rnd rndn
x:R
Hx:x <> 0%R
K:round beta fexp rndn x = 0%R

False
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:negligible_exp = None
rndn:R -> Z
Hrnd:Valid_rnd rndn
x:R
Hx:x <> 0%R
K:round beta fexp rndn x = 0%R

negligible_exp = None /\ (forall n : Z, (fexp n < n)%Z) -> False
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:negligible_exp = None
rndn:R -> Z
Hrnd:Valid_rnd rndn
x:R
Hx:x <> 0%R
K:round beta fexp rndn x = 0%R
(exists n : Z, negligible_exp = Some n /\ (n <= fexp n)%Z) -> False
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:negligible_exp = None
rndn:R -> Z
Hrnd:Valid_rnd rndn
x:R
Hx:x <> 0%R
K:round beta fexp rndn x = 0%R
Hn:forall n : Z, (fexp n < n)%Z

False
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:negligible_exp = None
rndn:R -> Z
Hrnd:Valid_rnd rndn
x:R
Hx:x <> 0%R
K:round beta fexp rndn x = 0%R
(exists n : Z, negligible_exp = Some n /\ (n <= fexp n)%Z) -> False
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:negligible_exp = None
rndn:R -> Z
Hrnd:Valid_rnd rndn
x:R
Hx:x <> 0%R
K:round beta fexp rndn x = 0%R
Hn:forall n : Z, (fexp n < n)%Z
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R

False
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:negligible_exp = None
rndn:R -> Z
Hrnd:Valid_rnd rndn
x:R
Hx:x <> 0%R
K:round beta fexp rndn x = 0%R
(exists n : Z, negligible_exp = Some n /\ (n <= fexp n)%Z) -> False
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:negligible_exp = None
rndn:R -> Z
Hrnd:Valid_rnd rndn
x:R
Hx:x <> 0%R
K:round beta fexp rndn x = 0%R
Hn:forall n : Z, (fexp n < n)%Z
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R

~ (fexp e < e)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:negligible_exp = None
rndn:R -> Z
Hrnd:Valid_rnd rndn
x:R
Hx:x <> 0%R
K:round beta fexp rndn x = 0%R
Hn:forall n : Z, (fexp n < n)%Z
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R
(fexp e < e)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:negligible_exp = None
rndn:R -> Z
Hrnd:Valid_rnd rndn
x:R
Hx:x <> 0%R
K:round beta fexp rndn x = 0%R
(exists n : Z, negligible_exp = Some n /\ (n <= fexp n)%Z) -> False
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:negligible_exp = None
rndn:R -> Z
Hrnd:Valid_rnd rndn
x:R
Hx:x <> 0%R
K:round beta fexp rndn x = 0%R
Hn:forall n : Z, (fexp n < n)%Z
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R

(e <= fexp e)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:negligible_exp = None
rndn:R -> Z
Hrnd:Valid_rnd rndn
x:R
Hx:x <> 0%R
K:round beta fexp rndn x = 0%R
Hn:forall n : Z, (fexp n < n)%Z
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R
(fexp e < e)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:negligible_exp = None
rndn:R -> Z
Hrnd:Valid_rnd rndn
x:R
Hx:x <> 0%R
K:round beta fexp rndn x = 0%R
(exists n : Z, negligible_exp = Some n /\ (n <= fexp n)%Z) -> False
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:negligible_exp = None
rndn:R -> Z
Hrnd:Valid_rnd rndn
x:R
Hx:x <> 0%R
K:round beta fexp rndn x = 0%R
Hn:forall n : Z, (fexp n < n)%Z
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R

(fexp e < e)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:negligible_exp = None
rndn:R -> Z
Hrnd:Valid_rnd rndn
x:R
Hx:x <> 0%R
K:round beta fexp rndn x = 0%R
(exists n : Z, negligible_exp = Some n /\ (n <= fexp n)%Z) -> False
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:negligible_exp = None
rndn:R -> Z
Hrnd:Valid_rnd rndn
x:R
Hx:x <> 0%R
K:round beta fexp rndn x = 0%R

(exists n : Z, negligible_exp = Some n /\ (n <= fexp n)%Z) -> False
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:negligible_exp = None
rndn:R -> Z
Hrnd:Valid_rnd rndn
x:R
Hx:x <> 0%R
K:round beta fexp rndn x = 0%R
n:Z
H1:negligible_exp = Some n

False
rewrite H in H1; discriminate. Qed.
allows rnd x to be 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

Monotone_exp fexp -> forall rnd : R -> Z, Valid_rnd rnd -> forall x : R, x <> 0%R -> (Rabs (round beta fexp rnd x - x) < ulp (round beta fexp rnd x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

Monotone_exp fexp -> forall rnd : R -> Z, Valid_rnd rnd -> forall x : R, x <> 0%R -> (Rabs (round beta fexp rnd x - x) < ulp (round beta fexp rnd x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp

forall rnd : R -> Z, Valid_rnd rnd -> forall x : R, x <> 0%R -> (Rabs (round beta fexp rnd x - x) < ulp (round beta fexp rnd x))%R
(* wlog *)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp

(forall rnd : R -> Z, Valid_rnd rnd -> forall x : R, (0 < x)%R -> (Rabs (round beta fexp rnd x - x) < ulp (round beta fexp rnd x))%R) -> forall rnd : R -> Z, Valid_rnd rnd -> forall x : R, x <> 0%R -> (Rabs (round beta fexp rnd x - x) < ulp (round beta fexp rnd x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
forall rnd : R -> Z, Valid_rnd rnd -> forall x : R, (0 < x)%R -> (Rabs (round beta fexp rnd x - x) < ulp (round beta fexp rnd x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
M:forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 < x0)%R -> (Rabs (round beta fexp rnd0 x0 - x0) < ulp (round beta fexp rnd0 x0))%R
rnd:R -> Z
Hrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R

(Rabs (round beta fexp rnd x - x) < ulp (round beta fexp rnd x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
forall rnd : R -> Z, Valid_rnd rnd -> forall x : R, (0 < x)%R -> (Rabs (round beta fexp rnd x - x) < ulp (round beta fexp rnd x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
M:forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 < x0)%R -> (Rabs (round beta fexp rnd0 x0 - x0) < ulp (round beta fexp rnd0 x0))%R
rnd:R -> Z
Hrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R

(0 <= x)%R -> (Rabs (round beta fexp rnd x - x) < ulp (round beta fexp rnd x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
M:forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 < x0)%R -> (Rabs (round beta fexp rnd0 x0 - x0) < ulp (round beta fexp rnd0 x0))%R
rnd:R -> Z
Hrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
(x < 0)%R -> (Rabs (round beta fexp rnd x - x) < ulp (round beta fexp rnd x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
forall rnd : R -> Z, Valid_rnd rnd -> forall x : R, (0 < x)%R -> (Rabs (round beta fexp rnd x - x) < ulp (round beta fexp rnd x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
M:forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 < x0)%R -> (Rabs (round beta fexp rnd0 x0 - x0) < ulp (round beta fexp rnd0 x0))%R
rnd:R -> Z
Hrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
H:(0 < x)%R

(Rabs (round beta fexp rnd x - x) < ulp (round beta fexp rnd x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
M:forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 < x0)%R -> (Rabs (round beta fexp rnd0 x0 - x0) < ulp (round beta fexp rnd0 x0))%R
rnd:R -> Z
Hrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
H:0%R = x
(Rabs (round beta fexp rnd x - x) < ulp (round beta fexp rnd x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
M:forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 < x0)%R -> (Rabs (round beta fexp rnd0 x0 - x0) < ulp (round beta fexp rnd0 x0))%R
rnd:R -> Z
Hrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
(x < 0)%R -> (Rabs (round beta fexp rnd x - x) < ulp (round beta fexp rnd x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
forall rnd : R -> Z, Valid_rnd rnd -> forall x : R, (0 < x)%R -> (Rabs (round beta fexp rnd x - x) < ulp (round beta fexp rnd x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
M:forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 < x0)%R -> (Rabs (round beta fexp rnd0 x0 - x0) < ulp (round beta fexp rnd0 x0))%R
rnd:R -> Z
Hrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
H:0%R = x

(Rabs (round beta fexp rnd x - x) < ulp (round beta fexp rnd x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
M:forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 < x0)%R -> (Rabs (round beta fexp rnd0 x0 - x0) < ulp (round beta fexp rnd0 x0))%R
rnd:R -> Z
Hrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
(x < 0)%R -> (Rabs (round beta fexp rnd x - x) < ulp (round beta fexp rnd x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
forall rnd : R -> Z, Valid_rnd rnd -> forall x : R, (0 < x)%R -> (Rabs (round beta fexp rnd x - x) < ulp (round beta fexp rnd x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
M:forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 < x0)%R -> (Rabs (round beta fexp rnd0 x0 - x0) < ulp (round beta fexp rnd0 x0))%R
rnd:R -> Z
Hrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R

(x < 0)%R -> (Rabs (round beta fexp rnd x - x) < ulp (round beta fexp rnd x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
forall rnd : R -> Z, Valid_rnd rnd -> forall x : R, (0 < x)%R -> (Rabs (round beta fexp rnd x - x) < ulp (round beta fexp rnd x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
M:forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 < x0)%R -> (Rabs (round beta fexp rnd0 x0 - x0) < ulp (round beta fexp rnd0 x0))%R
rnd:R -> Z
Hrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
H:(x < 0)%R

(Rabs (round beta fexp rnd x - x) < ulp (round beta fexp rnd x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
forall rnd : R -> Z, Valid_rnd rnd -> forall x : R, (0 < x)%R -> (Rabs (round beta fexp rnd x - x) < ulp (round beta fexp rnd x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
M:forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 < x0)%R -> (Rabs (round beta fexp rnd0 x0 - x0) < ulp (round beta fexp rnd0 x0))%R
rnd:R -> Z
Hrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
H:(x < 0)%R

(Rabs (round beta fexp rnd (- - x) - - - x) < ulp (round beta fexp rnd (- - x)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
forall rnd : R -> Z, Valid_rnd rnd -> forall x : R, (0 < x)%R -> (Rabs (round beta fexp rnd x - x) < ulp (round beta fexp rnd x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
M:forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 < x0)%R -> (Rabs (round beta fexp rnd0 x0 - x0) < ulp (round beta fexp rnd0 x0))%R
rnd:R -> Z
Hrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
H:(x < 0)%R

(Rabs (- round beta fexp (Zrnd_opp rnd) (- x) - - - x) < ulp (round beta fexp (Zrnd_opp rnd) (- x)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
forall rnd : R -> Z, Valid_rnd rnd -> forall x : R, (0 < x)%R -> (Rabs (round beta fexp rnd x - x) < ulp (round beta fexp rnd x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
M:forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 < x0)%R -> (Rabs (round beta fexp rnd0 x0 - x0) < ulp (round beta fexp rnd0 x0))%R
rnd:R -> Z
Hrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
H:(x < 0)%R

(Rabs (- (round beta fexp (Zrnd_opp rnd) (- x) - - x)) < ulp (round beta fexp (Zrnd_opp rnd) (- x)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
forall rnd : R -> Z, Valid_rnd rnd -> forall x : R, (0 < x)%R -> (Rabs (round beta fexp rnd x - x) < ulp (round beta fexp rnd x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
M:forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 < x0)%R -> (Rabs (round beta fexp rnd0 x0 - x0) < ulp (round beta fexp rnd0 x0))%R
rnd:R -> Z
Hrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
H:(x < 0)%R

(Rabs (round beta fexp (Zrnd_opp rnd) (- x) - - x) < ulp (round beta fexp (Zrnd_opp rnd) (- x)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
forall rnd : R -> Z, Valid_rnd rnd -> forall x : R, (0 < x)%R -> (Rabs (round beta fexp rnd x - x) < ulp (round beta fexp rnd x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
M:forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 < x0)%R -> (Rabs (round beta fexp rnd0 x0 - x0) < ulp (round beta fexp rnd0 x0))%R
rnd:R -> Z
Hrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
H:(x < 0)%R

Valid_rnd (Zrnd_opp rnd)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
M:forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 < x0)%R -> (Rabs (round beta fexp rnd0 x0 - x0) < ulp (round beta fexp rnd0 x0))%R
rnd:R -> Z
Hrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
H:(x < 0)%R
(0 < - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
forall rnd : R -> Z, Valid_rnd rnd -> forall x : R, (0 < x)%R -> (Rabs (round beta fexp rnd x - x) < ulp (round beta fexp rnd x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
M:forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 < x0)%R -> (Rabs (round beta fexp rnd0 x0 - x0) < ulp (round beta fexp rnd0 x0))%R
rnd:R -> Z
Hrnd:Valid_rnd rnd
x:R
Zx:x <> 0%R
H:(x < 0)%R

(0 < - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
forall rnd : R -> Z, Valid_rnd rnd -> forall x : R, (0 < x)%R -> (Rabs (round beta fexp rnd x - x) < ulp (round beta fexp rnd x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp

forall rnd : R -> Z, Valid_rnd rnd -> forall x : R, (0 < x)%R -> (Rabs (round beta fexp rnd x - x) < ulp (round beta fexp rnd x))%R
(* 0 < x *)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
rnd:R -> Z
Hrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R

(Rabs (round beta fexp rnd x - x) < ulp (round beta fexp rnd x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
rnd:R -> Z
Hrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R

(Rabs (round beta fexp rnd x - x) < ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
rnd:R -> Z
Hrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
(ulp x <= ulp (round beta fexp rnd x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
rnd:R -> Z
Hrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R

x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
rnd:R -> Z
Hrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
(ulp x <= ulp (round beta fexp rnd x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
rnd:R -> Z
Hrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R

(ulp x <= ulp (round beta fexp rnd x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
rnd:R -> Z
Hrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R

(ulp (round beta fexp Zfloor x) <= ulp (round beta fexp rnd x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
rnd:R -> Z
Hrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
rnd:R -> Z
Hrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R

(0 <= round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
rnd:R -> Z
Hrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
(round beta fexp Zfloor x <= round beta fexp rnd x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
rnd:R -> Z
Hrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
rnd:R -> Z
Hrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R

F 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
rnd:R -> Z
Hrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
rnd:R -> Z
Hrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
(round beta fexp Zfloor x <= round beta fexp rnd x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
rnd:R -> Z
Hrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
rnd:R -> Z
Hrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R

(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
rnd:R -> Z
Hrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
(round beta fexp Zfloor x <= round beta fexp rnd x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
rnd:R -> Z
Hrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
rnd:R -> Z
Hrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R

(round beta fexp Zfloor x <= round beta fexp rnd x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
rnd:R -> Z
Hrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
rnd:R -> Z
Hrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
V:round beta fexp rnd x = round beta fexp Zfloor x

(round beta fexp Zfloor x <= round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
rnd:R -> Z
Hrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
V:round beta fexp rnd x = round beta fexp Zceil x
(round beta fexp Zfloor x <= round beta fexp Zceil x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
rnd:R -> Z
Hrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
rnd:R -> Z
Hrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
V:round beta fexp rnd x = round beta fexp Zceil x

(round beta fexp Zfloor x <= round beta fexp Zceil x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
rnd:R -> Z
Hrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
rnd:R -> Z
Hrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
V:round beta fexp rnd x = round beta fexp Zceil x

(round beta fexp Zfloor x <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
rnd:R -> Z
Hrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
V:round beta fexp rnd x = round beta fexp Zceil x
(x <= round beta fexp Zceil x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
rnd:R -> Z
Hrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
rnd:R -> Z
Hrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
V:round beta fexp rnd x = round beta fexp Zceil x

(x <= round beta fexp Zceil x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
rnd:R -> Z
Hrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
rnd:R -> Z
Hrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R

(0 <= x)%R
now apply Rlt_le. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

Monotone_exp fexp -> forall rnd : R -> Z, Valid_rnd rnd -> forall x : R, (Rabs (round beta fexp rnd x - x) <= ulp (round beta fexp rnd x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

Monotone_exp fexp -> forall rnd : R -> Z, Valid_rnd rnd -> forall x : R, (Rabs (round beta fexp rnd x - x) <= ulp (round beta fexp rnd x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
rnd:R -> Z
Vrnd:Valid_rnd rnd
x:R

(Rabs (round beta fexp rnd x - x) <= ulp (round beta fexp rnd x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
rnd:R -> Z
Vrnd:Valid_rnd rnd
x:R
Zx:x = 0%R

(Rabs (round beta fexp rnd x - x) <= ulp (round beta fexp rnd x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
rnd:R -> Z
Vrnd:Valid_rnd rnd
x:R
Nzx:x <> 0%R
(Rabs (round beta fexp rnd x - x) <= ulp (round beta fexp rnd x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
rnd:R -> Z
Vrnd:Valid_rnd rnd
x:R
Zx:x = 0%R

(Rabs (round beta fexp rnd x - x) <= ulp (round beta fexp rnd x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
rnd:R -> Z
Vrnd:Valid_rnd rnd
x:R
Zx:x = 0%R

(Rabs (0 - 0) <= ulp 0)%R
unfold Rminus; rewrite Ropp_0, Rplus_0_l, Rabs_R0; apply ulp_ge_0.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
rnd:R -> Z
Vrnd:Valid_rnd rnd
x:R
Nzx:x <> 0%R

(Rabs (round beta fexp rnd x - x) <= ulp (round beta fexp rnd x))%R
now apply Rlt_le, error_lt_ulp_round. Qed.
allows both x and rnd x to be 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

Monotone_exp fexp -> forall (choice : Z -> bool) (x : R), (Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

Monotone_exp fexp -> forall (choice : Z -> bool) (x : R), (Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R

(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x = 0%R

(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
(* *)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x = 0%R
Hx:x = 0%R

(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x = 0%R
Hx:x <> 0%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x = 0%R
Hx:x = 0%R

(/ 2 * ulp x <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x = 0%R
Hx:x <> 0%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x = 0%R
Hx:x = 0%R

(/ 2 * ulp 0 <= / 2 * ulp 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x = 0%R
Hx:x <> 0%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x = 0%R
Hx:x <> 0%R

(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x = 0%R
Hx:x <> 0%R

(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp x)%R -> (Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x = 0%R
Hx:x <> 0%R

(Rabs (0 - x) <= / 2 * ulp x)%R -> (Rabs (0 - x) <= / 2 * ulp 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x = 0%R
Hx:x <> 0%R

(Rabs x <= / 2 * ulp x)%R -> (Rabs x <= / 2 * ulp 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x = 0%R
Hx:x <> 0%R
N:(Rabs x <= / 2 * ulp x)%R

(Rabs x <= / 2 * ulp 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x = 0%R
Hx:x <> 0%R
N:(Rabs x <= / 2 * ulp x)%R

(Rabs x <= / 2 * match negligible_exp with | Some n => bpow (fexp n) | None => 0 end)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x = 0%R
Hx:x <> 0%R
N:(Rabs x <= / 2 * ulp x)%R

negligible_exp = None /\ (forall n : Z, (fexp n < n)%Z) -> (Rabs x <= / 2 * match negligible_exp with | Some n => bpow (fexp n) | None => 0 end)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x = 0%R
Hx:x <> 0%R
N:(Rabs x <= / 2 * ulp x)%R
(exists n : Z, negligible_exp = Some n /\ (n <= fexp n)%Z) -> (Rabs x <= / 2 * match negligible_exp with | Some n => bpow (fexp n) | None => 0 end)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x = 0%R
Hx:x <> 0%R
N:(Rabs x <= / 2 * ulp x)%R
H1:negligible_exp = None
H2:forall n : Z, (fexp n < n)%Z

(Rabs x <= / 2 * match negligible_exp with | Some n => bpow (fexp n) | None => 0 end)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x = 0%R
Hx:x <> 0%R
N:(Rabs x <= / 2 * ulp x)%R
(exists n : Z, negligible_exp = Some n /\ (n <= fexp n)%Z) -> (Rabs x <= / 2 * match negligible_exp with | Some n => bpow (fexp n) | None => 0 end)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hx:x <> 0%R
N:(Rabs x <= / 2 * ulp x)%R
H1:negligible_exp = None
H2:forall n : Z, (fexp n < n)%Z

round beta fexp (Znearest choice) x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x = 0%R
Hx:x <> 0%R
N:(Rabs x <= / 2 * ulp x)%R
(exists n : Z, negligible_exp = Some n /\ (n <= fexp n)%Z) -> (Rabs x <= / 2 * match negligible_exp with | Some n => bpow (fexp n) | None => 0 end)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x = 0%R
Hx:x <> 0%R
N:(Rabs x <= / 2 * ulp x)%R

(exists n : Z, negligible_exp = Some n /\ (n <= fexp n)%Z) -> (Rabs x <= / 2 * match negligible_exp with | Some n => bpow (fexp n) | None => 0 end)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x = 0%R
Hx:x <> 0%R
N:(Rabs x <= / 2 * ulp x)%R
n:Z
H1:negligible_exp = Some n
Hn:(n <= fexp n)%Z

(Rabs x <= / 2 * bpow (fexp n))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x = 0%R
Hx:x <> 0%R
N:(Rabs x <= / 2 * ulp x)%R
n:Z
H1:negligible_exp = Some n
Hn:(n <= fexp n)%Z

(/ 2 * ulp x <= / 2 * bpow (fexp n))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x = 0%R
Hx:x <> 0%R
N:(Rabs x <= / 2 * ulp x)%R
n:Z
H1:negligible_exp = Some n
Hn:(n <= fexp n)%Z

ulp x = bpow (fexp n)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x = 0%R
Hx:x <> 0%R
N:(Rabs x <= / 2 * ulp x)%R
n:Z
H1:negligible_exp = Some n
Hn:(n <= fexp n)%Z

bpow (cexp beta fexp x) = bpow (fexp n)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x = 0%R
Hx:x <> 0%R
N:(Rabs x <= / 2 * ulp x)%R
n:Z
H1:negligible_exp = Some n
Hn:(n <= fexp n)%Z

cexp beta fexp x = fexp n
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x = 0%R
Hx:x <> 0%R
N:(Rabs x <= / 2 * ulp x)%R
n:Z
H1:negligible_exp = Some n
Hn:(n <= fexp n)%Z

fexp (mag beta x) = fexp n
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x = 0%R
Hx:x <> 0%R
N:(Rabs x <= / 2 * ulp x)%R
n:Z
H1:negligible_exp = Some n
Hn:(n <= fexp n)%Z

(mag beta x <= fexp n)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x = 0%R
Hx:x <> 0%R
N:(Rabs x <= / 2 * ulp x)%R
n:Z
H1:negligible_exp = Some n
Hn:(n <= fexp n)%Z

(mag beta x - 1 < fexp n)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x = 0%R
Hx:x <> 0%R
N:(Rabs x <= / 2 * ulp x)%R
n:Z
H1:negligible_exp = Some n
Hn:(n <= fexp n)%Z

(bpow (mag beta x - 1) < bpow (fexp n))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x = 0%R
Hx:x <> 0%R
N:(Rabs x <= / 2 * ulp x)%R
n:Z
H1:negligible_exp = Some n
Hn:(n <= fexp n)%Z
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R

(bpow (Build_mag_prop beta x e He - 1) < bpow (fexp n))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x = 0%R
Hx:x <> 0%R
N:(Rabs x <= / 2 * ulp x)%R
n:Z
H1:negligible_exp = Some n
Hn:(n <= fexp n)%Z
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R

(bpow (e - 1) < bpow (fexp n))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x = 0%R
Hx:x <> 0%R
N:(Rabs x <= / 2 * ulp x)%R
n:Z
H1:negligible_exp = Some n
Hn:(n <= fexp n)%Z
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R

(bpow (e - 1) <= Rabs x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x = 0%R
Hx:x <> 0%R
N:(Rabs x <= / 2 * ulp x)%R
n:Z
H1:negligible_exp = Some n
Hn:(n <= fexp n)%Z
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R
(Rabs x < bpow (fexp n))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x = 0%R
Hx:x <> 0%R
N:(Rabs x <= / 2 * ulp x)%R
n:Z
H1:negligible_exp = Some n
Hn:(n <= fexp n)%Z
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R

(Rabs x < bpow (fexp n))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x = 0%R
Hx:x <> 0%R
N:(Rabs x <= / 2 * ulp x)%R
n:Z
H1:negligible_exp = Some n
Hn:(n <= fexp n)%Z
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R

(Rabs x <= Rabs (round beta fexp (Znearest choice) x - x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x = 0%R
Hx:x <> 0%R
N:(Rabs x <= / 2 * ulp x)%R
n:Z
H1:negligible_exp = Some n
Hn:(n <= fexp n)%Z
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R
(Rabs (round beta fexp (Znearest choice) x - x) < bpow (fexp n))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x = 0%R
Hx:x <> 0%R
N:(Rabs x <= / 2 * ulp x)%R
n:Z
H1:negligible_exp = Some n
Hn:(n <= fexp n)%Z
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R

Rabs x = Rabs (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x = 0%R
Hx:x <> 0%R
N:(Rabs x <= / 2 * ulp x)%R
n:Z
H1:negligible_exp = Some n
Hn:(n <= fexp n)%Z
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R
(Rabs (round beta fexp (Znearest choice) x - x) < bpow (fexp n))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x = 0%R
Hx:x <> 0%R
N:(Rabs x <= / 2 * ulp x)%R
n:Z
H1:negligible_exp = Some n
Hn:(n <= fexp n)%Z
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R

(Rabs (round beta fexp (Znearest choice) x - x) < bpow (fexp n))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x = 0%R
Hx:x <> 0%R
N:(Rabs x <= / 2 * ulp x)%R
n:Z
H1:negligible_exp = Some n
Hn:(n <= fexp n)%Z
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R

(Rabs (round beta fexp (Znearest choice) x - x) < ulp 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x = 0%R
Hx:x <> 0%R
N:(Rabs x <= / 2 * ulp x)%R
n:Z
H1:negligible_exp = Some n
Hn:(n <= fexp n)%Z
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R
(ulp 0 <= bpow (fexp n))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x = 0%R
Hx:x <> 0%R
N:(Rabs x <= / 2 * ulp x)%R
n:Z
H1:negligible_exp = Some n
Hn:(n <= fexp n)%Z
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R

(Rabs (round beta fexp (Znearest choice) x - x) < ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x = 0%R
Hx:x <> 0%R
N:(Rabs x <= / 2 * ulp x)%R
n:Z
H1:negligible_exp = Some n
Hn:(n <= fexp n)%Z
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R
(ulp 0 <= bpow (fexp n))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x = 0%R
Hx:x <> 0%R
N:(Rabs x <= / 2 * ulp x)%R
n:Z
H1:negligible_exp = Some n
Hn:(n <= fexp n)%Z
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R

(ulp 0 <= bpow (fexp n))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x = 0%R
Hx:x <> 0%R
N:(Rabs x <= / 2 * ulp x)%R
n:Z
H1:negligible_exp = Some n
Hn:(n <= fexp n)%Z
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R

(bpow (fexp n) <= bpow (fexp n))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R

(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
(* *)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zfloor x

(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zceil x
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
(* . *)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zfloor x
H:(0 <= x)%R

(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zfloor x
H:(x < 0)%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zceil x
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zfloor x
H:(0 <= x)%R

(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp Zfloor x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zfloor x
H:(x < 0)%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zceil x
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zfloor x
H:(0 <= x)%R

(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zfloor x
H:(x < 0)%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zceil x
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zfloor x
H:(x < 0)%R

(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zceil x
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zfloor x
H:(x < 0)%R

(/ 2 * ulp x <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zceil x
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zfloor x
H:(x < 0)%R

(0 <= / 2)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zfloor x
H:(x < 0)%R
(ulp x <= ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zceil x
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zfloor x
H:(x < 0)%R

(ulp x <= ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zceil x
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zfloor x
H:(x < 0)%R

(Rabs x <= Rabs (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zceil x
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zfloor x
H:(x < 0)%R

(- x <= Rabs (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zceil x
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zfloor x
H:(x < 0)%R

(- x <= Rabs (round beta fexp Zfloor x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zceil x
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zfloor x
H:(x < 0)%R

(- x <= - round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zfloor x
H:(x < 0)%R
(round beta fexp Zfloor x <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zceil x
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zfloor x
H:(x < 0)%R

(round beta fexp Zfloor x <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zfloor x
H:(x < 0)%R
(round beta fexp Zfloor x <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zceil x
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zfloor x
H:(x < 0)%R

(round beta fexp Zfloor x <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zceil x
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zfloor x
H:(x < 0)%R

F 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zfloor x
H:(x < 0)%R
(x <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zceil x
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zfloor x
H:(x < 0)%R

(x <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zceil x
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zceil x

(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
(* . *)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zceil x
H:(0 <= x)%R

(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zceil x
H:(x < 0)%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zceil x
H:(0 <= x)%R

(/ 2 * ulp x <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zceil x
H:(x < 0)%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zceil x
H:(0 <= x)%R

(0 <= / 2)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zceil x
H:(0 <= x)%R
(ulp x <= ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zceil x
H:(x < 0)%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zceil x
H:(0 <= x)%R

(ulp x <= ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zceil x
H:(x < 0)%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zceil x
H:(0 <= x)%R

(x <= round beta fexp (Znearest choice) x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zceil x
H:(x < 0)%R
(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zceil x
H:(x < 0)%R

(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp (Znearest choice) x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zceil x
H:(x < 0)%R

(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (- round beta fexp Zceil x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zceil x
H:(x < 0)%R

(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (round beta fexp Zfloor (- x)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zceil x
H:(x < 0)%R

(Rabs (round beta fexp (Znearest choice) x - x) <= / 2 * ulp (- x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zceil x
H:(x < 0)%R
(0 <= - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zceil x
H:(x < 0)%R

(Rabs (round beta fexp (Znearest choice) (- - x) - - - x) <= / 2 * ulp (- x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zceil x
H:(x < 0)%R
(0 <= - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zceil x
H:(x < 0)%R

(Rabs (- round beta fexp (Znearest (fun t : Z => negb (choice (- (t + 1))%Z))) (- x) - - - x) <= / 2 * ulp (- x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zceil x
H:(x < 0)%R
(0 <= - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zceil x
H:(x < 0)%R

(Rabs (- round beta fexp (Znearest (fun t : Z => negb (choice (- (t + 1))%Z))) (- x) + - - - x) <= / 2 * ulp (- x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zceil x
H:(x < 0)%R
(0 <= - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zceil x
H:(x < 0)%R

(Rabs (round beta fexp (Znearest (fun t : Z => negb (choice (- (t + 1))%Z))) (- x) + - - x) <= / 2 * ulp (- x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zceil x
H:(x < 0)%R
(0 <= - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zceil x
H:(x < 0)%R

(0 <= - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zceil x
H:(x < 0)%R

(- 0 <= - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice:Z -> bool
x:R
Hfx:round beta fexp (Znearest choice) x <> 0%R
Hx:round beta fexp (Znearest choice) x = round beta fexp Zceil x
H:(x < 0)%R

(x <= 0)%R
now apply Rlt_le. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x y : R, F x -> F y -> (x <= y)%R -> (pred x <= pred y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x y : R, F x -> F y -> (x <= y)%R -> (pred x <= pred y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
Hxy:(x < y)%R

(pred x <= pred y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
y:R
Fx, Fy:F y
(pred y <= pred y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
Hxy:(x < y)%R

(pred x <= pred y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
Hxy:(x < y)%R

F (pred x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
Hxy:(x < y)%R
(pred x < y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
Hxy:(x < y)%R

(pred x < y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
Hxy:(x < y)%R

(pred x <= x)%R
apply pred_le_id. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x y : R, F x -> F y -> (x <= y)%R -> (succ x <= succ y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x y : R, F x -> F y -> (x <= y)%R -> (succ x <= succ y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
Hxy:(x <= y)%R

(succ x <= succ y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
Hxy:(x <= y)%R

(- succ y <= - succ x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
Hxy:(x <= y)%R

(pred (- y) <= pred (- x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
Hxy:(x <= y)%R

F (- y)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
Hxy:(x <= y)%R
F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
Hxy:(x <= y)%R
(- y <= - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
Hxy:(x <= y)%R

F (- x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
Hxy:(x <= y)%R
(- y <= - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
Hxy:(x <= y)%R

(- y <= - x)%R
now apply Ropp_le_contravar. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x y : R, F x -> F y -> (pred x <= pred y)%R -> (x <= y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x y : R, F x -> F y -> (pred x <= pred y)%R -> (x <= y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
Hxy:(pred x <= pred y)%R

(x <= y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
Hxy:(pred x <= pred y)%R

(succ (pred x) <= succ (pred y))%R
apply succ_le; trivial; now apply generic_format_pred. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x y : R, F x -> F y -> (succ x <= succ y)%R -> (x <= y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x y : R, F x -> F y -> (succ x <= succ y)%R -> (x <= y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
Hxy:(succ x <= succ y)%R

(x <= y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
Hxy:(succ x <= succ y)%R

(pred (succ x) <= pred (succ y))%R
apply pred_le; trivial; now apply generic_format_succ. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x y : R, F x -> F y -> (x < y)%R -> (pred x < pred y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x y : R, F x -> F y -> (x < y)%R -> (pred x < pred y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
Hxy:(x < y)%R

(pred x < pred y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
Hxy:(x < y)%R

~ (pred y <= pred x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
Hxy:(x < y)%R
H:(pred y <= pred x)%R

False
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
Hxy:(x < y)%R
H:(pred y <= pred x)%R

(y <= x)%R
now apply pred_le_inv. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x y : R, F x -> F y -> (x < y)%R -> (succ x < succ y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x y : R, F x -> F y -> (x < y)%R -> (succ x < succ y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
Hxy:(x < y)%R

(succ x < succ y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
Hxy:(x < y)%R

~ (succ y <= succ x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
Hxy:(x < y)%R
H:(succ y <= succ x)%R

False
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fx:F x
Fy:F y
Hxy:(x < y)%R
H:(succ y <= succ x)%R

(y <= x)%R
now apply succ_le_inv. Qed.
Adding ulp is a, somewhat reasonable, overapproximation of succ.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

Monotone_exp fexp -> forall x : R, (succ x <= x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

Monotone_exp fexp -> forall x : R, (succ x <= x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R

(succ x <= x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Nx:(x < 0)%R

(succ x <= x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Nx:(x < 0)%R

(succ x <= - (- x - ulp x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Nx:(x < 0)%R

(- pred_pos (- x) <= - (- x - ulp (- x)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Nx:(x < 0)%R

(- x - ulp (- x) <= (if Req_bool (- x) (bpow (mag beta (- x) - 1)) then - x - bpow (fexp (mag beta (- x) - 1)) else - x - ulp (- x)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)

(- x - ulp (- x) <= (if Req_bool (- x) (bpow (mag beta (- x) - 1)) then - x - bpow (fexp (mag beta (- x) - 1)) else - x - ulp (- x)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Nx:(x < 0)%R
Hx:(- x)%R <> bpow (mag beta (- x) - 1)
(- x - ulp (- x) <= (if Req_bool (- x) (bpow (mag beta (- x) - 1)) then - x - bpow (fexp (mag beta (- x) - 1)) else - x - ulp (- x)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)

(- x - ulp (- x) <= (if Req_bool (- x) (bpow (mag beta (- x) - 1)) then - x - bpow (fexp (mag beta (- x) - 1)) else - x - ulp (- x)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)

(- x - ulp (- x) <= - x - bpow (fexp (mag beta (- x) - 1)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)

(bpow (fexp (mag beta (- x) - 1)) <= ulp (- x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)

(bpow (fexp (mag beta (- x) - 1)) <= bpow (cexp beta fexp (- x)))%R
apply bpow_le, Mexp; lia.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Nx:(x < 0)%R
Hx:(- x)%R <> bpow (mag beta (- x) - 1)

(- x - ulp (- x) <= (if Req_bool (- x) (bpow (mag beta (- x) - 1)) then - x - bpow (fexp (mag beta (- x) - 1)) else - x - ulp (- x)))%R
now rewrite (Req_bool_false _ _ Hx); right. Qed.
And it also lies in the format.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

Monotone_exp fexp -> forall x : R, F x -> F (x + ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

Monotone_exp fexp -> forall x : R, F x -> F (x + ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F x

F (x + ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F x
Px:(0 <= x)%R

F (x + ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F x
Nx:(x < 0)%R
F (x + ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F x
Px:(0 <= x)%R

F (x + ulp x)
now rewrite <-(succ_eq_pos _ Px); apply generic_format_succ.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F x
Nx:(x < 0)%R

F (x + ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R

F (x + ulp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R

F (- (- x - ulp x))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R

F (- x - ulp (- x))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)

F (- x - ulp (- x))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R <> bpow (mag beta (- x) - 1)
F (- x - ulp (- x))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)

F (- x - ulp (- x))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)

F (- x - bpow (cexp beta fexp (- x)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)

F (bpow (mag beta (- x) - 1) - bpow (cexp beta fexp (- x)))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)

F (bpow (mag beta (- x) - 1) - bpow (fexp (mag beta (- x))))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)

F (bpow (e - 1) - bpow (fexp e))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)

(fexp e < e)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z
F (bpow (e - 1) - bpow (fexp e))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)

(fexp e < e)%Z
now apply mag_generic_gt; [|lra|].
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z

F (bpow (e - 1) - bpow (fexp e))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z

F (bpow (e - 1 - fexp e + fexp e) - bpow (fexp e))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z

F (bpow (e - 1 - fexp e) * bpow (fexp e) - bpow (fexp e))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z
m:=bpow (e - 1 - fexp e):R

F (m * bpow (fexp e) - bpow (fexp e))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z
m:=bpow (e - 1 - fexp e):R

F ((m - 1) * bpow (fexp e))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z
m:=bpow (e - 1 - fexp e):R

(e - 1 - fexp e)%Z = 0%Z -> F ((m - 1) * bpow (fexp e))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z
m:=bpow (e - 1 - fexp e):R
forall p : positive, (e - 1 - fexp e)%Z = Z.pos p -> F ((m - 1) * bpow (fexp e))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z
m:=bpow (e - 1 - fexp e):R
forall p : positive, (e - 1 - fexp e)%Z = Z.neg p -> F ((m - 1) * bpow (fexp e))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z
m:=bpow (e - 1 - fexp e):R

(e - 1 - fexp e)%Z = 0%Z -> F ((m - 1) * bpow (fexp e))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z
m:=bpow (e - 1 - fexp e):R
He:(e - 1 - fexp e)%Z = 0%Z

F (0 * bpow (fexp e))
rewrite Rmult_0_l; apply generic_format_0.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z
m:=bpow (e - 1 - fexp e):R

forall p : positive, (e - 1 - fexp e)%Z = Z.pos p -> F ((m - 1) * bpow (fexp e))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z
m:=bpow (e - 1 - fexp e):R
forall p : positive, (e - 1 - fexp e)%Z = Z.neg p -> F ((m - 1) * bpow (fexp e))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z
m:=bpow (e - 1 - fexp e):R

forall p : positive, (e - 1 - fexp e)%Z = Z.pos p -> F ((m - 1) * bpow (fexp e))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z
m:=bpow (e - 1 - fexp e):R
p:positive
Hp:(e - 1 - fexp e)%Z = Z.pos p

F ((IZR (Z.pow_pos beta p) - 1) * bpow (fexp e))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z
m:=bpow (e - 1 - fexp e):R
p:positive
Hp:(e - 1 - fexp e)%Z = Z.pos p
f:={| Fnum := Z.pow_pos beta p - 1; Fexp := fexp e |} : float beta:float beta

F ((IZR (Z.pow_pos beta p) - 1) * bpow (fexp e))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z
m:=bpow (e - 1 - fexp e):R
p:positive
Hp:(e - 1 - fexp e)%Z = Z.pos p
f:={| Fnum := Z.pow_pos beta p - 1; Fexp := fexp e |} : float beta:float beta

F2R f = ((IZR (Z.pow_pos beta p) - 1) * bpow (fexp e))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z
m:=bpow (e - 1 - fexp e):R
p:positive
Hp:(e - 1 - fexp e)%Z = Z.pos p
f:={| Fnum := Z.pow_pos beta p - 1; Fexp := fexp e |} : float beta:float beta
Hm':((IZR (Z.pow_pos beta p) - 1) * bpow (fexp e))%R <> 0%R
(cexp beta fexp ((IZR (Z.pow_pos beta p) - 1) * bpow (fexp e)) <= fexp e)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z
m:=bpow (e - 1 - fexp e):R
p:positive
Hp:(e - 1 - fexp e)%Z = Z.pos p
f:={| Fnum := Z.pow_pos beta p - 1; Fexp := fexp e |} : float beta:float beta

F2R f = ((IZR (Z.pow_pos beta p) - 1) * bpow (fexp e))%R
now unfold Defs.F2R; simpl; rewrite minus_IZR.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z
m:=bpow (e - 1 - fexp e):R
p:positive
Hp:(e - 1 - fexp e)%Z = Z.pos p
f:={| Fnum := Z.pow_pos beta p - 1; Fexp := fexp e |} : float beta:float beta
Hm':((IZR (Z.pow_pos beta p) - 1) * bpow (fexp e))%R <> 0%R

(cexp beta fexp ((IZR (Z.pow_pos beta p) - 1) * bpow (fexp e)) <= fexp e)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z
m:=bpow (e - 1 - fexp e):R
p:positive
Hp:(e - 1 - fexp e)%Z = Z.pos p
f:={| Fnum := Z.pow_pos beta p - 1; Fexp := fexp e |} : float beta:float beta
Hm':((IZR (Z.pow_pos beta p) - 1) * bpow (fexp e))%R <> 0%R

(fexp (mag beta ((IZR (Z.pow_pos beta p) - 1) * bpow (fexp e))) <= fexp e)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z
m:=bpow (e - 1 - fexp e):R
p:positive
Hp:(e - 1 - fexp e)%Z = Z.pos p
f:={| Fnum := Z.pow_pos beta p - 1; Fexp := fexp e |} : float beta:float beta
Hm':((IZR (Z.pow_pos beta p) - 1) * bpow (fexp e))%R <> 0%R

(fexp (mag beta ((bpow (Z.pos p) - 1) * bpow (fexp e))) <= fexp e)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z
m:=bpow (e - 1 - fexp e):R
p:positive
Hp:(e - 1 - fexp e)%Z = Z.pos p
f:={| Fnum := Z.pow_pos beta p - 1; Fexp := fexp e |} : float beta:float beta
Hm':((IZR (Z.pow_pos beta p) - 1) * bpow (fexp e))%R <> 0%R

(fexp (mag beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e))) <= fexp e)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z
m:=bpow (e - 1 - fexp e):R
p:positive
Hp:(e - 1 - fexp e)%Z = Z.pos p
f:={| Fnum := Z.pow_pos beta p - 1; Fexp := fexp e |} : float beta:float beta
Hm':((IZR (Z.pow_pos beta p) - 1) * bpow (fexp e))%R <> 0%R
He:(1 <= e - 1 - fexp e)%Z

(fexp (mag beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e))) <= fexp e)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z
m:=bpow (e - 1 - fexp e):R
p:positive
Hp:(e - 1 - fexp e)%Z = Z.pos p
f:={| Fnum := Z.pow_pos beta p - 1; Fexp := fexp e |} : float beta:float beta
Hm':((IZR (Z.pow_pos beta p) - 1) * bpow (fexp e))%R <> 0%R
He:(1 <= e - 1 - fexp e)%Z
e':=mag beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e)):mag_prop beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e))

(fexp e' <= fexp e)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z
m:=bpow (e - 1 - fexp e):R
p:positive
Hp:(e - 1 - fexp e)%Z = Z.pos p
f:={| Fnum := Z.pow_pos beta p - 1; Fexp := fexp e |} : float beta:float beta
Hm':((IZR (Z.pow_pos beta p) - 1) * bpow (fexp e))%R <> 0%R
He:(1 <= e - 1 - fexp e)%Z
e':=mag beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e)):mag_prop beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e))

e' = (e - 1)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z
m:=bpow (e - 1 - fexp e):R
p:positive
Hp:(e - 1 - fexp e)%Z = Z.pos p
f:={| Fnum := Z.pow_pos beta p - 1; Fexp := fexp e |} : float beta:float beta
Hm':((IZR (Z.pow_pos beta p) - 1) * bpow (fexp e))%R <> 0%R
He:(1 <= e - 1 - fexp e)%Z
e':=mag beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e)):mag_prop beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e))

(bpow (e - 1 - 1) <= Rabs ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e)) < bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z
m:=bpow (e - 1 - fexp e):R
p:positive
Hp:(e - 1 - fexp e)%Z = Z.pos p
f:={| Fnum := Z.pow_pos beta p - 1; Fexp := fexp e |} : float beta:float beta
Hm':((IZR (Z.pow_pos beta p) - 1) * bpow (fexp e))%R <> 0%R
He:(1 <= e - 1 - fexp e)%Z
e':=mag beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e)):mag_prop beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e))

(bpow (e - 1 - 1) <= Rabs (bpow (e - 1 - fexp e) - 1) * bpow (fexp e) < bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z
m:=bpow (e - 1 - fexp e):R
p:positive
Hp:(e - 1 - fexp e)%Z = Z.pos p
f:={| Fnum := Z.pow_pos beta p - 1; Fexp := fexp e |} : float beta:float beta
Hm':((IZR (Z.pow_pos beta p) - 1) * bpow (fexp e))%R <> 0%R
He:(1 <= e - 1 - fexp e)%Z
e':=mag beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e)):mag_prop beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e))

(bpow (e - 1 - 1) <= (bpow (e - 1 - fexp e) - 1) * bpow (fexp e) < bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z
m:=bpow (e - 1 - fexp e):R
p:positive
Hp:(e - 1 - fexp e)%Z = Z.pos p
f:={| Fnum := Z.pow_pos beta p - 1; Fexp := fexp e |} : float beta:float beta
Hm':((IZR (Z.pow_pos beta p) - 1) * bpow (fexp e))%R <> 0%R
He:(1 <= e - 1 - fexp e)%Z
e':=mag beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e)):mag_prop beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e))

(0 < IZR beta)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z
m:=bpow (e - 1 - fexp e):R
p:positive
Hp:(e - 1 - fexp e)%Z = Z.pos p
f:={| Fnum := Z.pow_pos beta p - 1; Fexp := fexp e |} : float beta:float beta
Hm':((IZR (Z.pow_pos beta p) - 1) * bpow (fexp e))%R <> 0%R
He:(1 <= e - 1 - fexp e)%Z
e':=mag beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e)):mag_prop beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e))
beta_pos:(0 < IZR beta)%R
(bpow (e - 1 - 1) <= (bpow (e - 1 - fexp e) - 1) * bpow (fexp e) < bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z
m:=bpow (e - 1 - fexp e):R
p:positive
Hp:(e - 1 - fexp e)%Z = Z.pos p
f:={| Fnum := Z.pow_pos beta p - 1; Fexp := fexp e |} : float beta:float beta
Hm':((IZR (Z.pow_pos beta p) - 1) * bpow (fexp e))%R <> 0%R
He:(1 <= e - 1 - fexp e)%Z
e':=mag beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e)):mag_prop beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e))

(0 < IZR beta)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z
m:=bpow (e - 1 - fexp e):R
p:positive
Hp:(e - 1 - fexp e)%Z = Z.pos p
f:={| Fnum := Z.pow_pos beta p - 1; Fexp := fexp e |} : float beta:float beta
Hm':((IZR (Z.pow_pos beta p) - 1) * bpow (fexp e))%R <> 0%R
He:(1 <= e - 1 - fexp e)%Z
e':=mag beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e)):mag_prop beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e))

(2 <= IZR beta)%R
apply IZR_le, Z.leb_le, radix_prop.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z
m:=bpow (e - 1 - fexp e):R
p:positive
Hp:(e - 1 - fexp e)%Z = Z.pos p
f:={| Fnum := Z.pow_pos beta p - 1; Fexp := fexp e |} : float beta:float beta
Hm':((IZR (Z.pow_pos beta p) - 1) * bpow (fexp e))%R <> 0%R
He:(1 <= e - 1 - fexp e)%Z
e':=mag beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e)):mag_prop beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e))
beta_pos:(0 < IZR beta)%R

(bpow (e - 1 - 1) <= (bpow (e - 1 - fexp e) - 1) * bpow (fexp e) < bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z
m:=bpow (e - 1 - fexp e):R
p:positive
Hp:(e - 1 - fexp e)%Z = Z.pos p
f:={| Fnum := Z.pow_pos beta p - 1; Fexp := fexp e |} : float beta:float beta
Hm':((IZR (Z.pow_pos beta p) - 1) * bpow (fexp e))%R <> 0%R
He:(1 <= e - 1 - fexp e)%Z
e':=mag beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e)):mag_prop beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e))
beta_pos:(0 < IZR beta)%R

(bpow (e - 1 - 1) <= (bpow (e - 1 - fexp e) - 1) * bpow (fexp e))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z
m:=bpow (e - 1 - fexp e):R
p:positive
Hp:(e - 1 - fexp e)%Z = Z.pos p
f:={| Fnum := Z.pow_pos beta p - 1; Fexp := fexp e |} : float beta:float beta
Hm':((IZR (Z.pow_pos beta p) - 1) * bpow (fexp e))%R <> 0%R
He:(1 <= e - 1 - fexp e)%Z
e':=mag beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e)):mag_prop beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e))
beta_pos:(0 < IZR beta)%R
((bpow (e - 1 - fexp e) - 1) * bpow (fexp e) < bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z
m:=bpow (e - 1 - fexp e):R
p:positive
Hp:(e - 1 - fexp e)%Z = Z.pos p
f:={| Fnum := Z.pow_pos beta p - 1; Fexp := fexp e |} : float beta:float beta
Hm':((IZR (Z.pow_pos beta p) - 1) * bpow (fexp e))%R <> 0%R
He:(1 <= e - 1 - fexp e)%Z
e':=mag beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e)):mag_prop beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e))
beta_pos:(0 < IZR beta)%R

(bpow (e - 1 - 1) <= (bpow (e - 1 - fexp e) - 1) * bpow (fexp e))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z
m:=bpow (e - 1 - fexp e):R
p:positive
Hp:(e - 1 - fexp e)%Z = Z.pos p
f:={| Fnum := Z.pow_pos beta p - 1; Fexp := fexp e |} : float beta:float beta
Hm':((IZR (Z.pow_pos beta p) - 1) * bpow (fexp e))%R <> 0%R
He:(1 <= e - 1 - fexp e)%Z
e':=mag beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e)):mag_prop beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e))
beta_pos:(0 < IZR beta)%R

(bpow (e - 1 - fexp e + -1 + fexp e) <= (bpow (e - 1 - fexp e) - 1) * bpow (fexp e))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z
m:=bpow (e - 1 - fexp e):R
p:positive
Hp:(e - 1 - fexp e)%Z = Z.pos p
f:={| Fnum := Z.pow_pos beta p - 1; Fexp := fexp e |} : float beta:float beta
Hm':((IZR (Z.pow_pos beta p) - 1) * bpow (fexp e))%R <> 0%R
He:(1 <= e - 1 - fexp e)%Z
e':=mag beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e)):mag_prop beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e))
beta_pos:(0 < IZR beta)%R

(bpow (e - 1 - fexp e + -1) * bpow (fexp e) <= (bpow (e - 1 - fexp e) - 1) * bpow (fexp e))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z
m:=bpow (e - 1 - fexp e):R
p:positive
Hp:(e - 1 - fexp e)%Z = Z.pos p
f:={| Fnum := Z.pow_pos beta p - 1; Fexp := fexp e |} : float beta:float beta
Hm':((IZR (Z.pow_pos beta p) - 1) * bpow (fexp e))%R <> 0%R
He:(1 <= e - 1 - fexp e)%Z
e':=mag beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e)):mag_prop beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e))
beta_pos:(0 < IZR beta)%R

(bpow (e - 1 - fexp e + -1) <= bpow (e - 1 - fexp e) - 1)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z
m:=bpow (e - 1 - fexp e):R
p:positive
Hp:(e - 1 - fexp e)%Z = Z.pos p
f:={| Fnum := Z.pow_pos beta p - 1; Fexp := fexp e |} : float beta:float beta
Hm':((IZR (Z.pow_pos beta p) - 1) * bpow (fexp e))%R <> 0%R
He:(1 <= e - 1 - fexp e)%Z
e':=mag beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e)):mag_prop beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e))
beta_pos:(0 < IZR beta)%R

(bpow (e - 1 - fexp e) * / IZR (beta * 1) <= bpow (e - 1 - fexp e) - 1)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z
m:=bpow (e - 1 - fexp e):R
p:positive
Hp:(e - 1 - fexp e)%Z = Z.pos p
f:={| Fnum := Z.pow_pos beta p - 1; Fexp := fexp e |} : float beta:float beta
Hm':((IZR (Z.pow_pos beta p) - 1) * bpow (fexp e))%R <> 0%R
He:(1 <= e - 1 - fexp e)%Z
e':=mag beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e)):mag_prop beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e))
beta_pos:(0 < IZR beta)%R

(bpow (e - 1 - fexp e) * / IZR beta <= bpow (e - 1 - fexp e) - 1)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z
m:=bpow (e - 1 - fexp e):R
p:positive
Hp:(e - 1 - fexp e)%Z = Z.pos p
f:={| Fnum := Z.pow_pos beta p - 1; Fexp := fexp e |} : float beta:float beta
Hm':((IZR (Z.pow_pos beta p) - 1) * bpow (fexp e))%R <> 0%R
He:(1 <= e - 1 - fexp e)%Z
e':=mag beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e)):mag_prop beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e))
beta_pos:(0 < IZR beta)%R

(bpow (e - 1 - fexp e) * / IZR beta * IZR beta <= (bpow (e - 1 - fexp e) - 1) * IZR beta)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z
m:=bpow (e - 1 - fexp e):R
p:positive
Hp:(e - 1 - fexp e)%Z = Z.pos p
f:={| Fnum := Z.pow_pos beta p - 1; Fexp := fexp e |} : float beta:float beta
Hm':((IZR (Z.pow_pos beta p) - 1) * bpow (fexp e))%R <> 0%R
He:(1 <= e - 1 - fexp e)%Z
e':=mag beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e)):mag_prop beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e))
beta_pos:(0 < IZR beta)%R

(bpow (e - 1 - fexp e) <= (bpow (e - 1 - fexp e) - 1) * IZR beta)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z
m:=bpow (e - 1 - fexp e):R
p:positive
Hp:(e - 1 - fexp e)%Z = Z.pos p
f:={| Fnum := Z.pow_pos beta p - 1; Fexp := fexp e |} : float beta:float beta
Hm':((IZR (Z.pow_pos beta p) - 1) * bpow (fexp e))%R <> 0%R
He:(1 <= e - 1 - fexp e)%Z
e':=mag beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e)):mag_prop beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e))
beta_pos:(0 < IZR beta)%R

(bpow (e - 1 - fexp e) + IZR beta <= bpow (e - 1 - fexp e) * IZR beta)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z
m:=bpow (e - 1 - fexp e):R
p:positive
Hp:(e - 1 - fexp e)%Z = Z.pos p
f:={| Fnum := Z.pow_pos beta p - 1; Fexp := fexp e |} : float beta:float beta
Hm':((IZR (Z.pow_pos beta p) - 1) * bpow (fexp e))%R <> 0%R
He:(1 <= e - 1 - fexp e)%Z
e':=mag beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e)):mag_prop beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e))
beta_pos:(0 < IZR beta)%R

(bpow (e - 1 - fexp e) + IZR beta <= 2 * bpow (e - 1 - fexp e))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z
m:=bpow (e - 1 - fexp e):R
p:positive
Hp:(e - 1 - fexp e)%Z = Z.pos p
f:={| Fnum := Z.pow_pos beta p - 1; Fexp := fexp e |} : float beta:float beta
Hm':((IZR (Z.pow_pos beta p) - 1) * bpow (fexp e))%R <> 0%R
He:(1 <= e - 1 - fexp e)%Z
e':=mag beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e)):mag_prop beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e))
beta_pos:(0 < IZR beta)%R
(2 * bpow (e - 1 - fexp e) <= bpow (e - 1 - fexp e) * IZR beta)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z
m:=bpow (e - 1 - fexp e):R
p:positive
Hp:(e - 1 - fexp e)%Z = Z.pos p
f:={| Fnum := Z.pow_pos beta p - 1; Fexp := fexp e |} : float beta:float beta
Hm':((IZR (Z.pow_pos beta p) - 1) * bpow (fexp e))%R <> 0%R
He:(1 <= e - 1 - fexp e)%Z
e':=mag beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e)):mag_prop beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e))
beta_pos:(0 < IZR beta)%R

(bpow (e - 1 - fexp e) + IZR beta <= 2 * bpow (e - 1 - fexp e))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z
m:=bpow (e - 1 - fexp e):R
p:positive
Hp:(e - 1 - fexp e)%Z = Z.pos p
f:={| Fnum := Z.pow_pos beta p - 1; Fexp := fexp e |} : float beta:float beta
Hm':((IZR (Z.pow_pos beta p) - 1) * bpow (fexp e))%R <> 0%R
He:(1 <= e - 1 - fexp e)%Z
e':=mag beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e)):mag_prop beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e))
beta_pos:(0 < IZR beta)%R

(bpow (e - 1 - fexp e) + IZR beta <= bpow (e - 1 - fexp e) + bpow (e - 1 - fexp e))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z
m:=bpow (e - 1 - fexp e):R
p:positive
Hp:(e - 1 - fexp e)%Z = Z.pos p
f:={| Fnum := Z.pow_pos beta p - 1; Fexp := fexp e |} : float beta:float beta
Hm':((IZR (Z.pow_pos beta p) - 1) * bpow (fexp e))%R <> 0%R
He:(1 <= e - 1 - fexp e)%Z
e':=mag beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e)):mag_prop beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e))
beta_pos:(0 < IZR beta)%R

(IZR beta <= bpow (e - 1 - fexp e))%R
rewrite <-bpow_1; apply bpow_le; lia.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z
m:=bpow (e - 1 - fexp e):R
p:positive
Hp:(e - 1 - fexp e)%Z = Z.pos p
f:={| Fnum := Z.pow_pos beta p - 1; Fexp := fexp e |} : float beta:float beta
Hm':((IZR (Z.pow_pos beta p) - 1) * bpow (fexp e))%R <> 0%R
He:(1 <= e - 1 - fexp e)%Z
e':=mag beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e)):mag_prop beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e))
beta_pos:(0 < IZR beta)%R

(2 * bpow (e - 1 - fexp e) <= bpow (e - 1 - fexp e) * IZR beta)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z
m:=bpow (e - 1 - fexp e):R
p:positive
Hp:(e - 1 - fexp e)%Z = Z.pos p
f:={| Fnum := Z.pow_pos beta p - 1; Fexp := fexp e |} : float beta:float beta
Hm':((IZR (Z.pow_pos beta p) - 1) * bpow (fexp e))%R <> 0%R
He:(1 <= e - 1 - fexp e)%Z
e':=mag beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e)):mag_prop beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e))
beta_pos:(0 < IZR beta)%R

(2 <= IZR beta)%R
apply IZR_le, Z.leb_le, radix_prop.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z
m:=bpow (e - 1 - fexp e):R
p:positive
Hp:(e - 1 - fexp e)%Z = Z.pos p
f:={| Fnum := Z.pow_pos beta p - 1; Fexp := fexp e |} : float beta:float beta
Hm':((IZR (Z.pow_pos beta p) - 1) * bpow (fexp e))%R <> 0%R
He:(1 <= e - 1 - fexp e)%Z
e':=mag beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e)):mag_prop beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e))
beta_pos:(0 < IZR beta)%R

((bpow (e - 1 - fexp e) - 1) * bpow (fexp e) < bpow (e - 1))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z
m:=bpow (e - 1 - fexp e):R
p:positive
Hp:(e - 1 - fexp e)%Z = Z.pos p
f:={| Fnum := Z.pow_pos beta p - 1; Fexp := fexp e |} : float beta:float beta
Hm':((IZR (Z.pow_pos beta p) - 1) * bpow (fexp e))%R <> 0%R
He:(1 <= e - 1 - fexp e)%Z
e':=mag beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e)):mag_prop beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e))
beta_pos:(0 < IZR beta)%R

((bpow (e - 1 - fexp e) - 1) * bpow (fexp e) * bpow (- fexp e) < bpow (e - 1) * bpow (- fexp e))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z
m:=bpow (e - 1 - fexp e):R
p:positive
Hp:(e - 1 - fexp e)%Z = Z.pos p
f:={| Fnum := Z.pow_pos beta p - 1; Fexp := fexp e |} : float beta:float beta
Hm':((IZR (Z.pow_pos beta p) - 1) * bpow (fexp e))%R <> 0%R
He:(1 <= e - 1 - fexp e)%Z
e':=mag beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e)):mag_prop beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e))
beta_pos:(0 < IZR beta)%R

((bpow (e - 1 - fexp e) - 1) * bpow (fexp e + - fexp e) < bpow (e - 1 + - fexp e))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z
m:=bpow (e - 1 - fexp e):R
p:positive
Hp:(e - 1 - fexp e)%Z = Z.pos p
f:={| Fnum := Z.pow_pos beta p - 1; Fexp := fexp e |} : float beta:float beta
Hm':((IZR (Z.pow_pos beta p) - 1) * bpow (fexp e))%R <> 0%R
He:(1 <= e - 1 - fexp e)%Z
e':=mag beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e)):mag_prop beta ((bpow (e - 1 - fexp e) - 1) * bpow (fexp e))
beta_pos:(0 < IZR beta)%R

((bpow (e - 1 - fexp e) - 1) * 1 < bpow (e - 1 + - fexp e))%R
rewrite Rmult_1_r; unfold Zminus; lra.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R = bpow (mag beta (- x) - 1)
e:=mag beta (- x):mag_prop beta (- x)
Hfe:(fexp e < e)%Z
m:=bpow (e - 1 - fexp e):R

forall p : positive, (e - 1 - fexp e)%Z = Z.neg p -> F ((m - 1) * bpow (fexp e))
intros p Hp; exfalso; lia.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R <> bpow (mag beta (- x) - 1)

F (- x - ulp (- x))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R <> bpow (mag beta (- x) - 1)

F (pred_pos (- x))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R <> bpow (mag beta (- x) - 1)
pred_pos (- x) = (- x - ulp (- x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R <> bpow (mag beta (- x) - 1)

F (pred_pos (- x))
now apply generic_format_pred_pos; [|lra].
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Mexp:Monotone_exp fexp
x:R
Fx:F (- x)
Nx:(x < 0)%R
Hx:(- x)%R <> bpow (mag beta (- x) - 1)

pred_pos (- x) = (- x - ulp (- x))%R
now unfold pred_pos; rewrite Req_bool_false. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x y : R, F y -> (y < round beta fexp Zceil x)%R -> (y <= round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x y : R, F y -> (y < round beta fexp Zceil x)%R -> (y <= round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fy:F y
Hlt:(y < round beta fexp Zceil x)%R

(y <= round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fy:F y
Hlt:(y < round beta fexp Zceil x)%R

(y <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fy:F y
Hlt:(y < round beta fexp Zceil x)%R

~ (x < y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fy:F y
Hlt:(x < y)%R

~ (y < round beta fexp Zceil x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fy:F y
Hlt:(x < y)%R

(round beta fexp Zceil x <= y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fy:F y
Hlt:(x < y)%R

(x <= y)%R
now apply Rlt_le. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x y : R, F y -> (round beta fexp Zfloor x < y)%R -> (round beta fexp Zceil x <= y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x y : R, F y -> (round beta fexp Zfloor x < y)%R -> (round beta fexp Zceil x <= y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fy:F y
Hlt:(round beta fexp Zfloor x < y)%R

(round beta fexp Zceil x <= y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fy:F y
Hlt:(round beta fexp Zfloor x < y)%R

(x <= y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fy:F y
Hlt:(round beta fexp Zfloor x < y)%R

~ (y < x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fy:F y
Hlt:(y < x)%R

~ (round beta fexp Zfloor x < y)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fy:F y
Hlt:(y < x)%R

(y <= round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, y:R
Fy:F y
Hlt:(y < x)%R

(y <= x)%R
now apply Rlt_le. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, (pred (round beta fexp Zceil x) <= round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, (pred (round beta fexp Zceil x) <= round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R

(pred (round beta fexp Zceil x) <= round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x

(pred (round beta fexp Zceil x) <= round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
(pred (round beta fexp Zceil x) <= round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x

(pred x <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
(pred (round beta fexp Zceil x) <= round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x

(pred (round beta fexp Zceil x) <= round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
Zx:round beta fexp Zceil x = 0%R

(pred (round beta fexp Zceil x) <= round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
Zx:round beta fexp Zceil x <> 0%R
(pred (round beta fexp Zceil x) <= round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
Zx:round beta fexp Zceil x = 0%R

(- succ 0 <= round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
Zx:round beta fexp Zceil x <> 0%R
(pred (round beta fexp Zceil x) <= round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
Zx:round beta fexp Zceil x = 0%R

(- (0 + ulp 0) <= round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
Zx:round beta fexp Zceil x <> 0%R
(pred (round beta fexp Zceil x) <= round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
Zx:round beta fexp Zceil x = 0%R

(- match negligible_exp with | Some n => bpow (fexp n) | None => 0 end <= round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
Zx:round beta fexp Zceil x <> 0%R
(pred (round beta fexp Zceil x) <= round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
Zx:round beta fexp Zceil x = 0%R

negligible_exp = None /\ (forall n : Z, (fexp n < n)%Z) -> (- match negligible_exp with | Some n => bpow (fexp n) | None => 0 end <= round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
Zx:round beta fexp Zceil x = 0%R
(exists n : Z, negligible_exp = Some n /\ (n <= fexp n)%Z) -> (- match negligible_exp with | Some n => bpow (fexp n) | None => 0 end <= round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
Zx:round beta fexp Zceil x <> 0%R
(pred (round beta fexp Zceil x) <= round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
Zx:round beta fexp Zceil x = 0%R
H1:negligible_exp = None
H2:forall n : Z, (fexp n < n)%Z

(- match negligible_exp with | Some n => bpow (fexp n) | None => 0 end <= round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
Zx:round beta fexp Zceil x = 0%R
(exists n : Z, negligible_exp = Some n /\ (n <= fexp n)%Z) -> (- match negligible_exp with | Some n => bpow (fexp n) | None => 0 end <= round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
Zx:round beta fexp Zceil x <> 0%R
(pred (round beta fexp Zceil x) <= round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
H1:negligible_exp = None
H2:forall n : Z, (fexp n < n)%Z

x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
Zx:round beta fexp Zceil x = 0%R
(exists n : Z, negligible_exp = Some n /\ (n <= fexp n)%Z) -> (- match negligible_exp with | Some n => bpow (fexp n) | None => 0 end <= round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
Zx:round beta fexp Zceil x <> 0%R
(pred (round beta fexp Zceil x) <= round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
Zx:round beta fexp Zceil x = 0%R

(exists n : Z, negligible_exp = Some n /\ (n <= fexp n)%Z) -> (- match negligible_exp with | Some n => bpow (fexp n) | None => 0 end <= round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
Zx:round beta fexp Zceil x <> 0%R
(pred (round beta fexp Zceil x) <= round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
Zx:round beta fexp Zceil x = 0%R
n:Z
H1:negligible_exp = Some n
Hn:(n <= fexp n)%Z

(- bpow (fexp n) <= round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
Zx:round beta fexp Zceil x <> 0%R
(pred (round beta fexp Zceil x) <= round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
Zx:round beta fexp Zceil x = 0%R
n:Z
H1:negligible_exp = Some n
Hn:(n <= fexp n)%Z
K:(round beta fexp Zfloor x < - bpow (fexp n))%R

(- bpow (fexp n) <= round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
Zx:round beta fexp Zceil x <> 0%R
(pred (round beta fexp Zceil x) <= round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
Zx:round beta fexp Zceil x = 0%R
n:Z
H1:negligible_exp = Some n
Hn:(n <= fexp n)%Z
K:(round beta fexp Zfloor x < - bpow (fexp n))%R

~ (round beta fexp Zceil x <= - bpow (fexp n))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
Zx:round beta fexp Zceil x = 0%R
n:Z
H1:negligible_exp = Some n
Hn:(n <= fexp n)%Z
K:(round beta fexp Zfloor x < - bpow (fexp n))%R
(round beta fexp Zceil x <= - bpow (fexp n))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
Zx:round beta fexp Zceil x <> 0%R
(pred (round beta fexp Zceil x) <= round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
Zx:round beta fexp Zceil x = 0%R
n:Z
H1:negligible_exp = Some n
Hn:(n <= fexp n)%Z
K:(round beta fexp Zfloor x < - bpow (fexp n))%R

(- bpow (fexp n) < round beta fexp Zceil x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
Zx:round beta fexp Zceil x = 0%R
n:Z
H1:negligible_exp = Some n
Hn:(n <= fexp n)%Z
K:(round beta fexp Zfloor x < - bpow (fexp n))%R
(round beta fexp Zceil x <= - bpow (fexp n))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
Zx:round beta fexp Zceil x <> 0%R
(pred (round beta fexp Zceil x) <= round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
Zx:round beta fexp Zceil x = 0%R
n:Z
H1:negligible_exp = Some n
Hn:(n <= fexp n)%Z
K:(round beta fexp Zfloor x < - bpow (fexp n))%R

(- bpow (fexp n) < - 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
Zx:round beta fexp Zceil x = 0%R
n:Z
H1:negligible_exp = Some n
Hn:(n <= fexp n)%Z
K:(round beta fexp Zfloor x < - bpow (fexp n))%R
(round beta fexp Zceil x <= - bpow (fexp n))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
Zx:round beta fexp Zceil x <> 0%R
(pred (round beta fexp Zceil x) <= round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
Zx:round beta fexp Zceil x = 0%R
n:Z
H1:negligible_exp = Some n
Hn:(n <= fexp n)%Z
K:(round beta fexp Zfloor x < - bpow (fexp n))%R

(round beta fexp Zceil x <= - bpow (fexp n))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
Zx:round beta fexp Zceil x <> 0%R
(pred (round beta fexp Zceil x) <= round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
Zx:round beta fexp Zceil x = 0%R
n:Z
H1:negligible_exp = Some n
Hn:(n <= fexp n)%Z
K:(round beta fexp Zfloor x < - bpow (fexp n))%R

F (- bpow (fexp n))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
Zx:round beta fexp Zceil x <> 0%R
(pred (round beta fexp Zceil x) <= round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
Zx:round beta fexp Zceil x = 0%R
n:Z
H1:negligible_exp = Some n
Hn:(n <= fexp n)%Z
K:(round beta fexp Zfloor x < - bpow (fexp n))%R

(fexp (fexp n + 1) <= fexp n)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
Zx:round beta fexp Zceil x <> 0%R
(pred (round beta fexp Zceil x) <= round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
Zx:round beta fexp Zceil x <> 0%R

(pred (round beta fexp Zceil x) <= round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
Zx:round beta fexp Zceil x <> 0%R

let u := round beta fexp Zceil x in (pred u < u)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
Zx:round beta fexp Zceil x <> 0%R
Hup:let u := round beta fexp Zceil x in (pred u < u)%R
(pred (round beta fexp Zceil x) <= round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
Zx:round beta fexp Zceil x <> 0%R
Hup:let u := round beta fexp Zceil x in (pred u < u)%R

(pred (round beta fexp Zceil x) <= round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
Zx:round beta fexp Zceil x <> 0%R
Hup:let u := round beta fexp Zceil x in (pred u < u)%R

F (pred (round beta fexp Zceil x))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
Zx:round beta fexp Zceil x <> 0%R
Hup:let u := round beta fexp Zceil x in (pred u < u)%R

F (round beta fexp Zceil x)
now apply round_UP_pt. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, (round beta fexp Zceil x <= succ (round beta fexp Zfloor x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, (round beta fexp Zceil x <= succ (round beta fexp Zfloor x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R

(round beta fexp Zceil x <= succ (round beta fexp Zfloor x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R

(round beta fexp Zceil (- - x) <= succ (round beta fexp Zfloor (- - x)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R

(- round beta fexp Zfloor (- x) <= - pred (round beta fexp Zceil (- x)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R

(pred (round beta fexp Zceil (- x)) <= round beta fexp Zfloor (- x))%R
apply pred_UP_le_DN. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, ~ F x -> pred (round beta fexp Zceil x) = round beta fexp Zfloor x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, ~ F x -> pred (round beta fexp Zceil x) = round beta fexp Zfloor x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x

pred (round beta fexp Zceil x) = round beta fexp Zfloor x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x

(pred (round beta fexp Zceil x) <= round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
(round beta fexp Zfloor x <= pred (round beta fexp Zceil x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x

(round beta fexp Zfloor x <= pred (round beta fexp Zceil x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x

(round beta fexp Zfloor x < round beta fexp Zceil x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x
HE:(round beta fexp Zfloor x < x < round beta fexp Zceil x)%R

(round beta fexp Zfloor x < round beta fexp Zceil x)%R
now apply Rlt_trans with (1 := proj1 HE) (2 := proj2 HE). Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, ~ F x -> succ (round beta fexp Zfloor x) = round beta fexp Zceil x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, ~ F x -> succ (round beta fexp Zfloor x) = round beta fexp Zceil x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x

succ (round beta fexp Zfloor x) = round beta fexp Zceil x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x

succ (pred (round beta fexp Zceil x)) = round beta fexp Zceil x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:~ F x

F (round beta fexp Zceil x)
apply generic_format_round... Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x d : R, F d -> (d <= x < succ d)%R -> round beta fexp Zfloor x = d
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x d : R, F d -> (d <= x < succ d)%R -> round beta fexp Zfloor x = d
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, d:R
Fd:F d
Hxd1:(d <= x)%R
Hxd2:(x < succ d)%R

round beta fexp Zfloor x = d
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, d:R
Fd:F d
Hxd1:(d <= x)%R
Hxd2:(x < succ d)%R
T1:F (round beta fexp Zfloor x)
T2:(round beta fexp Zfloor x <= x)%R
T3:forall g : R, F g -> (g <= x)%R -> (g <= round beta fexp Zfloor x)%R

round beta fexp Zfloor x = d
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, d:R
Fd:F d
Hxd1:(d <= x)%R
Hxd2:(x < succ d)%R
T1:F (round beta fexp Zfloor x)
T2:(round beta fexp Zfloor x <= x)%R
T3:forall g : R, F g -> (g <= x)%R -> (g <= round beta fexp Zfloor x)%R

(d <= round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, d:R
Fd:F d
Hxd1:(d <= x)%R
Hxd2:(x < succ d)%R
T1:F (round beta fexp Zfloor x)
T2:(round beta fexp Zfloor x <= x)%R
T3:forall g : R, F g -> (g <= x)%R -> (g <= round beta fexp Zfloor x)%R
(round beta fexp Zfloor x <= d)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, d:R
Fd:F d
Hxd1:(d <= x)%R
Hxd2:(x < succ d)%R
T1:F (round beta fexp Zfloor x)
T2:(round beta fexp Zfloor x <= x)%R
T3:forall g : R, F g -> (g <= x)%R -> (g <= round beta fexp Zfloor x)%R

(round beta fexp Zfloor x <= d)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, d:R
Fd:F d
Hxd1:(d <= x)%R
Hxd2:(x < succ d)%R
T1:F (round beta fexp Zfloor x)
T2:(round beta fexp Zfloor x <= x)%R
T3:forall g : R, F g -> (g <= x)%R -> (g <= round beta fexp Zfloor x)%R
Fx:F x

(round beta fexp Zfloor x <= d)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, d:R
Fd:F d
Hxd1:(d <= x)%R
Hxd2:(x < succ d)%R
T1:F (round beta fexp Zfloor x)
T2:(round beta fexp Zfloor x <= x)%R
T3:forall g : R, F g -> (g <= x)%R -> (g <= round beta fexp Zfloor x)%R
NFx:~ F x
(round beta fexp Zfloor x <= d)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, d:R
Fd:F d
Hxd1:(d <= x)%R
Hxd2:(x < succ d)%R
T1:F (round beta fexp Zfloor x)
T2:(round beta fexp Zfloor x <= x)%R
T3:forall g : R, F g -> (g <= x)%R -> (g <= round beta fexp Zfloor x)%R
Fx:F x

(x <= d)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, d:R
Fd:F d
Hxd1:(d <= x)%R
Hxd2:(x < succ d)%R
T1:F (round beta fexp Zfloor x)
T2:(round beta fexp Zfloor x <= x)%R
T3:forall g : R, F g -> (g <= x)%R -> (g <= round beta fexp Zfloor x)%R
NFx:~ F x
(round beta fexp Zfloor x <= d)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, d:R
Fd:F d
Hxd1:(d <= x)%R
Hxd2:(x < succ d)%R
T1:F (round beta fexp Zfloor x)
T2:(round beta fexp Zfloor x <= x)%R
T3:forall g : R, F g -> (g <= x)%R -> (g <= round beta fexp Zfloor x)%R
Fx:F x

(succ x <= succ d)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, d:R
Fd:F d
Hxd1:(d <= x)%R
Hxd2:(x < succ d)%R
T1:F (round beta fexp Zfloor x)
T2:(round beta fexp Zfloor x <= x)%R
T3:forall g : R, F g -> (g <= x)%R -> (g <= round beta fexp Zfloor x)%R
NFx:~ F x
(round beta fexp Zfloor x <= d)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, d:R
Fd:F d
Hxd1:(d <= x)%R
Hxd2:(x < succ d)%R
T1:F (round beta fexp Zfloor x)
T2:(round beta fexp Zfloor x <= x)%R
T3:forall g : R, F g -> (g <= x)%R -> (g <= round beta fexp Zfloor x)%R
Fx:F x

F (succ d)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, d:R
Fd:F d
Hxd1:(d <= x)%R
Hxd2:(x < succ d)%R
T1:F (round beta fexp Zfloor x)
T2:(round beta fexp Zfloor x <= x)%R
T3:forall g : R, F g -> (g <= x)%R -> (g <= round beta fexp Zfloor x)%R
NFx:~ F x
(round beta fexp Zfloor x <= d)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, d:R
Fd:F d
Hxd1:(d <= x)%R
Hxd2:(x < succ d)%R
T1:F (round beta fexp Zfloor x)
T2:(round beta fexp Zfloor x <= x)%R
T3:forall g : R, F g -> (g <= x)%R -> (g <= round beta fexp Zfloor x)%R
NFx:~ F x

(round beta fexp Zfloor x <= d)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, d:R
Fd:F d
Hxd1:(d <= x)%R
Hxd2:(x < succ d)%R
T1:F (round beta fexp Zfloor x)
T2:(round beta fexp Zfloor x <= x)%R
T3:forall g : R, F g -> (g <= x)%R -> (g <= round beta fexp Zfloor x)%R
NFx:~ F x

(succ (round beta fexp Zfloor x) <= succ d)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, d:R
Fd:F d
Hxd1:(d <= x)%R
Hxd2:(x < succ d)%R
T1:F (round beta fexp Zfloor x)
T2:(round beta fexp Zfloor x <= x)%R
T3:forall g : R, F g -> (g <= x)%R -> (g <= round beta fexp Zfloor x)%R
NFx:~ F x

(round beta fexp Zceil x <= succ d)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, d:R
Fd:F d
Hxd1:(d <= x)%R
Hxd2:(x < succ d)%R
T1:F (round beta fexp Zfloor x)
T2:(round beta fexp Zfloor x <= x)%R
T3:forall g : R, F g -> (g <= x)%R -> (g <= round beta fexp Zfloor x)%R
NFx:~ F x

F (succ d)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, d:R
Fd:F d
Hxd1:(d <= x)%R
Hxd2:(x < succ d)%R
T1:F (round beta fexp Zfloor x)
T2:(round beta fexp Zfloor x <= x)%R
T3:forall g : R, F g -> (g <= x)%R -> (g <= round beta fexp Zfloor x)%R
NFx:~ F x
(x <= succ d)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, d:R
Fd:F d
Hxd1:(d <= x)%R
Hxd2:(x < succ d)%R
T1:F (round beta fexp Zfloor x)
T2:(round beta fexp Zfloor x <= x)%R
T3:forall g : R, F g -> (g <= x)%R -> (g <= round beta fexp Zfloor x)%R
NFx:~ F x

(x <= succ d)%R
now left. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x u : R, F u -> (pred u < x <= u)%R -> round beta fexp Zceil x = u
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x u : R, F u -> (pred u < x <= u)%R -> round beta fexp Zceil x = u
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, u:R
Fu:F u
Hux:(pred u < x <= u)%R

round beta fexp Zceil x = u
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, u:R
Fu:F u
Hux:(pred u < x <= u)%R

(- - round beta fexp Zceil x)%R = u
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, u:R
Fu:F u
Hux:(pred u < x <= u)%R

(- round beta fexp Zfloor (- x))%R = u
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, u:R
Fu:F u
Hux:(pred u < x <= u)%R

(- round beta fexp Zfloor (- x))%R = (- - u)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, u:R
Fu:F u
Hux:(pred u < x <= u)%R

round beta fexp Zfloor (- x) = (- u)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, u:R
Fu:F u
Hux:(pred u < x <= u)%R

F (- u)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, u:R
Fu:F u
Hux:(pred u < x <= u)%R
(- u <= - x < succ (- u))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, u:R
Fu:F u
Hux:(pred u < x <= u)%R

(- u <= - x < succ (- u))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, u:R
Fu:F u
Hux:(pred u < x <= u)%R

(- x < succ (- u))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x, u:R
Fu:F u
Hux:(pred u < x <= u)%R

(- x < - pred u)%R
now apply Ropp_lt_contravar. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

Exp_not_FTZ fexp -> ulp (ulp 0) = ulp 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

Exp_not_FTZ fexp -> ulp (ulp 0) = ulp 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:Exp_not_FTZ fexp

negligible_exp = None /\ (forall n : Z, (fexp n < n)%Z) -> ulp (ulp 0) = ulp 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:Exp_not_FTZ fexp
(exists n : Z, negligible_exp = Some n /\ (n <= fexp n)%Z) -> ulp (ulp 0) = ulp 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:Exp_not_FTZ fexp
K1:negligible_exp = None
K2:forall n : Z, (fexp n < n)%Z

ulp (ulp 0) = ulp 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:Exp_not_FTZ fexp
(exists n : Z, negligible_exp = Some n /\ (n <= fexp n)%Z) -> ulp (ulp 0) = ulp 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:Exp_not_FTZ fexp
K1:negligible_exp = None
K2:forall n : Z, (fexp n < n)%Z

0%R = ulp 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:Exp_not_FTZ fexp
(exists n : Z, negligible_exp = Some n /\ (n <= fexp n)%Z) -> ulp (ulp 0) = ulp 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:Exp_not_FTZ fexp
K1:negligible_exp = None
K2:forall n : Z, (fexp n < n)%Z

match negligible_exp with | Some n => bpow (fexp n) | None => 0%R end = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:Exp_not_FTZ fexp
(exists n : Z, negligible_exp = Some n /\ (n <= fexp n)%Z) -> ulp (ulp 0) = ulp 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:Exp_not_FTZ fexp

(exists n : Z, negligible_exp = Some n /\ (n <= fexp n)%Z) -> ulp (ulp 0) = ulp 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:Exp_not_FTZ fexp
n:Z
Hn1:negligible_exp = Some n
Hn2:(n <= fexp n)%Z

ulp (ulp 0) = ulp 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:Exp_not_FTZ fexp
n:Z
Hn1:negligible_exp = Some n
Hn2:(n <= fexp n)%Z

(ulp (ulp 0) <= ulp 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:Exp_not_FTZ fexp
n:Z
Hn1:negligible_exp = Some n
Hn2:(n <= fexp n)%Z
(ulp 0 <= ulp (ulp 0))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:Exp_not_FTZ fexp
n:Z
Hn1:negligible_exp = Some n
Hn2:(n <= fexp n)%Z

(ulp (bpow (fexp n)) <= bpow (fexp n))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:Exp_not_FTZ fexp
n:Z
Hn1:negligible_exp = Some n
Hn2:(n <= fexp n)%Z
bpow (fexp n) = ulp 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:Exp_not_FTZ fexp
n:Z
Hn1:negligible_exp = Some n
Hn2:(n <= fexp n)%Z
(ulp 0 <= ulp (ulp 0))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:Exp_not_FTZ fexp
n:Z
Hn1:negligible_exp = Some n
Hn2:(n <= fexp n)%Z

(bpow (fexp (fexp n + 1)) <= bpow (fexp n))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:Exp_not_FTZ fexp
n:Z
Hn1:negligible_exp = Some n
Hn2:(n <= fexp n)%Z
bpow (fexp n) = ulp 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:Exp_not_FTZ fexp
n:Z
Hn1:negligible_exp = Some n
Hn2:(n <= fexp n)%Z
(ulp 0 <= ulp (ulp 0))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:Exp_not_FTZ fexp
n:Z
Hn1:negligible_exp = Some n
Hn2:(n <= fexp n)%Z

(fexp (fexp n + 1) <= fexp n)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:Exp_not_FTZ fexp
n:Z
Hn1:negligible_exp = Some n
Hn2:(n <= fexp n)%Z
bpow (fexp n) = ulp 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:Exp_not_FTZ fexp
n:Z
Hn1:negligible_exp = Some n
Hn2:(n <= fexp n)%Z
(ulp 0 <= ulp (ulp 0))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:Exp_not_FTZ fexp
n:Z
Hn1:negligible_exp = Some n
Hn2:(n <= fexp n)%Z

bpow (fexp n) = ulp 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:Exp_not_FTZ fexp
n:Z
Hn1:negligible_exp = Some n
Hn2:(n <= fexp n)%Z
(ulp 0 <= ulp (ulp 0))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:Exp_not_FTZ fexp
n:Z
Hn1:negligible_exp = Some n
Hn2:(n <= fexp n)%Z

bpow (fexp n) = match negligible_exp with | Some n0 => bpow (fexp n0) | None => 0%R end
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:Exp_not_FTZ fexp
n:Z
Hn1:negligible_exp = Some n
Hn2:(n <= fexp n)%Z
(ulp 0 <= ulp (ulp 0))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
H:Exp_not_FTZ fexp
n:Z
Hn1:negligible_exp = Some n
Hn2:(n <= fexp n)%Z

(ulp 0 <= ulp (ulp 0))%R
now apply ulp_ge_ulp_0. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, F x -> (0 < x)%R -> ulp (succ x) = ulp x \/ succ x = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, F x -> (0 < x)%R -> ulp (succ x) = ulp x \/ succ x = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R

ulp (succ x) = ulp x \/ succ x = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R
Hx':(0 <= x)%R

ulp (succ x) = ulp x \/ succ x = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R
Hx':(0 <= x)%R

ulp (x + ulp x) = ulp x \/ (x + ulp x)%R = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R
Hx':(0 <= x)%R
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R

ulp (x + ulp x) = ulp x \/ (x + ulp x)%R = bpow e
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R
Hx':(0 <= x)%R
e:Z
He:x <> 0%R -> (bpow (e - 1) <= x < bpow e)%R

ulp (x + ulp x) = ulp x \/ (x + ulp x)%R = bpow e
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R
Hx':(0 <= x)%R
e:Z
He:(bpow (e - 1) <= x < bpow e)%R

ulp (x + ulp x) = ulp x \/ (x + ulp x)%R = bpow e
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R
Hx':(0 <= x)%R
e:Z
He:(bpow (e - 1) <= x < bpow e)%R

(x + ulp x <= bpow e)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R
Hx':(0 <= x)%R
e:Z
He:(bpow (e - 1) <= x < bpow e)%R
H:(x + ulp x <= bpow e)%R
ulp (x + ulp x) = ulp x \/ (x + ulp x)%R = bpow e
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R
Hx':(0 <= x)%R
e:Z
He:(bpow (e - 1) <= x < bpow e)%R

(x < bpow e)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R
Hx':(0 <= x)%R
e:Z
He:(bpow (e - 1) <= x < bpow e)%R
H:(x + ulp x <= bpow e)%R
ulp (x + ulp x) = ulp x \/ (x + ulp x)%R = bpow e
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R
Hx':(0 <= x)%R
e:Z
He:(bpow (e - 1) <= x < bpow e)%R
H:(x + ulp x <= bpow e)%R

ulp (x + ulp x) = ulp x \/ (x + ulp x)%R = bpow e
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R
Hx':(0 <= x)%R
e:Z
He:(bpow (e - 1) <= x < bpow e)%R
H:(x + ulp x < bpow e)%R

ulp (x + ulp x) = ulp x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R
Hx':(0 <= x)%R
e:Z
He:(bpow (e - 1) <= x < bpow e)%R
H:(x + ulp x < bpow e)%R

bpow (cexp beta fexp (x + ulp x)) = ulp x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R
Hx':(0 <= x)%R
e:Z
He:(bpow (e - 1) <= x < bpow e)%R
H:(x + ulp x < bpow e)%R
(x + ulp x)%R <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R
Hx':(0 <= x)%R
e:Z
He:(bpow (e - 1) <= x < bpow e)%R
H:(x + ulp x < bpow e)%R

bpow (cexp beta fexp (x + ulp x)) = ulp x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R
Hx':(0 <= x)%R
e:Z
He:(bpow (e - 1) <= x < bpow e)%R
H:(x + ulp x < bpow e)%R
(x <= x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R
Hx':(0 <= x)%R
e:Z
He:(bpow (e - 1) <= x < bpow e)%R
H:(x + ulp x < bpow e)%R

bpow (cexp beta fexp (x + ulp x)) = ulp x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R
Hx':(0 <= x)%R
e:Z
He:(bpow (e - 1) <= x < bpow e)%R
H:(x + ulp x < bpow e)%R
(0 <= ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R
Hx':(0 <= x)%R
e:Z
He:(bpow (e - 1) <= x < bpow e)%R
H:(x + ulp x < bpow e)%R

bpow (cexp beta fexp (x + ulp x)) = ulp x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R
Hx':(0 <= x)%R
e:Z
He:(bpow (e - 1) <= x < bpow e)%R
H:(x + ulp x < bpow e)%R

bpow (cexp beta fexp (x + ulp x)) = bpow (cexp beta fexp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R
Hx':(0 <= x)%R
e:Z
He:(bpow (e - 1) <= x < bpow e)%R
H:(x + ulp x < bpow e)%R
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R
Hx':(0 <= x)%R
e:Z
He:(bpow (e - 1) <= x < bpow e)%R
H:(x + ulp x < bpow e)%R

bpow (cexp beta fexp (x + ulp x)) = bpow (cexp beta fexp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R
Hx':(0 <= x)%R
e:Z
He:(bpow (e - 1) <= x < bpow e)%R
H:(x + ulp x < bpow e)%R

mag beta (x + ulp x) = mag beta x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R
Hx':(0 <= x)%R
e:Z
He:(bpow (e - 1) <= x < bpow e)%R
H:(x + ulp x < bpow e)%R

mag beta (x + ulp x) = e
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R
Hx':(0 <= x)%R
e:Z
He:(bpow (e - 1) <= x < bpow e)%R
H:(x + ulp x < bpow e)%R
e = mag beta x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R
Hx':(0 <= x)%R
e:Z
He:(bpow (e - 1) <= x < bpow e)%R
H:(x + ulp x < bpow e)%R

(bpow (e - 1) <= x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R
Hx':(0 <= x)%R
e:Z
He:(bpow (e - 1) <= x < bpow e)%R
H:(x + ulp x < bpow e)%R
e = mag beta x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R
Hx':(0 <= x)%R
e:Z
He:(bpow (e - 1) <= x < bpow e)%R
H:(x + ulp x < bpow e)%R

(x <= x + ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R
Hx':(0 <= x)%R
e:Z
He:(bpow (e - 1) <= x < bpow e)%R
H:(x + ulp x < bpow e)%R
e = mag beta x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R
Hx':(0 <= x)%R
e:Z
He:(bpow (e - 1) <= x < bpow e)%R
H:(x + ulp x < bpow e)%R

(0 <= ulp x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R
Hx':(0 <= x)%R
e:Z
He:(bpow (e - 1) <= x < bpow e)%R
H:(x + ulp x < bpow e)%R
e = mag beta x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < x)%R
Hx':(0 <= x)%R
e:Z
He:(bpow (e - 1) <= x < bpow e)%R
H:(x + ulp x < bpow e)%R

e = mag beta x
now apply sym_eq, mag_unique_pos. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, F x -> (0 < pred x)%R -> ulp (pred x) = ulp x \/ x = bpow (mag beta x - 1)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall x : R, F x -> (0 < pred x)%R -> ulp (pred x) = ulp x \/ x = bpow (mag beta x - 1)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < pred x)%R

ulp (pred x) = ulp x \/ x = bpow (mag beta x - 1)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < pred x)%R

(0 < x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < pred x)%R
Hx':(0 < x)%R
ulp (pred x) = ulp x \/ x = bpow (mag beta x - 1)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < pred x)%R

(pred x <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < pred x)%R
Hx':(0 < x)%R
ulp (pred x) = ulp x \/ x = bpow (mag beta x - 1)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < pred x)%R
Hx':(0 < x)%R

ulp (pred x) = ulp x \/ x = bpow (mag beta x - 1)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < pred x)%R
Hx':(0 < x)%R

x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < pred x)%R
Hx':(0 < x)%R
Zx:x <> 0%R
ulp (pred x) = ulp x \/ x = bpow (mag beta x - 1)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < pred x)%R
Hx':(0 < x)%R
Zx:x <> 0%R

ulp (pred x) = ulp x \/ x = bpow (mag beta x - 1)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < pred x)%R
Hx':(0 < x)%R
Zx:x <> 0%R

ulp (pred x) = bpow (cexp beta fexp x) \/ x = bpow (mag beta x - 1)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < pred x)%R
Hx':(0 < x)%R
Zx:x <> 0%R

ulp (pred x) = bpow (fexp (mag beta x)) \/ x = bpow (mag beta x - 1)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < pred x)%R
Hx':(0 < x)%R
Zx:x <> 0%R
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R

ulp (pred x) = bpow (fexp (Build_mag_prop beta x e He)) \/ x = bpow (Build_mag_prop beta x e He - 1)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < pred x)%R
Hx':(0 < x)%R
Zx:x <> 0%R
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R

ulp (pred x) = bpow (fexp e) \/ x = bpow (e - 1)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < pred x)%R
Hx':(0 < x)%R
Zx:x <> 0%R
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R

(bpow (e - 1) <= x < bpow e)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < pred x)%R
Hx':(0 < x)%R
Zx:x <> 0%R
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R
H:(bpow (e - 1) <= x < bpow e)%R
ulp (pred x) = bpow (fexp e) \/ x = bpow (e - 1)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < pred x)%R
Hx':(0 < x)%R
Zx:x <> 0%R
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R

(bpow (e - 1) <= Rabs x < bpow e)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < pred x)%R
Hx':(0 < x)%R
Zx:x <> 0%R
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R
H:(bpow (e - 1) <= x < bpow e)%R
ulp (pred x) = bpow (fexp e) \/ x = bpow (e - 1)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < pred x)%R
Hx':(0 < x)%R
Zx:x <> 0%R
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R
H:(bpow (e - 1) <= x < bpow e)%R

ulp (pred x) = bpow (fexp e) \/ x = bpow (e - 1)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < pred x)%R
Hx':(0 < x)%R
Zx:x <> 0%R
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R
H:(bpow (e - 1) <= x < bpow e)%R
H1:(bpow (e - 1) < x)%R

ulp (pred x) = bpow (fexp e) \/ x = bpow (e - 1)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < pred x)%R
Hx':(0 < x)%R
Zx:x <> 0%R
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R
H:(bpow (e - 1) <= x < bpow e)%R
H1:bpow (e - 1) = x
ulp (pred x) = bpow (fexp e) \/ x = bpow (e - 1)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < pred x)%R
Hx':(0 < x)%R
Zx:x <> 0%R
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R
H:(bpow (e - 1) <= x < bpow e)%R
H1:(bpow (e - 1) < x)%R

ulp (pred x) = bpow (fexp e) \/ x = bpow (e - 1)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < pred x)%R
Hx':(0 < x)%R
Zx:x <> 0%R
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R
H:(bpow (e - 1) <= x < bpow e)%R
H1:(bpow (e - 1) < x)%R

ulp (pred x) = bpow (fexp e)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < pred x)%R
Hx':(0 < x)%R
Zx:x <> 0%R
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R
H:(bpow (e - 1) <= x < bpow e)%R
H1:(bpow (e - 1) <= pred x)%R

ulp (pred x) = bpow (fexp e)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < pred x)%R
Hx':(0 < x)%R
Zx:x <> 0%R
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R
H:(bpow (e - 1) <= x < bpow e)%R
H1:(bpow (e - 1) < x)%R
F (bpow (e - 1))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < pred x)%R
Hx':(0 < x)%R
Zx:x <> 0%R
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R
H:(bpow (e - 1) <= x < bpow e)%R
H1:(bpow (e - 1) <= pred x)%R

bpow (cexp beta fexp (pred x)) = bpow (fexp e)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < pred x)%R
Hx':(0 < x)%R
Zx:x <> 0%R
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R
H:(bpow (e - 1) <= x < bpow e)%R
H1:(bpow (e - 1) < x)%R
F (bpow (e - 1))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < pred x)%R
Hx':(0 < x)%R
Zx:x <> 0%R
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R
H:(bpow (e - 1) <= x < bpow e)%R
H1:(bpow (e - 1) <= pred x)%R

mag beta (pred x) = e
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < pred x)%R
Hx':(0 < x)%R
Zx:x <> 0%R
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R
H:(bpow (e - 1) <= x < bpow e)%R
H1:(bpow (e - 1) < x)%R
F (bpow (e - 1))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < pred x)%R
Hx':(0 < x)%R
Zx:x <> 0%R
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R
H:(bpow (e - 1) <= x < bpow e)%R
H1:(bpow (e - 1) <= pred x)%R

(bpow (e - 1) <= pred x < bpow e)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < pred x)%R
Hx':(0 < x)%R
Zx:x <> 0%R
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R
H:(bpow (e - 1) <= x < bpow e)%R
H1:(bpow (e - 1) < x)%R
F (bpow (e - 1))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < pred x)%R
Hx':(0 < x)%R
Zx:x <> 0%R
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R
H:(bpow (e - 1) <= x < bpow e)%R
H1:(bpow (e - 1) <= pred x)%R

(pred x < bpow e)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < pred x)%R
Hx':(0 < x)%R
Zx:x <> 0%R
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R
H:(bpow (e - 1) <= x < bpow e)%R
H1:(bpow (e - 1) < x)%R
F (bpow (e - 1))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < pred x)%R
Hx':(0 < x)%R
Zx:x <> 0%R
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R
H:(bpow (e - 1) <= x < bpow e)%R
H1:(bpow (e - 1) <= pred x)%R

(pred x <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < pred x)%R
Hx':(0 < x)%R
Zx:x <> 0%R
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R
H:(bpow (e - 1) <= x < bpow e)%R
H1:(bpow (e - 1) < x)%R
F (bpow (e - 1))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < pred x)%R
Hx':(0 < x)%R
Zx:x <> 0%R
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R
H:(bpow (e - 1) <= x < bpow e)%R
H1:(bpow (e - 1) < x)%R

F (bpow (e - 1))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < pred x)%R
Hx':(0 < x)%R
Zx:x <> 0%R
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R
H:(bpow (e - 1) <= x < bpow e)%R
H1:(bpow (e - 1) < x)%R

(fexp (e - 1 + 1) <= e - 1)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < pred x)%R
Hx':(0 < x)%R
Zx:x <> 0%R
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R
H:(bpow (e - 1) <= x < bpow e)%R
H1:(bpow (e - 1) < x)%R

(fexp (e - 1 + 1) < e)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < pred x)%R
Hx':(0 < x)%R
Zx:x <> 0%R
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R
H:(bpow (e - 1) <= x < bpow e)%R
H1:(bpow (e - 1) < x)%R

(fexp e < e)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
x:R
Fx:F x
Hx:(0 < pred x)%R
Hx':(0 < x)%R
Zx:x <> 0%R
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R
H:(bpow (e - 1) <= x < bpow e)%R
H1:(bpow (e - 1) < x)%R

(fexp (mag beta x) < mag beta x)%Z
now apply mag_generic_gt. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

Exp_not_FTZ fexp -> forall rnd : R -> Z, Valid_rnd rnd -> forall x : R, (0 < x)%R -> ulp (round beta fexp rnd x) = ulp x \/ round beta fexp rnd x = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

Exp_not_FTZ fexp -> forall rnd : R -> Z, Valid_rnd rnd -> forall x : R, (0 < x)%R -> ulp (round beta fexp rnd x) = ulp x \/ round beta fexp rnd x = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R

ulp (round beta fexp rnd x) = ulp x \/ round beta fexp rnd x = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:F x

ulp (round beta fexp rnd x) = ulp x \/ round beta fexp rnd x = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
ulp (round beta fexp rnd x) = ulp x \/ round beta fexp rnd x = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x

ulp (round beta fexp rnd x) = ulp x \/ round beta fexp rnd x = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zfloor x

ulp (round beta fexp Zfloor x) = ulp x \/ round beta fexp Zfloor x = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
ulp (round beta fexp Zceil x) = ulp x \/ round beta fexp Zceil x = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zfloor x

ulp (round beta fexp Zfloor x) = ulp x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
ulp (round beta fexp Zceil x) = ulp x \/ round beta fexp Zceil x = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x

ulp (round beta fexp Zceil x) = ulp x \/ round beta fexp Zceil x = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x

(0 <= round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:(0 <= round beta fexp Zfloor x)%R
ulp (round beta fexp Zceil x) = ulp x \/ round beta fexp Zceil x = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x

F 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:(0 <= round beta fexp Zfloor x)%R
ulp (round beta fexp Zceil x) = ulp x \/ round beta fexp Zceil x = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x

(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:(0 <= round beta fexp Zfloor x)%R
ulp (round beta fexp Zceil x) = ulp x \/ round beta fexp Zceil x = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:(0 <= round beta fexp Zfloor x)%R

ulp (round beta fexp Zceil x) = ulp x \/ round beta fexp Zceil x = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:(0 < round beta fexp Zfloor x)%R

ulp (round beta fexp Zceil x) = ulp x \/ round beta fexp Zceil x = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:0%R = round beta fexp Zfloor x
ulp (round beta fexp Zceil x) = ulp x \/ round beta fexp Zceil x = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:(0 < round beta fexp Zfloor x)%R

ulp (succ (round beta fexp Zfloor x)) = ulp x \/ succ (round beta fexp Zfloor x) = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:0%R = round beta fexp Zfloor x
ulp (round beta fexp Zceil x) = ulp x \/ round beta fexp Zceil x = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:(0 < round beta fexp Zfloor x)%R

F (round beta fexp Zfloor x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:(0 < round beta fexp Zfloor x)%R
(0 < round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:(0 < round beta fexp Zfloor x)%R
Y:ulp (succ (round beta fexp Zfloor x)) = ulp (round beta fexp Zfloor x)
ulp (succ (round beta fexp Zfloor x)) = ulp x \/ succ (round beta fexp Zfloor x) = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:(0 < round beta fexp Zfloor x)%R
Y:succ (round beta fexp Zfloor x) = bpow (mag beta (round beta fexp Zfloor x))
ulp (succ (round beta fexp Zfloor x)) = ulp x \/ succ (round beta fexp Zfloor x) = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:0%R = round beta fexp Zfloor x
ulp (round beta fexp Zceil x) = ulp x \/ round beta fexp Zceil x = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:(0 < round beta fexp Zfloor x)%R

(0 < round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:(0 < round beta fexp Zfloor x)%R
Y:ulp (succ (round beta fexp Zfloor x)) = ulp (round beta fexp Zfloor x)
ulp (succ (round beta fexp Zfloor x)) = ulp x \/ succ (round beta fexp Zfloor x) = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:(0 < round beta fexp Zfloor x)%R
Y:succ (round beta fexp Zfloor x) = bpow (mag beta (round beta fexp Zfloor x))
ulp (succ (round beta fexp Zfloor x)) = ulp x \/ succ (round beta fexp Zfloor x) = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:0%R = round beta fexp Zfloor x
ulp (round beta fexp Zceil x) = ulp x \/ round beta fexp Zceil x = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:(0 < round beta fexp Zfloor x)%R
Y:ulp (succ (round beta fexp Zfloor x)) = ulp (round beta fexp Zfloor x)

ulp (succ (round beta fexp Zfloor x)) = ulp x \/ succ (round beta fexp Zfloor x) = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:(0 < round beta fexp Zfloor x)%R
Y:succ (round beta fexp Zfloor x) = bpow (mag beta (round beta fexp Zfloor x))
ulp (succ (round beta fexp Zfloor x)) = ulp x \/ succ (round beta fexp Zfloor x) = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:0%R = round beta fexp Zfloor x
ulp (round beta fexp Zceil x) = ulp x \/ round beta fexp Zceil x = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:(0 < round beta fexp Zfloor x)%R
Y:ulp (succ (round beta fexp Zfloor x)) = ulp (round beta fexp Zfloor x)

(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:(0 < round beta fexp Zfloor x)%R
Y:succ (round beta fexp Zfloor x) = bpow (mag beta (round beta fexp Zfloor x))
ulp (succ (round beta fexp Zfloor x)) = ulp x \/ succ (round beta fexp Zfloor x) = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:0%R = round beta fexp Zfloor x
ulp (round beta fexp Zceil x) = ulp x \/ round beta fexp Zceil x = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:(0 < round beta fexp Zfloor x)%R
Y:succ (round beta fexp Zfloor x) = bpow (mag beta (round beta fexp Zfloor x))

ulp (succ (round beta fexp Zfloor x)) = ulp x \/ succ (round beta fexp Zfloor x) = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:0%R = round beta fexp Zfloor x
ulp (round beta fexp Zceil x) = ulp x \/ round beta fexp Zceil x = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:(0 < round beta fexp Zfloor x)%R
Y:succ (round beta fexp Zfloor x) = bpow (mag beta (round beta fexp Zfloor x))

bpow (mag beta (round beta fexp Zfloor x)) = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:0%R = round beta fexp Zfloor x
ulp (round beta fexp Zceil x) = ulp x \/ round beta fexp Zceil x = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:0%R = round beta fexp Zfloor x

ulp (round beta fexp Zceil x) = ulp x \/ round beta fexp Zceil x = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:0%R = round beta fexp Zfloor x

ulp (succ (round beta fexp Zfloor x)) = ulp x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:0%R = round beta fexp Zfloor x

ulp (ulp 0) = ulp x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:0%R = round beta fexp Zfloor x

ulp 0 = ulp x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:0%R = round beta fexp Zfloor x

negligible_exp = None /\ (forall n : Z, (fexp n < n)%Z) -> ulp 0 = ulp x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:0%R = round beta fexp Zfloor x
(exists n : Z, negligible_exp = Some n /\ (n <= fexp n)%Z) -> ulp 0 = ulp x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:0%R = round beta fexp Zfloor x
K1:negligible_exp = None
K2:forall n : Z, (fexp n < n)%Z

ulp 0 = ulp x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:0%R = round beta fexp Zfloor x
(exists n : Z, negligible_exp = Some n /\ (n <= fexp n)%Z) -> ulp 0 = ulp x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:0%R = round beta fexp Zfloor x
K1:negligible_exp = None
K2:forall n : Z, (fexp n < n)%Z

x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:0%R = round beta fexp Zfloor x
K1:negligible_exp = None
K2:forall n : Z, (fexp n < n)%Z
x = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:0%R = round beta fexp Zfloor x
(exists n : Z, negligible_exp = Some n /\ (n <= fexp n)%Z) -> ulp 0 = ulp x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:0%R = round beta fexp Zfloor x
K1:negligible_exp = None
K2:forall n : Z, (fexp n < n)%Z

x = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:0%R = round beta fexp Zfloor x
(exists n : Z, negligible_exp = Some n /\ (n <= fexp n)%Z) -> ulp 0 = ulp x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:0%R = round beta fexp Zfloor x

(exists n : Z, negligible_exp = Some n /\ (n <= fexp n)%Z) -> ulp 0 = ulp x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:0%R = round beta fexp Zfloor x
n:Z
Hn1:negligible_exp = Some n
Hn2:(n <= fexp n)%Z

ulp 0 = ulp x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:0%R = round beta fexp Zfloor x
n:Z
Hn1:negligible_exp = Some n
Hn2:(n <= fexp n)%Z

bpow (fexp n) = ulp x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:0%R = round beta fexp Zfloor x
n:Z
Hn1:negligible_exp = Some n
Hn2:(n <= fexp n)%Z
bpow (fexp n) = ulp 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:0%R = round beta fexp Zfloor x
n:Z
Hn1:negligible_exp = Some n
Hn2:(n <= fexp n)%Z

bpow (fexp n) = ulp x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:0%R = round beta fexp Zfloor x
n:Z
Hn1:negligible_exp = Some n
Hn2:(n <= fexp n)%Z
bpow (fexp n) = match negligible_exp with | Some n0 => bpow (fexp n0) | None => 0%R end
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:0%R = round beta fexp Zfloor x
n:Z
Hn1:negligible_exp = Some n
Hn2:(n <= fexp n)%Z

bpow (fexp n) = ulp x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:0%R = round beta fexp Zfloor x
n:Z
Hn1:negligible_exp = Some n
Hn2:(n <= fexp n)%Z

bpow (fexp n) = bpow (cexp beta fexp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:0%R = round beta fexp Zfloor x
n:Z
Hn1:negligible_exp = Some n
Hn2:(n <= fexp n)%Z
x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:0%R = round beta fexp Zfloor x
n:Z
Hn1:negligible_exp = Some n
Hn2:(n <= fexp n)%Z

bpow (fexp n) = bpow (cexp beta fexp x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:0%R = round beta fexp Zfloor x
n:Z
Hn1:negligible_exp = Some n
Hn2:(n <= fexp n)%Z

fexp n = fexp (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:0%R = round beta fexp Zfloor x
n:Z
Hn1:negligible_exp = Some n
Hn2:(n <= fexp n)%Z
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R

fexp n = fexp e
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:0%R = round beta fexp Zfloor x
n:Z
Hn1:negligible_exp = Some n
Hn2:(n <= fexp n)%Z
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R

(e <= fexp n)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:0%R = round beta fexp Zfloor x
n:Z
Hn1:negligible_exp = Some n
Hn2:(n <= fexp n)%Z
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R

(e <= fexp e)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:0%R = round beta fexp Zfloor x
n:Z
Hn1:negligible_exp = Some n
Hn2:(n <= fexp n)%Z
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R
H:(e <= fexp e)%Z
(e <= fexp n)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:0%R = round beta fexp Zfloor x
n:Z
Hn1:negligible_exp = Some n
Hn2:(n <= fexp n)%Z
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R

(bpow (e - 1) <= x < bpow e)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:0%R = round beta fexp Zfloor x
n:Z
Hn1:negligible_exp = Some n
Hn2:(n <= fexp n)%Z
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R
H:(e <= fexp e)%Z
(e <= fexp n)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:0%R = round beta fexp Zfloor x
n:Z
Hn1:negligible_exp = Some n
Hn2:(n <= fexp n)%Z
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R

(bpow (e - 1) <= Rabs x < bpow e)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:0%R = round beta fexp Zfloor x
n:Z
Hn1:negligible_exp = Some n
Hn2:(n <= fexp n)%Z
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:0%R = round beta fexp Zfloor x
n:Z
Hn1:negligible_exp = Some n
Hn2:(n <= fexp n)%Z
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R
H:(e <= fexp e)%Z
(e <= fexp n)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:0%R = round beta fexp Zfloor x
n:Z
Hn1:negligible_exp = Some n
Hn2:(n <= fexp n)%Z
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R

(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:0%R = round beta fexp Zfloor x
n:Z
Hn1:negligible_exp = Some n
Hn2:(n <= fexp n)%Z
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R
H:(e <= fexp e)%Z
(e <= fexp n)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:0%R = round beta fexp Zfloor x
n:Z
Hn1:negligible_exp = Some n
Hn2:(n <= fexp n)%Z
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R
H:(e <= fexp e)%Z

(e <= fexp n)%Z
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Fx:~ F x
Hr:round beta fexp rnd x = round beta fexp Zceil x
M:0%R = round beta fexp Zfloor x
n:Z
Hn1:negligible_exp = Some n
Hn2:(n <= fexp n)%Z
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R
H:(e <= fexp e)%Z

fexp e = fexp n
now apply fexp_negligible_exp_eq. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

Exp_not_FTZ fexp -> forall rnd : R -> Z, Valid_rnd rnd -> forall x : R, ulp (round beta fexp rnd x) = ulp x \/ Rabs (round beta fexp rnd x) = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

Exp_not_FTZ fexp -> forall rnd : R -> Z, Valid_rnd rnd -> forall x : R, ulp (round beta fexp rnd x) = ulp x \/ Rabs (round beta fexp rnd x) = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R

ulp (round beta fexp rnd x) = ulp x \/ Rabs (round beta fexp rnd x) = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:(x < 0)%R

ulp (round beta fexp rnd x) = ulp x \/ Rabs (round beta fexp rnd x) = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x = 0%R \/ (x > 0)%R
ulp (round beta fexp rnd x) = ulp x \/ Rabs (round beta fexp rnd x) = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:(x < 0)%R

(0 < - x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:(x < 0)%R
ulp (round beta fexp (Zrnd_opp rnd) (- x)) = ulp (- x) -> ulp (round beta fexp rnd x) = ulp x \/ Rabs (round beta fexp rnd x) = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:(x < 0)%R
round beta fexp (Zrnd_opp rnd) (- x) = bpow (mag beta (- x)) -> ulp (round beta fexp rnd x) = ulp x \/ Rabs (round beta fexp rnd x) = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x = 0%R \/ (x > 0)%R
ulp (round beta fexp rnd x) = ulp x \/ Rabs (round beta fexp rnd x) = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:(x < 0)%R

ulp (round beta fexp (Zrnd_opp rnd) (- x)) = ulp (- x) -> ulp (round beta fexp rnd x) = ulp x \/ Rabs (round beta fexp rnd x) = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:(x < 0)%R
round beta fexp (Zrnd_opp rnd) (- x) = bpow (mag beta (- x)) -> ulp (round beta fexp rnd x) = ulp x \/ Rabs (round beta fexp rnd x) = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x = 0%R \/ (x > 0)%R
ulp (round beta fexp rnd x) = ulp x \/ Rabs (round beta fexp rnd x) = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:(x < 0)%R

ulp (- round beta fexp (Zrnd_opp rnd) (- x)) = ulp x -> ulp (round beta fexp rnd x) = ulp x \/ Rabs (round beta fexp rnd x) = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:(x < 0)%R
round beta fexp (Zrnd_opp rnd) (- x) = bpow (mag beta (- x)) -> ulp (round beta fexp rnd x) = ulp x \/ Rabs (round beta fexp rnd x) = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x = 0%R \/ (x > 0)%R
ulp (round beta fexp rnd x) = ulp x \/ Rabs (round beta fexp rnd x) = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:(x < 0)%R

ulp (round beta fexp rnd x) = ulp x -> ulp (round beta fexp rnd x) = ulp x \/ Rabs (round beta fexp rnd x) = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:(x < 0)%R
round beta fexp (Zrnd_opp rnd) (- x) = bpow (mag beta (- x)) -> ulp (round beta fexp rnd x) = ulp x \/ Rabs (round beta fexp rnd x) = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x = 0%R \/ (x > 0)%R
ulp (round beta fexp rnd x) = ulp x \/ Rabs (round beta fexp rnd x) = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:(x < 0)%R

round beta fexp (Zrnd_opp rnd) (- x) = bpow (mag beta (- x)) -> ulp (round beta fexp rnd x) = ulp x \/ Rabs (round beta fexp rnd x) = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x = 0%R \/ (x > 0)%R
ulp (round beta fexp rnd x) = ulp x \/ Rabs (round beta fexp rnd x) = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:(x < 0)%R

round beta fexp (Zrnd_opp rnd) (- x) = bpow (mag beta x) -> ulp (round beta fexp rnd x) = ulp x \/ Rabs (round beta fexp rnd x) = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x = 0%R \/ (x > 0)%R
ulp (round beta fexp rnd x) = ulp x \/ Rabs (round beta fexp rnd x) = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:(x < 0)%R
Y:round beta fexp (Zrnd_opp rnd) (- x) = bpow (mag beta x)

Rabs (round beta fexp rnd x) = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x = 0%R \/ (x > 0)%R
ulp (round beta fexp rnd x) = ulp x \/ Rabs (round beta fexp rnd x) = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:(x < 0)%R
Y:round beta fexp (Zrnd_opp rnd) (- x) = bpow (mag beta x)

Rabs (round beta fexp rnd (- - x)) = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x = 0%R \/ (x > 0)%R
ulp (round beta fexp rnd x) = ulp x \/ Rabs (round beta fexp rnd x) = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:(x < 0)%R
Y:round beta fexp (Zrnd_opp rnd) (- x) = bpow (mag beta x)

Rabs (- bpow (mag beta x)) = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x = 0%R \/ (x > 0)%R
ulp (round beta fexp rnd x) = ulp x \/ Rabs (round beta fexp rnd x) = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:(x < 0)%R
Y:round beta fexp (Zrnd_opp rnd) (- x) = bpow (mag beta x)

(bpow (mag beta x) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x = 0%R \/ (x > 0)%R
ulp (round beta fexp rnd x) = ulp x \/ Rabs (round beta fexp rnd x) = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x = 0%R \/ (x > 0)%R

ulp (round beta fexp rnd x) = ulp x \/ Rabs (round beta fexp rnd x) = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:x = 0%R

ulp (round beta fexp rnd x) = ulp x \/ Rabs (round beta fexp rnd x) = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:(x > 0)%R
ulp (round beta fexp rnd x) = ulp x \/ Rabs (round beta fexp rnd x) = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:(x > 0)%R

ulp (round beta fexp rnd x) = ulp x \/ Rabs (round beta fexp rnd x) = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:(x > 0)%R

ulp (round beta fexp rnd x) = ulp x \/ round beta fexp rnd x = bpow (mag beta x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:(x > 0)%R
(round beta fexp rnd x >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:(x > 0)%R

(round beta fexp rnd x >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:(x > 0)%R

F 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:(x > 0)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Not_FTZ_:Exp_not_FTZ fexp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Zx:(x > 0)%R

(0 <= x)%R
now apply Rlt_le. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall rnd : R -> Z, Valid_rnd rnd -> forall x : R, (x <= succ (round beta fexp rnd x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall rnd : R -> Z, Valid_rnd rnd -> forall x : R, (x <= succ (round beta fexp rnd x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Vrnd:Valid_rnd rnd
x:R

(x <= succ (round beta fexp rnd x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Vrnd:Valid_rnd rnd
x:R

(x <= round beta fexp Zceil x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Vrnd:Valid_rnd rnd
x:R
(round beta fexp Zceil x <= succ (round beta fexp rnd x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Vrnd:Valid_rnd rnd
x:R

(x <= round beta fexp Zceil x)%R
now apply round_UP_pt.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Vrnd:Valid_rnd rnd
x:R

(round beta fexp Zceil x <= succ (round beta fexp rnd x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Vrnd:Valid_rnd rnd
x:R
Hr:round beta fexp rnd x = round beta fexp Zfloor x

(round beta fexp Zceil x <= succ (round beta fexp Zfloor x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Vrnd:Valid_rnd rnd
x:R
Hr:round beta fexp rnd x = round beta fexp Zceil x
(round beta fexp Zceil x <= succ (round beta fexp Zceil x))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Vrnd:Valid_rnd rnd
x:R
Hr:round beta fexp rnd x = round beta fexp Zfloor x

(round beta fexp Zceil x <= succ (round beta fexp Zfloor x))%R
now apply UP_le_succ_DN.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Vrnd:Valid_rnd rnd
x:R
Hr:round beta fexp rnd x = round beta fexp Zceil x

(round beta fexp Zceil x <= succ (round beta fexp Zceil x))%R
apply succ_ge_id. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall rnd : R -> Z, Valid_rnd rnd -> forall x : R, (pred (round beta fexp rnd x) <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall rnd : R -> Z, Valid_rnd rnd -> forall x : R, (pred (round beta fexp rnd x) <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Vrnd:Valid_rnd rnd
x:R

(pred (round beta fexp rnd x) <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Vrnd:Valid_rnd rnd
x:R

(pred (round beta fexp rnd x) <= round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Vrnd:Valid_rnd rnd
x:R
(round beta fexp Zfloor x <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Vrnd:Valid_rnd rnd
x:R

(pred (round beta fexp rnd x) <= round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Vrnd:Valid_rnd rnd
x:R
Hr:round beta fexp rnd x = round beta fexp Zfloor x

(pred (round beta fexp Zfloor x) <= round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Vrnd:Valid_rnd rnd
x:R
Hr:round beta fexp rnd x = round beta fexp Zceil x
(pred (round beta fexp Zceil x) <= round beta fexp Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
rnd:R -> Z
Vrnd:Valid_rnd rnd
x:R
Hr:round beta fexp rnd x = round beta fexp Zfloor x

(pred (round beta fexp Zfloor x) <= round beta fexp Zfloor x)%R
apply pred_le_id. Qed.
Properties of rounding to nearest and ulp
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall (choice : Z -> bool) (u v : R), F u -> (v < (u + succ u) / 2)%R -> (round beta fexp (Znearest choice) v <= u)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall (choice : Z -> bool) (u v : R), F u -> (v < (u + succ u) / 2)%R -> (round beta fexp (Znearest choice) v <= u)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R

(round beta fexp (Znearest choice) v <= u)%R
(* . *)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R

succ u = 0%R /\ u = 0%R \/ (u < succ u)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
V:succ u = 0%R /\ u = 0%R \/ (u < succ u)%R
(round beta fexp (Znearest choice) v <= u)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
P:(u < succ u)%R

succ u = 0%R /\ u = 0%R \/ (u < succ u)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
P:u = succ u
succ u = 0%R /\ u = 0%R \/ (u < succ u)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
V:succ u = 0%R /\ u = 0%R \/ (u < succ u)%R
(round beta fexp (Znearest choice) v <= u)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
P:u = succ u

succ u = 0%R /\ u = 0%R \/ (u < succ u)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
V:succ u = 0%R /\ u = 0%R \/ (u < succ u)%R
(round beta fexp (Znearest choice) v <= u)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
P:u = succ u
Zu:u = 0%R

succ u = 0%R /\ u = 0%R \/ (u < succ u)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
P:u = succ u
Zu:u <> 0%R
succ u = 0%R /\ u = 0%R \/ (u < succ u)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
V:succ u = 0%R /\ u = 0%R \/ (u < succ u)%R
(round beta fexp (Znearest choice) v <= u)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
P:u = succ u
Zu:u = 0%R

succ u = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
P:u = succ u
Zu:u <> 0%R
succ u = 0%R /\ u = 0%R \/ (u < succ u)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
V:succ u = 0%R /\ u = 0%R \/ (u < succ u)%R
(round beta fexp (Znearest choice) v <= u)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
P:u = succ u
Zu:u <> 0%R

succ u = 0%R /\ u = 0%R \/ (u < succ u)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
V:succ u = 0%R /\ u = 0%R \/ (u < succ u)%R
(round beta fexp (Znearest choice) v <= u)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
V:succ u = 0%R /\ u = 0%R \/ (u < succ u)%R

(round beta fexp (Znearest choice) v <= u)%R
(* *)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
V1:succ u = 0%R
V2:u = 0%R

(round beta fexp (Znearest choice) v <= u)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
V:(u < succ u)%R
(round beta fexp (Znearest choice) v <= u)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
V1:succ u = 0%R
V2:u = 0%R

F 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
V1:succ u = 0%R
V2:u = 0%R
(v <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
V:(u < succ u)%R
(round beta fexp (Znearest choice) v <= u)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
V1:succ u = 0%R
V2:u = 0%R

(v <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
V:(u < succ u)%R
(round beta fexp (Znearest choice) v <= u)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
V1:succ u = 0%R
V2:u = 0%R

((u + succ u) / 2 <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
V:(u < succ u)%R
(round beta fexp (Znearest choice) v <= u)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
V:(u < succ u)%R

(round beta fexp (Znearest choice) v <= u)%R
(* *)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
V:(u < succ u)%R
T:(u < (u + succ u) / 2 < succ u)%R

(round beta fexp (Znearest choice) v <= u)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
V:(u < succ u)%R
T1:(u < (u + succ u) / 2)%R
T2:((u + succ u) / 2 < succ u)%R

(round beta fexp (Znearest choice) v <= u)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
V:(u < succ u)%R
T1:(u < (u + succ u) / 2)%R
T2:((u + succ u) / 2 < succ u)%R

Rnd_N_pt F v (round beta fexp (Znearest choice) v)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
V:(u < succ u)%R
T1:(u < (u + succ u) / 2)%R
T2:((u + succ u) / 2 < succ u)%R
Rnd_N_pt F ((u + succ u) / 2) u
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
V:(u < succ u)%R
T1:(u < (u + succ u) / 2)%R
T2:((u + succ u) / 2 < succ u)%R

Rnd_N_pt F ((u + succ u) / 2) u
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
V:(u < succ u)%R
T1:(u < (u + succ u) / 2)%R
T2:((u + succ u) / 2 < succ u)%R

Rnd_DN_pt F ((u + succ u) / 2) u
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
V:(u < succ u)%R
T1:(u < (u + succ u) / 2)%R
T2:((u + succ u) / 2 < succ u)%R
Rnd_UP_pt F ((u + succ u) / 2) (succ u)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
V:(u < succ u)%R
T1:(u < (u + succ u) / 2)%R
T2:((u + succ u) / 2 < succ u)%R
((u + succ u) / 2 - u <= succ u - (u + succ u) / 2)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
V:(u < succ u)%R
T1:(u < (u + succ u) / 2)%R
T2:((u + succ u) / 2 < succ u)%R

Rnd_DN_pt F ((u + succ u) / 2) (round beta fexp Zfloor ((u + succ u) / 2))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
V:(u < succ u)%R
T1:(u < (u + succ u) / 2)%R
T2:((u + succ u) / 2 < succ u)%R
round beta fexp Zfloor ((u + succ u) / 2) = u
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
V:(u < succ u)%R
T1:(u < (u + succ u) / 2)%R
T2:((u + succ u) / 2 < succ u)%R
Rnd_UP_pt F ((u + succ u) / 2) (succ u)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
V:(u < succ u)%R
T1:(u < (u + succ u) / 2)%R
T2:((u + succ u) / 2 < succ u)%R
((u + succ u) / 2 - u <= succ u - (u + succ u) / 2)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
V:(u < succ u)%R
T1:(u < (u + succ u) / 2)%R
T2:((u + succ u) / 2 < succ u)%R

round beta fexp Zfloor ((u + succ u) / 2) = u
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
V:(u < succ u)%R
T1:(u < (u + succ u) / 2)%R
T2:((u + succ u) / 2 < succ u)%R
Rnd_UP_pt F ((u + succ u) / 2) (succ u)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
V:(u < succ u)%R
T1:(u < (u + succ u) / 2)%R
T2:((u + succ u) / 2 < succ u)%R
((u + succ u) / 2 - u <= succ u - (u + succ u) / 2)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
V:(u < succ u)%R
T1:(u < (u + succ u) / 2)%R
T2:((u + succ u) / 2 < succ u)%R

(u <= (u + succ u) / 2 < succ u)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
V:(u < succ u)%R
T1:(u < (u + succ u) / 2)%R
T2:((u + succ u) / 2 < succ u)%R
Rnd_UP_pt F ((u + succ u) / 2) (succ u)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
V:(u < succ u)%R
T1:(u < (u + succ u) / 2)%R
T2:((u + succ u) / 2 < succ u)%R
((u + succ u) / 2 - u <= succ u - (u + succ u) / 2)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
V:(u < succ u)%R
T1:(u < (u + succ u) / 2)%R
T2:((u + succ u) / 2 < succ u)%R

Rnd_UP_pt F ((u + succ u) / 2) (succ u)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
V:(u < succ u)%R
T1:(u < (u + succ u) / 2)%R
T2:((u + succ u) / 2 < succ u)%R
((u + succ u) / 2 - u <= succ u - (u + succ u) / 2)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
V:(u < succ u)%R
T1:(u < (u + succ u) / 2)%R
T2:((u + succ u) / 2 < succ u)%R

Rnd_UP_pt F ((u + succ u) / 2) (round beta fexp Zceil ((u + succ u) / 2))
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
V:(u < succ u)%R
T1:(u < (u + succ u) / 2)%R
T2:((u + succ u) / 2 < succ u)%R
round beta fexp Zceil ((u + succ u) / 2) = succ u
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
V:(u < succ u)%R
T1:(u < (u + succ u) / 2)%R
T2:((u + succ u) / 2 < succ u)%R
((u + succ u) / 2 - u <= succ u - (u + succ u) / 2)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
V:(u < succ u)%R
T1:(u < (u + succ u) / 2)%R
T2:((u + succ u) / 2 < succ u)%R

round beta fexp Zceil ((u + succ u) / 2) = succ u
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
V:(u < succ u)%R
T1:(u < (u + succ u) / 2)%R
T2:((u + succ u) / 2 < succ u)%R
((u + succ u) / 2 - u <= succ u - (u + succ u) / 2)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
V:(u < succ u)%R
T1:(u < (u + succ u) / 2)%R
T2:((u + succ u) / 2 < succ u)%R

F (succ u)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
V:(u < succ u)%R
T1:(u < (u + succ u) / 2)%R
T2:((u + succ u) / 2 < succ u)%R
(pred (succ u) < (u + succ u) / 2 <= succ u)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
V:(u < succ u)%R
T1:(u < (u + succ u) / 2)%R
T2:((u + succ u) / 2 < succ u)%R
((u + succ u) / 2 - u <= succ u - (u + succ u) / 2)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
V:(u < succ u)%R
T1:(u < (u + succ u) / 2)%R
T2:((u + succ u) / 2 < succ u)%R

(pred (succ u) < (u + succ u) / 2 <= succ u)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
V:(u < succ u)%R
T1:(u < (u + succ u) / 2)%R
T2:((u + succ u) / 2 < succ u)%R
((u + succ u) / 2 - u <= succ u - (u + succ u) / 2)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
V:(u < succ u)%R
T1:(u < (u + succ u) / 2)%R
T2:((u + succ u) / 2 < succ u)%R

(u < (u + succ u) / 2 <= succ u)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
V:(u < succ u)%R
T1:(u < (u + succ u) / 2)%R
T2:((u + succ u) / 2 < succ u)%R
((u + succ u) / 2 - u <= succ u - (u + succ u) / 2)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:(v < (u + succ u) / 2)%R
V:(u < succ u)%R
T1:(u < (u + succ u) / 2)%R
T2:((u + succ u) / 2 < succ u)%R

((u + succ u) / 2 - u <= succ u - (u + succ u) / 2)%R
right; field. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall (choice : Z -> bool) (u v : R), F u -> ((u + pred u) / 2 < v)%R -> (u <= round beta fexp (Znearest choice) v)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall (choice : Z -> bool) (u v : R), F u -> ((u + pred u) / 2 < v)%R -> (u <= round beta fexp (Znearest choice) v)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:((u + pred u) / 2 < v)%R

(u <= round beta fexp (Znearest choice) v)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:((u + pred u) / 2 < v)%R

(u <= round beta fexp (Znearest choice) (- - v))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:((u + pred u) / 2 < v)%R

(u <= - round beta fexp (Znearest (fun t : Z => negb (choice (- (t + 1))%Z))) (- v))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:((u + pred u) / 2 < v)%R

(- - u <= - round beta fexp (Znearest (fun t : Z => negb (choice (- (t + 1))%Z))) (- v))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:((u + pred u) / 2 < v)%R

(round beta fexp (Znearest (fun t : Z => negb (choice (- (t + 1))%Z))) (- v) <= - u)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:((u + pred u) / 2 < v)%R

F (- u)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:((u + pred u) / 2 < v)%R
(- v < (- u + succ (- u)) / 2)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:((u + pred u) / 2 < v)%R

(- v < (- u + succ (- u)) / 2)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:((u + pred u) / 2 < v)%R

(- ((- u + succ (- u)) / 2) < - - v)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:((u + pred u) / 2 < v)%R

(- ((- u + succ (- u)) / 2) < v)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:((u + pred u) / 2 < v)%R

(- ((- u + succ (- u)) / 2) <= (u + pred u) / 2)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Fu:F u
H:((u + pred u) / 2 < v)%R

(- ((- u + succ (- u)) / 2) <= (u + - succ (- u)) / 2)%R
right; field. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall (choice : Z -> bool) (u v : R), F u -> (u <= round beta fexp (Znearest choice) v)%R -> ((u + pred u) / 2 <= v)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall (choice : Z -> bool) (u v : R), F u -> (u <= round beta fexp (Znearest choice) v)%R -> ((u + pred u) / 2 <= v)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R

((u + pred u) / 2 <= v)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R

u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
((u + pred u) / 2 <= v)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
Zu:u = 0%R

u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
Zu:u <> 0%R
u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
((u + pred u) / 2 <= v)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
Zu:u = 0%R

forall z : Z, negligible_exp = Some z -> u = 0%R /\ Some z = None \/ (pred u < u)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
Zu:u = 0%R
negligible_exp = None -> u = 0%R /\ None = None \/ (pred u < u)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
Zu:u <> 0%R
u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
((u + pred u) / 2 <= v)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
Zu:u = 0%R
n:Z
Hn:negligible_exp = Some n

(pred u < u)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
Zu:u = 0%R
negligible_exp = None -> u = 0%R /\ None = None \/ (pred u < u)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
Zu:u <> 0%R
u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
((u + pred u) / 2 <= v)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
Zu:u = 0%R
n:Z
Hn:negligible_exp = Some n

(- ulp 0 < 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
Zu:u = 0%R
negligible_exp = None -> u = 0%R /\ None = None \/ (pred u < u)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
Zu:u <> 0%R
u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
((u + pred u) / 2 <= v)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
Zu:u = 0%R
n:Z
Hn:negligible_exp = Some n

(- bpow (fexp n) < 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
Zu:u = 0%R
negligible_exp = None -> u = 0%R /\ None = None \/ (pred u < u)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
Zu:u <> 0%R
u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
((u + pred u) / 2 <= v)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
Zu:u = 0%R
n:Z
Hn:negligible_exp = Some n

(- bpow (fexp n) < - 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
Zu:u = 0%R
negligible_exp = None -> u = 0%R /\ None = None \/ (pred u < u)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
Zu:u <> 0%R
u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
((u + pred u) / 2 <= v)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
Zu:u = 0%R

negligible_exp = None -> u = 0%R /\ None = None \/ (pred u < u)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
Zu:u <> 0%R
u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
((u + pred u) / 2 <= v)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
Zu:u <> 0%R

u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
((u + pred u) / 2 <= v)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
Zu:u <> 0%R

(pred u < u)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
((u + pred u) / 2 <= v)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R

((u + pred u) / 2 <= v)%R
(* *)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R

u = 0%R /\ negligible_exp = None -> ((u + pred u) / 2 <= v)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
(pred u < u)%R -> ((u + pred u) / 2 <= v)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
K1:u = 0%R
K2:negligible_exp = None

((u + pred u) / 2 <= v)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
(pred u < u)%R -> ((u + pred u) / 2 <= v)%R
(* . *)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
K1:u = 0%R
K2:negligible_exp = None

((0 + - ulp 0) / 2 <= v)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
(pred u < u)%R -> ((u + pred u) / 2 <= v)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
K1:u = 0%R
K2:negligible_exp = None

((0 + - 0) / 2 <= v)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
(pred u < u)%R -> ((u + pred u) / 2 <= v)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
K1:u = 0%R
K2:negligible_exp = None

(0 <= v)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
(pred u < u)%R -> ((u + pred u) / 2 <= v)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
K1:u = 0%R
K2:negligible_exp = None

(v < 0)%R -> (0 <= v)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
(pred u < u)%R -> ((u + pred u) / 2 <= v)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
K1:u = 0%R
K2:negligible_exp = None
H3:(v < 0)%R

~ (u <= round beta fexp (Znearest choice) v)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
(pred u < u)%R -> ((u + pred u) / 2 <= v)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
K1:u = 0%R
K2:negligible_exp = None
H3:(v < 0)%R

(round beta fexp (Znearest choice) v < 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
(pred u < u)%R -> ((u + pred u) / 2 <= v)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
K1:u = 0%R
K2:negligible_exp = None
H3:(v < 0)%R

(round beta fexp (Znearest choice) v <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
K1:u = 0%R
K2:negligible_exp = None
H3:(v < 0)%R
H4:(round beta fexp (Znearest choice) v <= 0)%R
(round beta fexp (Znearest choice) v < 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
(pred u < u)%R -> ((u + pred u) / 2 <= v)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
K1:u = 0%R
K2:negligible_exp = None
H3:(v < 0)%R

F 0
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
K1:u = 0%R
K2:negligible_exp = None
H3:(v < 0)%R
(v <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
K1:u = 0%R
K2:negligible_exp = None
H3:(v < 0)%R
H4:(round beta fexp (Znearest choice) v <= 0)%R
(round beta fexp (Znearest choice) v < 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
(pred u < u)%R -> ((u + pred u) / 2 <= v)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
K1:u = 0%R
K2:negligible_exp = None
H3:(v < 0)%R

(v <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
K1:u = 0%R
K2:negligible_exp = None
H3:(v < 0)%R
H4:(round beta fexp (Znearest choice) v <= 0)%R
(round beta fexp (Znearest choice) v < 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
(pred u < u)%R -> ((u + pred u) / 2 <= v)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
K1:u = 0%R
K2:negligible_exp = None
H3:(v < 0)%R
H4:(round beta fexp (Znearest choice) v <= 0)%R

(round beta fexp (Znearest choice) v < 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
(pred u < u)%R -> ((u + pred u) / 2 <= v)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
K1:u = 0%R
K2:negligible_exp = None
H3:(v < 0)%R
H4:(round beta fexp (Znearest choice) v <= 0)%R

round beta fexp (Znearest choice) v = 0%R -> (round beta fexp (Znearest choice) v < 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
(pred u < u)%R -> ((u + pred u) / 2 <= v)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
K1:u = 0%R
K2:negligible_exp = None
H3:(v < 0)%R
H4:(round beta fexp (Znearest choice) v <= 0)%R
H5:round beta fexp (Znearest choice) v = 0%R

(round beta fexp (Znearest choice) v < 0)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
(pred u < u)%R -> ((u + pred u) / 2 <= v)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
K1:u = 0%R
K2:negligible_exp = None
H3:(v < 0)%R
H4:(round beta fexp (Znearest choice) v <= 0)%R
H5:round beta fexp (Znearest choice) v = 0%R

v = 0%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
(pred u < u)%R -> ((u + pred u) / 2 <= v)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R

(pred u < u)%R -> ((u + pred u) / 2 <= v)%R
(* . *)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
K1:(pred u < u)%R

((u + pred u) / 2 <= v)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
K1:(pred u < u)%R

(v < (u + pred u) / 2)%R -> ((u + pred u) / 2 <= v)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
K1:(pred u < u)%R
H3:(v < (u + pred u) / 2)%R

((u + pred u) / 2 <= v)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
K1:(pred u < u)%R
H3:(v < (u + pred u) / 2)%R

~ (u <= round beta fexp (Znearest choice) v)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
K1:(pred u < u)%R
H3:(v < (u + pred u) / 2)%R

(round beta fexp (Znearest choice) v < u)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
K1:(pred u < u)%R
H3:(v < (u + pred u) / 2)%R

(round beta fexp (Znearest choice) v <= pred u)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
K1:(pred u < u)%R
H3:(v < (u + pred u) / 2)%R

F (pred u)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
K1:(pred u < u)%R
H3:(v < (u + pred u) / 2)%R
(v < (pred u + succ (pred u)) / 2)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
K1:(pred u < u)%R
H3:(v < (u + pred u) / 2)%R

(v < (pred u + succ (pred u)) / 2)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
K1:(pred u < u)%R
H3:(v < (u + pred u) / 2)%R

(v < (pred u + u) / 2)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(u <= round beta fexp (Znearest choice) v)%R
K:u = 0%R /\ negligible_exp = None \/ (pred u < u)%R
K1:(pred u < u)%R
H3:(v < (u + pred u) / 2)%R

((u + pred u) / 2 <= (pred u + u) / 2)%R
right; f_equal; ring. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall (choice : Z -> bool) (u v : R), F u -> (round beta fexp (Znearest choice) v <= u)%R -> (v <= (u + succ u) / 2)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall (choice : Z -> bool) (u v : R), F u -> (round beta fexp (Znearest choice) v <= u)%R -> (v <= (u + succ u) / 2)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(round beta fexp (Znearest choice) v <= u)%R

(v <= (u + succ u) / 2)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(round beta fexp (Znearest choice) v <= u)%R

(- ((u + succ u) / 2) <= - v)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(round beta fexp (Znearest choice) v <= u)%R

(- ((u + succ u) / 2) <= (- u + pred (- u)) / 2)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(round beta fexp (Znearest choice) v <= u)%R
((- u + pred (- u)) / 2 <= - v)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(round beta fexp (Znearest choice) v <= u)%R

((- u + pred (- u)) / 2 <= - v)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(round beta fexp (Znearest choice) v <= u)%R

F (- u)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(round beta fexp (Znearest choice) v <= u)%R
(- u <= round beta fexp (Znearest (fun t : Z => negb (choice (- (t + 1))%Z))) (- v))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(round beta fexp (Znearest choice) v <= u)%R

(- u <= round beta fexp (Znearest (fun t : Z => negb (choice (- (t + 1))%Z))) (- v))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(round beta fexp (Znearest choice) v <= u)%R

(- u <= - - round beta fexp (Znearest (fun t : Z => negb (choice (- (t + 1))%Z))) (- v))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
u, v:R
Hu:F u
H2:(round beta fexp (Znearest choice) v <= u)%R

(- u <= - round beta fexp (Znearest choice) v)%R
apply Ropp_le_contravar; easy. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall (choice : Z -> bool) (x : R), let d := round beta fexp Zfloor x in let u := round beta fexp Zceil x in (x < (d + u) / 2)%R -> round beta fexp (Znearest choice) x = d
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall (choice : Z -> bool) (x : R), let d := round beta fexp Zfloor x in let u := round beta fexp Zceil x in (x < (d + u) / 2)%R -> round beta fexp (Znearest choice) x = d
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
d:=round beta fexp Zfloor x:R
u:=round beta fexp Zceil x:R
H:(x < (d + u) / 2)%R

round beta fexp (Znearest choice) x = d
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
d:=round beta fexp Zfloor x:R
u:=round beta fexp Zceil x:R
H:(x < (d + u) / 2)%R

(round beta fexp (Znearest choice) x <= d)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
d:=round beta fexp Zfloor x:R
u:=round beta fexp Zceil x:R
H:(x < (d + u) / 2)%R
(d <= round beta fexp (Znearest choice) x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
d:=round beta fexp Zfloor x:R
u:=round beta fexp Zceil x:R
H:(x < (d + u) / 2)%R
Fx:F x

(round beta fexp (Znearest choice) x <= d)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
d:=round beta fexp Zfloor x:R
u:=round beta fexp Zceil x:R
H:(x < (d + u) / 2)%R
Fx:~ F x
(round beta fexp (Znearest choice) x <= d)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
d:=round beta fexp Zfloor x:R
u:=round beta fexp Zceil x:R
H:(x < (d + u) / 2)%R
(d <= round beta fexp (Znearest choice) x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
d:=round beta fexp Zfloor x:R
u:=round beta fexp Zceil x:R
H:(x < (d + u) / 2)%R
Fx:F x

(x <= d)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
d:=round beta fexp Zfloor x:R
u:=round beta fexp Zceil x:R
H:(x < (d + u) / 2)%R
Fx:~ F x
(round beta fexp (Znearest choice) x <= d)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
d:=round beta fexp Zfloor x:R
u:=round beta fexp Zceil x:R
H:(x < (d + u) / 2)%R
(d <= round beta fexp (Znearest choice) x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
d:=round beta fexp Zfloor x:R
u:=round beta fexp Zceil x:R
H:(x < (d + u) / 2)%R
Fx:~ F x

(round beta fexp (Znearest choice) x <= d)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
d:=round beta fexp Zfloor x:R
u:=round beta fexp Zceil x:R
H:(x < (d + u) / 2)%R
(d <= round beta fexp (Znearest choice) x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
d:=round beta fexp Zfloor x:R
u:=round beta fexp Zceil x:R
H:(x < (d + u) / 2)%R
Fx:~ F x

F d
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
d:=round beta fexp Zfloor x:R
u:=round beta fexp Zceil x:R
H:(x < (d + u) / 2)%R
Fx:~ F x
(x < (d + succ d) / 2)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
d:=round beta fexp Zfloor x:R
u:=round beta fexp Zceil x:R
H:(x < (d + u) / 2)%R
(d <= round beta fexp (Znearest choice) x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
d:=round beta fexp Zfloor x:R
u:=round beta fexp Zceil x:R
H:(x < (d + u) / 2)%R
Fx:~ F x

(x < (d + succ d) / 2)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
d:=round beta fexp Zfloor x:R
u:=round beta fexp Zceil x:R
H:(x < (d + u) / 2)%R
(d <= round beta fexp (Znearest choice) x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
d:=round beta fexp Zfloor x:R
u:=round beta fexp Zceil x:R
H:(x < (d + u) / 2)%R
Fx:~ F x

((d + u) / 2 <= (d + succ d) / 2)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
d:=round beta fexp Zfloor x:R
u:=round beta fexp Zceil x:R
H:(x < (d + u) / 2)%R
(d <= round beta fexp (Znearest choice) x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
d:=round beta fexp Zfloor x:R
u:=round beta fexp Zceil x:R
H:(x < (d + u) / 2)%R
Fx:~ F x

u = succ d
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
d:=round beta fexp Zfloor x:R
u:=round beta fexp Zceil x:R
H:(x < (d + u) / 2)%R
(d <= round beta fexp (Znearest choice) x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
d:=round beta fexp Zfloor x:R
u:=round beta fexp Zceil x:R
H:(x < (d + u) / 2)%R

(d <= round beta fexp (Znearest choice) x)%R
apply round_ge_generic; try apply round_DN_pt... Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall (choice : Z -> bool) (x d u : R), Rnd_DN_pt F x d -> Rnd_UP_pt F x u -> (x < (d + u) / 2)%R -> round beta fexp (Znearest choice) x = d
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall (choice : Z -> bool) (x d u : R), Rnd_DN_pt F x d -> Rnd_UP_pt F x u -> (x < (d + u) / 2)%R -> round beta fexp (Znearest choice) x = d
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x, d, u:R
Hd:Rnd_DN_pt F x d
Hu:Rnd_UP_pt F x u
H:(x < (d + u) / 2)%R

round beta fexp (Znearest choice) x = d
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x, d, u:R
Hd:Rnd_DN_pt F x d
Hu:Rnd_UP_pt F x u
H:(x < (d + u) / 2)%R

d = round beta fexp Zfloor x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x, d, u:R
Hd:Rnd_DN_pt F x d
Hu:Rnd_UP_pt F x u
H:(x < (d + u) / 2)%R
H0:d = round beta fexp Zfloor x
round beta fexp (Znearest choice) x = d
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x, d, u:R
Hd:Rnd_DN_pt F x d
Hu:Rnd_UP_pt F x u
H:(x < (d + u) / 2)%R

Rnd_DN_pt F x (round beta fexp Zfloor x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x, d, u:R
Hd:Rnd_DN_pt F x d
Hu:Rnd_UP_pt F x u
H:(x < (d + u) / 2)%R
H0:d = round beta fexp Zfloor x
round beta fexp (Znearest choice) x = d
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x, d, u:R
Hd:Rnd_DN_pt F x d
Hu:Rnd_UP_pt F x u
H:(x < (d + u) / 2)%R
H0:d = round beta fexp Zfloor x

round beta fexp (Znearest choice) x = d
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x, d, u:R
Hd:Rnd_DN_pt F x d
Hu:Rnd_UP_pt F x u
H:(x < (d + u) / 2)%R
H0:d = round beta fexp Zfloor x

round beta fexp (Znearest choice) x = round beta fexp Zfloor x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x, d, u:R
Hd:Rnd_DN_pt F x d
Hu:Rnd_UP_pt F x u
H:(x < (d + u) / 2)%R
H0:d = round beta fexp Zfloor x

(x < (round beta fexp Zfloor x + round beta fexp Zceil x) / 2)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x, d, u:R
Hd:Rnd_DN_pt F x d
Hu:Rnd_UP_pt F x u
H:(x < (d + u) / 2)%R
H0:d = round beta fexp Zfloor x

(x < (d + round beta fexp Zceil x) / 2)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x, d, u:R
Hd:Rnd_DN_pt F x d
Hu:Rnd_UP_pt F x u
H:(x < (d + u) / 2)%R
H0:d = round beta fexp Zfloor x

Rnd_UP_pt F x (round beta fexp Zceil x)
apply round_UP_pt... Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall (choice : Z -> bool) (x : R), let d := round beta fexp Zfloor x in let u := round beta fexp Zceil x in ((d + u) / 2 < x)%R -> round beta fexp (Znearest choice) x = u
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall (choice : Z -> bool) (x : R), let d := round beta fexp Zfloor x in let u := round beta fexp Zceil x in ((d + u) / 2 < x)%R -> round beta fexp (Znearest choice) x = u
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
d:=round beta fexp Zfloor x:R
u:=round beta fexp Zceil x:R
H:((d + u) / 2 < x)%R

round beta fexp (Znearest choice) x = u
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
d:=round beta fexp Zfloor x:R
u:=round beta fexp Zceil x:R
H:((d + u) / 2 < x)%R

(round beta fexp (Znearest choice) x <= u)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
d:=round beta fexp Zfloor x:R
u:=round beta fexp Zceil x:R
H:((d + u) / 2 < x)%R
(u <= round beta fexp (Znearest choice) x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
d:=round beta fexp Zfloor x:R
u:=round beta fexp Zceil x:R
H:((d + u) / 2 < x)%R

(u <= round beta fexp (Znearest choice) x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
d:=round beta fexp Zfloor x:R
u:=round beta fexp Zceil x:R
H:((d + u) / 2 < x)%R
Fx:F x

(u <= round beta fexp (Znearest choice) x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
d:=round beta fexp Zfloor x:R
u:=round beta fexp Zceil x:R
H:((d + u) / 2 < x)%R
Fx:~ F x
(u <= round beta fexp (Znearest choice) x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
d:=round beta fexp Zfloor x:R
u:=round beta fexp Zceil x:R
H:((d + u) / 2 < x)%R
Fx:F x

(u <= x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
d:=round beta fexp Zfloor x:R
u:=round beta fexp Zceil x:R
H:((d + u) / 2 < x)%R
Fx:~ F x
(u <= round beta fexp (Znearest choice) x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
d:=round beta fexp Zfloor x:R
u:=round beta fexp Zceil x:R
H:((d + u) / 2 < x)%R
Fx:~ F x

(u <= round beta fexp (Znearest choice) x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
d:=round beta fexp Zfloor x:R
u:=round beta fexp Zceil x:R
H:((d + u) / 2 < x)%R
Fx:~ F x

F u
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
d:=round beta fexp Zfloor x:R
u:=round beta fexp Zceil x:R
H:((d + u) / 2 < x)%R
Fx:~ F x
((u + pred u) / 2 < x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
d:=round beta fexp Zfloor x:R
u:=round beta fexp Zceil x:R
H:((d + u) / 2 < x)%R
Fx:~ F x

((u + pred u) / 2 < x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
d:=round beta fexp Zfloor x:R
u:=round beta fexp Zceil x:R
H:((d + u) / 2 < x)%R
Fx:~ F x

((u + pred u) / 2 <= (d + u) / 2)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x:R
d:=round beta fexp Zfloor x:R
u:=round beta fexp Zceil x:R
H:((d + u) / 2 < x)%R
Fx:~ F x

pred u = d
now apply pred_UP_eq_DN. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall (choice : Z -> bool) (x d u : R), Rnd_DN_pt F x d -> Rnd_UP_pt F x u -> ((d + u) / 2 < x)%R -> round beta fexp (Znearest choice) x = u
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall (choice : Z -> bool) (x d u : R), Rnd_DN_pt F x d -> Rnd_UP_pt F x u -> ((d + u) / 2 < x)%R -> round beta fexp (Znearest choice) x = u
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x, d, u:R
Hd:Rnd_DN_pt F x d
Hu:Rnd_UP_pt F x u
H:((d + u) / 2 < x)%R

round beta fexp (Znearest choice) x = u
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x, d, u:R
Hd:Rnd_DN_pt F x d
Hu:Rnd_UP_pt F x u
H:((d + u) / 2 < x)%R

u = round beta fexp Zceil x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x, d, u:R
Hd:Rnd_DN_pt F x d
Hu:Rnd_UP_pt F x u
H:((d + u) / 2 < x)%R
H0:u = round beta fexp Zceil x
round beta fexp (Znearest choice) x = u
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x, d, u:R
Hd:Rnd_DN_pt F x d
Hu:Rnd_UP_pt F x u
H:((d + u) / 2 < x)%R

Rnd_UP_pt F x (round beta fexp Zceil x)
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x, d, u:R
Hd:Rnd_DN_pt F x d
Hu:Rnd_UP_pt F x u
H:((d + u) / 2 < x)%R
H0:u = round beta fexp Zceil x
round beta fexp (Znearest choice) x = u
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x, d, u:R
Hd:Rnd_DN_pt F x d
Hu:Rnd_UP_pt F x u
H:((d + u) / 2 < x)%R
H0:u = round beta fexp Zceil x

round beta fexp (Znearest choice) x = u
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x, d, u:R
Hd:Rnd_DN_pt F x d
Hu:Rnd_UP_pt F x u
H:((d + u) / 2 < x)%R
H0:u = round beta fexp Zceil x

round beta fexp (Znearest choice) x = round beta fexp Zceil x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x, d, u:R
Hd:Rnd_DN_pt F x d
Hu:Rnd_UP_pt F x u
H:((d + u) / 2 < x)%R
H0:u = round beta fexp Zceil x

((round beta fexp Zfloor x + round beta fexp Zceil x) / 2 < x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x, d, u:R
Hd:Rnd_DN_pt F x d
Hu:Rnd_UP_pt F x u
H:((d + u) / 2 < x)%R
H0:u = round beta fexp Zceil x

((round beta fexp Zfloor x + u) / 2 < x)%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
choice:Z -> bool
x, d, u:R
Hd:Rnd_DN_pt F x d
Hu:Rnd_UP_pt F x u
H:((d + u) / 2 < x)%R
H0:u = round beta fexp Zceil x

Rnd_DN_pt F x (round beta fexp Zfloor x)
apply round_DN_pt... Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

Monotone_exp fexp -> forall (choice1 choice2 : Z -> bool) (x : R), let rx := round beta fexp (Znearest choice2) x in (x <= round beta fexp (Znearest choice1) (rx + ulp rx))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

Monotone_exp fexp -> forall (choice1 choice2 : Z -> bool) (x : R), let rx := round beta fexp (Znearest choice2) x in (x <= round beta fexp (Znearest choice1) (rx + ulp rx))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice1, choice2:Z -> bool
x:R

let rx := round beta fexp (Znearest choice2) x in (x <= round beta fexp (Znearest choice1) (rx + ulp rx))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice1, choice2:Z -> bool
x:R

(x <= round beta fexp (Znearest choice1) (round beta fexp (Znearest choice2) x + ulp (round beta fexp (Znearest choice2) x)))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice1, choice2:Z -> bool
x:R
rx:=round beta fexp (Znearest choice2) x:R

(x <= round beta fexp (Znearest choice1) (rx + ulp rx))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice1, choice2:Z -> bool
x:R
rx:=round beta fexp (Znearest choice2) x:R
Vrnd1:Valid_rnd (Znearest choice1)

(x <= round beta fexp (Znearest choice1) (rx + ulp rx))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice1, choice2:Z -> bool
x:R
rx:=round beta fexp (Znearest choice2) x:R
Vrnd1:Valid_rnd (Znearest choice1)
Vrnd2:Valid_rnd (Znearest choice2)

(x <= round beta fexp (Znearest choice1) (rx + ulp rx))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice1, choice2:Z -> bool
x:R
rx:=round beta fexp (Znearest choice2) x:R
Vrnd1:Valid_rnd (Znearest choice1)
Vrnd2:Valid_rnd (Znearest choice2)

(succ rx <= round beta fexp (Znearest choice1) (rx + ulp rx))%R
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
Hm:Monotone_exp fexp
choice1, choice2:Z -> bool
x:R
rx:=round beta fexp (Znearest choice2) x:R
Vrnd1:Valid_rnd (Znearest choice1)
Vrnd2:Valid_rnd (Znearest choice2)

F (rx + ulp rx)
now apply generic_format_plus_ulp, generic_format_round. Qed.
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall (c1 c2 : Z -> bool) (x : R), (x - round beta fexp Zfloor x)%R <> (round beta fexp Zceil x - x)%R -> round beta fexp (Znearest c1) x = round beta fexp (Znearest c2) x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp

forall (c1 c2 : Z -> bool) (x : R), (x - round beta fexp Zfloor x)%R <> (round beta fexp Zceil x - x)%R -> round beta fexp (Znearest c1) x = round beta fexp (Znearest c2) x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
c1, c2:Z -> bool
x:R

(x - round beta fexp Zfloor x)%R <> (round beta fexp Zceil x - x)%R -> round beta fexp (Znearest c1) x = round beta fexp (Znearest c2) x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
c1, c2:Z -> bool
x:R
d:=round beta fexp Zfloor x:R
u:=round beta fexp Zceil x:R
H:(x - d)%R <> (u - x)%R

round beta fexp (Znearest c1) x = round beta fexp (Znearest c2) x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
c1, c2:Z -> bool
x:R
d:=round beta fexp Zfloor x:R
u:=round beta fexp Zceil x:R
H:(x - d)%R <> (u - x)%R
L:((d + u) / 2 <= x)%R

round beta fexp (Znearest c1) x = round beta fexp (Znearest c2) x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
c1, c2:Z -> bool
x:R
d:=round beta fexp Zfloor x:R
u:=round beta fexp Zceil x:R
H:(x - d)%R <> (u - x)%R
L:(x < (d + u) / 2)%R
round beta fexp (Znearest c1) x = round beta fexp (Znearest c2) x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
c1, c2:Z -> bool
x:R
d:=round beta fexp Zfloor x:R
u:=round beta fexp Zceil x:R
H:(x - d)%R <> (u - x)%R
L:((d + u) / 2 <= x)%R

round beta fexp (Znearest c1) x = round beta fexp (Znearest c2) x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
c1, c2:Z -> bool
x:R
d:=round beta fexp Zfloor x:R
u:=round beta fexp Zceil x:R
H:(x - d)%R <> (u - x)%R
L:((d + u) / 2 < x)%R

round beta fexp (Znearest c1) x = round beta fexp (Znearest c2) x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
c1, c2:Z -> bool
x:R
d:=round beta fexp Zfloor x:R
u:=round beta fexp Zceil x:R
H:(x - d)%R <> (u - x)%R
L:((d + u) / 2)%R = x
round beta fexp (Znearest c1) x = round beta fexp (Znearest c2) x
beta:radix
fexp:Z -> Z
valid_exp:Valid_exp fexp
c1, c2:Z -> bool
x:R
d:=round beta fexp Zfloor x:R
u:=round beta fexp Zceil x:R
H:(x - d)%R <> (u - x)%R
L:((d + u) / 2)%R = x

round beta fexp (Znearest c1) x = round beta fexp (Znearest c2) x
contradict H; rewrite <- L; field. Qed. End Fcore_ulp.