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.

What is a real number belonging to a format, and many properties.

Require Import Raux Defs Round_pred Float_prop.

Section Generic.

Variable beta : radix.

Notation bpow e := (bpow beta e).

Section Format.

Variable fexp : Z -> Z.
To be a good fexp
Class Valid_exp :=
  valid_exp :
  forall k : Z,
  ( (fexp k < k)%Z -> (fexp (k + 1) <= k)%Z ) /\
  ( (k <= fexp k)%Z ->
    (fexp (fexp k + 1) <= fexp k)%Z /\
    forall l : Z, (l <= fexp k)%Z -> fexp l = fexp k ).

Context { valid_exp_ : Valid_exp }.

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

forall k l : Z, (fexp k < k)%Z -> (k <= l)%Z -> (fexp l < l)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall k l : Z, (fexp k < k)%Z -> (k <= l)%Z -> (fexp l < l)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
k, l:Z
Hk:(fexp k < k)%Z
H:(k <= l)%Z

(fexp l < l)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
k, l:Z
Hk:(fexp k < k)%Z
H:(k <= l)%Z

~ (fexp l >= l)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
k, l:Z
Hk:(fexp k < k)%Z
H:(k <= l)%Z
Hl:(fexp l >= l)%Z

False
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
k, l:Z
Hk:(fexp k < k)%Z
H:(k <= l)%Z
Hl:(l <= fexp l)%Z

False
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
k, l:Z
Hk:(fexp k < k)%Z
H:(k <= l)%Z
Hl:(l <= fexp l)%Z
H':(k <= fexp l)%Z -> fexp k = fexp l

False
omega. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall k l : Z, (fexp k < k)%Z -> (l <= k)%Z -> (fexp l < k)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall k l : Z, (fexp k < k)%Z -> (l <= k)%Z -> (fexp l < k)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
k, l:Z
Hk:(fexp k < k)%Z
H:(l <= k)%Z

(fexp l < k)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
k, l:Z
Hk:(fexp k < k)%Z
H:(l <= k)%Z

~ (fexp l >= k)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
k, l:Z
Hk:(fexp k < k)%Z
H:(l <= k)%Z
H':(fexp l >= k)%Z

False
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
k, l:Z
Hk:(fexp k < k)%Z
H:(l <= k)%Z
H':(k <= fexp l)%Z

False
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
k, l:Z
Hk:(fexp k < k)%Z
H:(l <= k)%Z
H':(k <= fexp l)%Z
Hl:(l <= fexp l)%Z

False
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
k, l:Z
Hk:(fexp k < k)%Z
H:(l <= k)%Z
H':(k <= fexp l)%Z
Hl:(fexp (fexp l + 1) <= fexp l)%Z /\ (forall l0 : Z, (l0 <= fexp l)%Z -> fexp l0 = fexp l)

False
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
k, l:Z
Hk:(fexp k < k)%Z
H:(l <= k)%Z
H':(k <= fexp l)%Z
Hl:(fexp (fexp l + 1) <= fexp l)%Z /\ (forall l0 : Z, (l0 <= fexp l)%Z -> fexp l0 = fexp l)
H1:fexp k = fexp l

False
omega. Qed. Definition cexp x := fexp (mag beta x). Definition canonical (f : float beta) := Fexp f = cexp (F2R f). Definition scaled_mantissa x := (x * bpow (- cexp x))%R. Definition generic_format (x : R) := x = F2R (Float beta (Ztrunc (scaled_mantissa x)) (cexp x)).
Basic facts
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

generic_format 0
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

generic_format 0
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

0%R = F2R {| Fnum := Ztrunc (0 * bpow (- cexp 0)); Fexp := cexp 0 |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

0%R = F2R {| Fnum := Ztrunc 0; Fexp := cexp 0 |}
now rewrite Ztrunc_IZR, F2R_0. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, cexp (- x) = cexp x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, cexp (- x) = cexp x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

cexp (- x) = cexp x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

fexp (mag beta (- x)) = fexp (mag beta x)
now rewrite mag_opp. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

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

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

cexp (Rabs x) = cexp x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

fexp (mag beta (Rabs x)) = fexp (mag beta x)
now rewrite mag_abs. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, generic_format x -> exists f : float beta, x = F2R f /\ canonical f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, generic_format x -> exists f : float beta, x = F2R f /\ canonical f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:generic_format x

exists f : float beta, x = F2R f /\ canonical f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:generic_format x

exists f : float beta, F2R {| Fnum := Ztrunc (scaled_mantissa x); Fexp := cexp x |} = F2R f /\ canonical f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:generic_format x

F2R {| Fnum := Ztrunc (scaled_mantissa x); Fexp := cexp x |} = F2R ?f /\ canonical ?f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:generic_format x

canonical {| Fnum := Ztrunc (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:generic_format x

Fexp {| Fnum := Ztrunc (scaled_mantissa x); Fexp := cexp x |} = cexp (F2R {| Fnum := Ztrunc (scaled_mantissa x); Fexp := cexp x |})
now rewrite <- Hx. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall e : Z, (fexp (e + 1) <= e)%Z -> generic_format (bpow e)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall e : Z, (fexp (e + 1) <= e)%Z -> generic_format (bpow e)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
e:Z
H:(fexp (e + 1) <= e)%Z

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

bpow e = F2R {| Fnum := Ztrunc (bpow e * bpow (- fexp (mag beta (bpow e)))); Fexp := fexp (mag beta (bpow e)) |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
e:Z
H:(fexp (e + 1) <= e)%Z

bpow e = F2R {| Fnum := Ztrunc (bpow e * bpow (- fexp (e + 1))); Fexp := fexp (e + 1) |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
e:Z
H:(fexp (e + 1) <= e)%Z

bpow e = F2R {| Fnum := Ztrunc (bpow (e + - fexp (e + 1))); Fexp := fexp (e + 1) |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
e:Z
H:(fexp (e + 1) <= e)%Z

bpow e = F2R {| Fnum := Ztrunc (IZR (beta ^ (e + - fexp (e + 1)))); Fexp := fexp (e + 1) |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
e:Z
H:(fexp (e + 1) <= e)%Z
(0 <= e + - fexp (e + 1))%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
e:Z
H:(fexp (e + 1) <= e)%Z

bpow e = F2R {| Fnum := beta ^ (e + - fexp (e + 1)); Fexp := fexp (e + 1) |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
e:Z
H:(fexp (e + 1) <= e)%Z
(0 <= e + - fexp (e + 1))%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
e:Z
H:(fexp (e + 1) <= e)%Z

F2R {| Fnum := 1; Fexp := e |} = F2R {| Fnum := beta ^ (e + - fexp (e + 1)); Fexp := fexp (e + 1) |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
e:Z
H:(fexp (e + 1) <= e)%Z
(0 <= e + - fexp (e + 1))%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
e:Z
H:(fexp (e + 1) <= e)%Z

F2R {| Fnum := 1 * beta ^ (e - fexp (e + 1)); Fexp := fexp (e + 1) |} = F2R {| Fnum := beta ^ (e + - fexp (e + 1)); Fexp := fexp (e + 1) |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
e:Z
H:(fexp (e + 1) <= e)%Z
(0 <= e + - fexp (e + 1))%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
e:Z
H:(fexp (e + 1) <= e)%Z

(0 <= e + - fexp (e + 1))%Z
now apply Zle_minus_le_0. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall e : Z, (fexp e <= e)%Z -> generic_format (bpow e)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall e : Z, (fexp e <= e)%Z -> generic_format (bpow e)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
e:Z
He:(fexp e <= e)%Z

generic_format (bpow e)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
e:Z
He:(fexp e <= e)%Z

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

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

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

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

(e <= fexp e)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
e:Z
He:(fexp e <= e)%Z
H:fexp e = e

(e <= e)%Z
apply Z.le_refl. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall m e : Z, (m <> 0%Z -> (cexp (F2R {| Fnum := m; Fexp := e |}) <= e)%Z) -> generic_format (F2R {| Fnum := m; Fexp := e |})
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall m e : Z, (m <> 0%Z -> (cexp (F2R {| Fnum := m; Fexp := e |}) <= e)%Z) -> generic_format (F2R {| Fnum := m; Fexp := e |})
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
m, e:Z

(m <> 0%Z -> (cexp (F2R {| Fnum := m; Fexp := e |}) <= e)%Z) -> generic_format (F2R {| Fnum := m; Fexp := e |})
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
m, e:Z
Zm:m = 0%Z

(m <> 0%Z -> (cexp (F2R {| Fnum := m; Fexp := e |}) <= e)%Z) -> generic_format (F2R {| Fnum := m; Fexp := e |})
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
m, e:Z
Zm:m <> 0%Z
(m <> 0%Z -> (cexp (F2R {| Fnum := m; Fexp := e |}) <= e)%Z) -> generic_format (F2R {| Fnum := m; Fexp := e |})
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
m, e:Z
Zm:m = 0%Z

generic_format (F2R {| Fnum := m; Fexp := e |})
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
m, e:Z
Zm:m <> 0%Z
(m <> 0%Z -> (cexp (F2R {| Fnum := m; Fexp := e |}) <= e)%Z) -> generic_format (F2R {| Fnum := m; Fexp := e |})
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
m, e:Z
Zm:m = 0%Z

generic_format 0
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
m, e:Z
Zm:m <> 0%Z
(m <> 0%Z -> (cexp (F2R {| Fnum := m; Fexp := e |}) <= e)%Z) -> generic_format (F2R {| Fnum := m; Fexp := e |})
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
m, e:Z
Zm:m <> 0%Z

(m <> 0%Z -> (cexp (F2R {| Fnum := m; Fexp := e |}) <= e)%Z) -> generic_format (F2R {| Fnum := m; Fexp := e |})
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
m, e:Z
Zm:m <> 0%Z

(m <> 0%Z -> (cexp (F2R {| Fnum := m; Fexp := e |}) <= e)%Z) -> F2R {| Fnum := m; Fexp := e |} = F2R {| Fnum := Ztrunc (F2R {| Fnum := m; Fexp := e |} * bpow (- cexp (F2R {| Fnum := m; Fexp := e |}))); Fexp := cexp (F2R {| Fnum := m; Fexp := e |}) |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
m, e:Z
Zm:m <> 0%Z
e':=cexp (F2R {| Fnum := m; Fexp := e |}):Z

(m <> 0%Z -> (e' <= e)%Z) -> F2R {| Fnum := m; Fexp := e |} = F2R {| Fnum := Ztrunc (F2R {| Fnum := m; Fexp := e |} * bpow (- e')); Fexp := e' |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
m, e:Z
Zm:m <> 0%Z
e':=cexp (F2R {| Fnum := m; Fexp := e |}):Z
He:m <> 0%Z -> (e' <= e)%Z

F2R {| Fnum := m; Fexp := e |} = F2R {| Fnum := Ztrunc (F2R {| Fnum := m; Fexp := e |} * bpow (- e')); Fexp := e' |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
m, e:Z
Zm:m <> 0%Z
e':=cexp (F2R {| Fnum := m; Fexp := e |}):Z
He:(e' <= e)%Z

F2R {| Fnum := m; Fexp := e |} = F2R {| Fnum := Ztrunc (F2R {| Fnum := m; Fexp := e |} * bpow (- e')); Fexp := e' |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
m, e:Z
Zm:m <> 0%Z
e':=cexp (F2R {| Fnum := m; Fexp := e |}):Z
He:(e' <= e)%Z

F2R {| Fnum := m; Fexp := e |} = F2R {| Fnum := Ztrunc (IZR (Fnum {| Fnum := m; Fexp := e |}) * bpow (Fexp {| Fnum := m; Fexp := e |}) * bpow (- e')); Fexp := e' |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
m, e:Z
Zm:m <> 0%Z
e':=cexp (F2R {| Fnum := m; Fexp := e |}):Z
He:(e' <= e)%Z

F2R {| Fnum := m; Fexp := e |} = F2R {| Fnum := Ztrunc (IZR m * bpow e * bpow (- e')); Fexp := e' |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
m, e:Z
Zm:m <> 0%Z
e':=cexp (F2R {| Fnum := m; Fexp := e |}):Z
He:(e' <= e)%Z

F2R {| Fnum := m * beta ^ (e - e'); Fexp := e' |} = F2R {| Fnum := Ztrunc (IZR m * bpow e * bpow (- e')); Fexp := e' |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
m, e:Z
Zm:m <> 0%Z
e':=cexp (F2R {| Fnum := m; Fexp := e |}):Z
He:(e' <= e)%Z

(m * beta ^ (e - e'))%Z = Ztrunc (IZR m * bpow e * bpow (- e'))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
m, e:Z
Zm:m <> 0%Z
e':=cexp (F2R {| Fnum := m; Fexp := e |}):Z
He:(e' <= e)%Z

(m * beta ^ (e - e'))%Z = Ztrunc (IZR (m * beta ^ (e + - e')))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
m, e:Z
Zm:m <> 0%Z
e':=cexp (F2R {| Fnum := m; Fexp := e |}):Z
He:(e' <= e)%Z
(0 <= e + - e')%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
m, e:Z
Zm:m <> 0%Z
e':=cexp (F2R {| Fnum := m; Fexp := e |}):Z
He:(e' <= e)%Z

(0 <= e + - e')%Z
now apply Zle_left. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall (x : R) (f : float beta), F2R f = x -> (x <> 0%R -> (cexp x <= Fexp f)%Z) -> generic_format x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall (x : R) (f : float beta), F2R f = x -> (x <> 0%R -> (cexp x <= Fexp f)%Z) -> generic_format x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:float beta
H1:F2R f = x
H2:x <> 0%R -> (cexp x <= Fexp f)%Z

generic_format x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
m, e:Z
H1:F2R {| Fnum := m; Fexp := e |} = x
H2:x <> 0%R -> (cexp x <= Fexp {| Fnum := m; Fexp := e |})%Z

generic_format (F2R {| Fnum := m; Fexp := e |})
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
m, e:Z
H1:F2R {| Fnum := m; Fexp := e |} = x
H2:x <> 0%R -> (cexp x <= Fexp {| Fnum := m; Fexp := e |})%Z

m <> 0%Z -> (cexp (F2R {| Fnum := m; Fexp := e |}) <= e)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
m, e:Z
H1:F2R {| Fnum := m; Fexp := e |} = x
H2:x <> 0%R -> (cexp x <= e)%Z
H3:m <> 0%Z

(cexp (F2R {| Fnum := m; Fexp := e |}) <= e)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
m, e:Z
H1:F2R {| Fnum := m; Fexp := e |} = x
H2:x <> 0%R -> (cexp x <= e)%Z
H3:m <> 0%Z

x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
m, e:Z
H1:F2R {| Fnum := m; Fexp := e |} = x
H2:x <> 0%R -> (cexp x <= e)%Z
H3:m <> 0%Z
Y:x = 0%R

m = 0%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
m, e:Z
H1:F2R {| Fnum := m; Fexp := e |} = x
H2:x <> 0%R -> (cexp x <= e)%Z
H3:m <> 0%Z
Y:x = 0%R

F2R {| Fnum := m; Fexp := e |} = 0%R
now rewrite H1. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall m e : Z, canonical {| Fnum := m; Fexp := e |} -> canonical {| Fnum := - m; Fexp := e |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall m e : Z, canonical {| Fnum := m; Fexp := e |} -> canonical {| Fnum := - m; Fexp := e |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
m, e:Z
H:canonical {| Fnum := m; Fexp := e |}

canonical {| Fnum := - m; Fexp := e |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
m, e:Z
H:canonical {| Fnum := m; Fexp := e |}

Fexp {| Fnum := - m; Fexp := e |} = cexp (F2R {| Fnum := - m; Fexp := e |})
now rewrite F2R_Zopp, cexp_opp. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall m e : Z, canonical {| Fnum := m; Fexp := e |} -> canonical {| Fnum := Z.abs m; Fexp := e |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall m e : Z, canonical {| Fnum := m; Fexp := e |} -> canonical {| Fnum := Z.abs m; Fexp := e |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
m, e:Z
H:canonical {| Fnum := m; Fexp := e |}

canonical {| Fnum := Z.abs m; Fexp := e |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
m, e:Z
H:canonical {| Fnum := m; Fexp := e |}

Fexp {| Fnum := Z.abs m; Fexp := e |} = cexp (F2R {| Fnum := Z.abs m; Fexp := e |})
now rewrite F2R_Zabs, cexp_abs. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

canonical {| Fnum := 0; Fexp := fexp (mag beta 0) |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

canonical {| Fnum := 0; Fexp := fexp (mag beta 0) |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

fexp (mag beta 0) = fexp (mag beta (F2R {| Fnum := 0; Fexp := fexp (mag beta 0) |}))
now rewrite F2R_0. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall f1 f2 : float beta, canonical f1 -> canonical f2 -> F2R f1 = F2R f2 -> f1 = f2
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall f1 f2 : float beta, canonical f1 -> canonical f2 -> F2R f1 = F2R f2 -> f1 = f2
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
m1, e1, m2, e2:Z

canonical {| Fnum := m1; Fexp := e1 |} -> canonical {| Fnum := m2; Fexp := e2 |} -> F2R {| Fnum := m1; Fexp := e1 |} = F2R {| Fnum := m2; Fexp := e2 |} -> {| Fnum := m1; Fexp := e1 |} = {| Fnum := m2; Fexp := e2 |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
m1, e1, m2, e2:Z

Fexp {| Fnum := m1; Fexp := e1 |} = cexp (F2R {| Fnum := m1; Fexp := e1 |}) -> Fexp {| Fnum := m2; Fexp := e2 |} = cexp (F2R {| Fnum := m2; Fexp := e2 |}) -> F2R {| Fnum := m1; Fexp := e1 |} = F2R {| Fnum := m2; Fexp := e2 |} -> {| Fnum := m1; Fexp := e1 |} = {| Fnum := m2; Fexp := e2 |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
m1, e1, m2, e2:Z

e1 = cexp (F2R {| Fnum := m1; Fexp := e1 |}) -> e2 = cexp (F2R {| Fnum := m2; Fexp := e2 |}) -> F2R {| Fnum := m1; Fexp := e1 |} = F2R {| Fnum := m2; Fexp := e2 |} -> {| Fnum := m1; Fexp := e1 |} = {| Fnum := m2; Fexp := e2 |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
m1, e1, m2, e2:Z
H1:e1 = cexp (F2R {| Fnum := m1; Fexp := e1 |})
H2:e2 = cexp (F2R {| Fnum := m2; Fexp := e2 |})
H:F2R {| Fnum := m1; Fexp := e1 |} = F2R {| Fnum := m2; Fexp := e2 |}

{| Fnum := m1; Fexp := e1 |} = {| Fnum := m2; Fexp := e2 |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
m1, e1, m2, e2:Z
H1:e1 = cexp (F2R {| Fnum := m2; Fexp := e2 |})
H2:e2 = cexp (F2R {| Fnum := m2; Fexp := e2 |})
H:F2R {| Fnum := m1; Fexp := e1 |} = F2R {| Fnum := m2; Fexp := e2 |}

{| Fnum := m1; Fexp := e1 |} = {| Fnum := m2; Fexp := e2 |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
m1, e1, m2, e2:Z
H1:e1 = e2
H2:e2 = cexp (F2R {| Fnum := m2; Fexp := e2 |})
H:F2R {| Fnum := m1; Fexp := e1 |} = F2R {| Fnum := m2; Fexp := e2 |}

{| Fnum := m1; Fexp := e1 |} = {| Fnum := m2; Fexp := e2 |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
m1, e1, m2, e2:Z
H1:e1 = e2
H:F2R {| Fnum := m1; Fexp := e1 |} = F2R {| Fnum := m2; Fexp := e2 |}

{| Fnum := m1; Fexp := e1 |} = {| Fnum := m2; Fexp := e2 |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
m1, e1, m2, e2:Z
H1:e1 = e2
H:F2R {| Fnum := m1; Fexp := e2 |} = F2R {| Fnum := m2; Fexp := e2 |}

{| Fnum := m1; Fexp := e2 |} = {| Fnum := m2; Fexp := e2 |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
m1, e1, m2, e2:Z
H1:e1 = e2
H:F2R {| Fnum := m1; Fexp := e2 |} = F2R {| Fnum := m2; Fexp := e2 |}

m1 = m2
apply eq_F2R with (1 := H). Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, generic_format x -> scaled_mantissa x = IZR (Ztrunc (scaled_mantissa x))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, generic_format x -> scaled_mantissa x = IZR (Ztrunc (scaled_mantissa x))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:generic_format x

scaled_mantissa x = IZR (Ztrunc (scaled_mantissa x))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:generic_format x

(x * bpow (- cexp x))%R = IZR (Ztrunc (x * bpow (- cexp x)))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:generic_format x

(F2R {| Fnum := Ztrunc (scaled_mantissa x); Fexp := cexp x |} * bpow (- cexp x))%R = IZR (Ztrunc (F2R {| Fnum := Ztrunc (scaled_mantissa x); Fexp := cexp x |} * bpow (- cexp x)))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:generic_format x

(IZR (Fnum {| Fnum := Ztrunc (scaled_mantissa x); Fexp := cexp x |}) * bpow (Fexp {| Fnum := Ztrunc (scaled_mantissa x); Fexp := cexp x |}) * bpow (- cexp x))%R = IZR (Ztrunc (IZR (Fnum {| Fnum := Ztrunc (scaled_mantissa x); Fexp := cexp x |}) * bpow (Fexp {| Fnum := Ztrunc (scaled_mantissa x); Fexp := cexp x |}) * bpow (- cexp x)))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:generic_format x

(IZR (Ztrunc (scaled_mantissa x)) * bpow (cexp x) * bpow (- cexp x))%R = IZR (Ztrunc (IZR (Ztrunc (scaled_mantissa x)) * bpow (cexp x) * bpow (- cexp x)))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:generic_format x

IZR (Ztrunc (scaled_mantissa x)) = IZR (Ztrunc (IZR (Ztrunc (scaled_mantissa x))))
now rewrite Ztrunc_IZR. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, (scaled_mantissa x * bpow (cexp x))%R = x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, (scaled_mantissa x * bpow (cexp x))%R = x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

(scaled_mantissa x * bpow (cexp x))%R = x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

(x * bpow (- cexp x) * bpow (cexp x))%R = x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

(x * bpow 0)%R = x
apply Rmult_1_r. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

scaled_mantissa 0 = 0%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

scaled_mantissa 0 = 0%R
apply Rmult_0_l. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, scaled_mantissa (- x) = (- scaled_mantissa x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, scaled_mantissa (- x) = (- scaled_mantissa x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

scaled_mantissa (- x) = (- scaled_mantissa x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

(- x * bpow (- cexp (- x)))%R = (- (x * bpow (- cexp x)))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

(- x * bpow (- cexp x))%R = (- (x * bpow (- cexp x)))%R
now rewrite Ropp_mult_distr_l_reverse. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

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

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

scaled_mantissa (Rabs x) = Rabs (scaled_mantissa x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

(Rabs x * bpow (- cexp (Rabs x)))%R = Rabs (x * bpow (- cexp x))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

(Rabs x * bpow (- cexp x))%R = (Rabs x * Rabs (bpow (- cexp x)))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

bpow (- cexp x) = Rabs (bpow (- cexp x))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

Rabs (bpow (- cexp x)) = bpow (- cexp x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

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

forall x : R, generic_format x -> generic_format (- x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, generic_format x -> generic_format (- x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:generic_format x

generic_format (- x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:generic_format x

(- x)%R = F2R {| Fnum := Ztrunc (scaled_mantissa (- x)); Fexp := cexp (- x) |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:generic_format x

(- x)%R = F2R {| Fnum := Ztrunc (- scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:generic_format x

(- x)%R = F2R {| Fnum := - Ztrunc (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:generic_format x

(- x)%R = (- F2R {| Fnum := Ztrunc (scaled_mantissa x); Fexp := cexp x |})%R
now apply f_equal. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, generic_format x -> generic_format (Rabs x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, generic_format x -> generic_format (Rabs x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:generic_format x

generic_format (Rabs x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:generic_format x

Rabs x = F2R {| Fnum := Ztrunc (scaled_mantissa (Rabs x)); Fexp := cexp (Rabs x) |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:generic_format x

Rabs x = F2R {| Fnum := Ztrunc (Rabs (scaled_mantissa x)); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:generic_format x

Rabs x = F2R {| Fnum := Z.abs (Ztrunc (scaled_mantissa x)); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:generic_format x

Rabs x = Rabs (F2R {| Fnum := Ztrunc (scaled_mantissa x); Fexp := cexp x |})
now apply f_equal. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, generic_format (Rabs x) -> generic_format x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, generic_format (Rabs x) -> generic_format x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

generic_format (Rabs x) -> generic_format x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

(if Rcase_abs x then (- x)%R else x) = F2R {| Fnum := Ztrunc (scaled_mantissa (if Rcase_abs x then (- x)%R else x)); Fexp := cexp (if Rcase_abs x then (- x)%R else x) |} -> x = F2R {| Fnum := Ztrunc (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

(- x)%R = F2R {| Fnum := Ztrunc (scaled_mantissa (- x)); Fexp := cexp (- x) |} -> x = F2R {| Fnum := Ztrunc (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
x = F2R {| Fnum := Ztrunc (scaled_mantissa x); Fexp := cexp x |} -> x = F2R {| Fnum := Ztrunc (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

(- x)%R = F2R {| Fnum := - Ztrunc (scaled_mantissa x); Fexp := cexp x |} -> x = F2R {| Fnum := Ztrunc (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
x = F2R {| Fnum := Ztrunc (scaled_mantissa x); Fexp := cexp x |} -> x = F2R {| Fnum := Ztrunc (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
H:(- x)%R = F2R {| Fnum := - Ztrunc (scaled_mantissa x); Fexp := cexp x |}

x = F2R {| Fnum := Ztrunc (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
x = F2R {| Fnum := Ztrunc (scaled_mantissa x); Fexp := cexp x |} -> x = F2R {| Fnum := Ztrunc (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
H:(- x)%R = F2R {| Fnum := - Ztrunc (scaled_mantissa x); Fexp := cexp x |}

(- - x)%R = F2R {| Fnum := Ztrunc (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
x = F2R {| Fnum := Ztrunc (scaled_mantissa x); Fexp := cexp x |} -> x = F2R {| Fnum := Ztrunc (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
H:(- x)%R = F2R {| Fnum := - Ztrunc (scaled_mantissa x); Fexp := cexp x |}

(- - F2R {| Fnum := Ztrunc (scaled_mantissa x); Fexp := cexp x |})%R = F2R {| Fnum := Ztrunc (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
x = F2R {| Fnum := Ztrunc (scaled_mantissa x); Fexp := cexp x |} -> x = F2R {| Fnum := Ztrunc (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

x = F2R {| Fnum := Ztrunc (scaled_mantissa x); Fexp := cexp x |} -> x = F2R {| Fnum := Ztrunc (scaled_mantissa x); Fexp := cexp x |}
easy. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall (x : R) (ex : Z), (bpow (ex - 1) <= Rabs x < bpow ex)%R -> cexp x = fexp ex
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall (x : R) (ex : Z), (bpow (ex - 1) <= Rabs x < bpow ex)%R -> cexp x = fexp ex
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Hx:(bpow (ex - 1) <= Rabs x < bpow ex)%R

cexp x = fexp ex
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Hx:(bpow (ex - 1) <= Rabs x < bpow ex)%R

fexp (mag beta x) = fexp ex
now rewrite mag_unique with (1 := Hx). Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall (x : R) (ex : Z), (bpow (ex - 1) <= x < bpow ex)%R -> cexp x = fexp ex
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall (x : R) (ex : Z), (bpow (ex - 1) <= x < bpow ex)%R -> cexp x = fexp ex
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R

cexp x = fexp ex
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R

(bpow (ex - 1) <= Rabs x < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R

(bpow (ex - 1) <= x < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R

(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R

(0 <= bpow (ex - 1))%R
apply bpow_ge_0. Qed.
Properties when the real number is "small" (kind of subnormal)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall (x : R) (ex : Z), (bpow (ex - 1) <= x < bpow ex)%R -> (ex <= fexp ex)%Z -> (0 < x * bpow (- fexp ex) < 1)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall (x : R) (ex : Z), (bpow (ex - 1) <= x < bpow ex)%R -> (ex <= fexp ex)%Z -> (0 < x * bpow (- fexp ex) < 1)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
He:(ex <= fexp ex)%Z

(0 < x * bpow (- fexp ex) < 1)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
He:(ex <= fexp ex)%Z

(0 < x * bpow (- fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
He:(ex <= fexp ex)%Z
(x * bpow (- fexp ex) < 1)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
He:(ex <= fexp ex)%Z

(0 < x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
He:(ex <= fexp ex)%Z
(0 < bpow (- fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
He:(ex <= fexp ex)%Z
(x * bpow (- fexp ex) < 1)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
He:(ex <= fexp ex)%Z

(0 < bpow (ex - 1))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
He:(ex <= fexp ex)%Z
(0 < bpow (- fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
He:(ex <= fexp ex)%Z
(x * bpow (- fexp ex) < 1)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
He:(ex <= fexp ex)%Z

(0 < bpow (- fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
He:(ex <= fexp ex)%Z
(x * bpow (- fexp ex) < 1)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
He:(ex <= fexp ex)%Z

(x * bpow (- fexp ex) < 1)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
He:(ex <= fexp ex)%Z

(0 < bpow (fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
He:(ex <= fexp ex)%Z
(x * bpow (- fexp ex) * bpow (fexp ex) < 1 * bpow (fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
He:(ex <= fexp ex)%Z

(x * bpow (- fexp ex) * bpow (fexp ex) < 1 * bpow (fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
He:(ex <= fexp ex)%Z

(x * bpow 0 < 1 * bpow (fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
He:(ex <= fexp ex)%Z

(x < bpow (fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
He:(ex <= fexp ex)%Z

(bpow ex <= bpow (fexp ex))%R
now apply bpow_le. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall (x : R) (ex : Z), (Rabs x < bpow ex)%R -> (ex <= fexp ex)%Z -> (Rabs (scaled_mantissa x) < 1)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall (x : R) (ex : Z), (Rabs x < bpow ex)%R -> (ex <= fexp ex)%Z -> (Rabs (scaled_mantissa x) < 1)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Ex:(Rabs x < bpow ex)%R
He:(ex <= fexp ex)%Z

(Rabs (scaled_mantissa x) < 1)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Ex:(Rabs x < bpow ex)%R
He:(ex <= fexp ex)%Z
Zx:x = 0%R

(Rabs (scaled_mantissa x) < 1)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Ex:(Rabs x < bpow ex)%R
He:(ex <= fexp ex)%Z
Zx:x <> 0%R
(Rabs (scaled_mantissa x) < 1)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Ex:(Rabs x < bpow ex)%R
He:(ex <= fexp ex)%Z
Zx:x = 0%R

(0 < 1)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Ex:(Rabs x < bpow ex)%R
He:(ex <= fexp ex)%Z
Zx:x <> 0%R
(Rabs (scaled_mantissa x) < 1)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Ex:(Rabs x < bpow ex)%R
He:(ex <= fexp ex)%Z
Zx:x <> 0%R

(Rabs (scaled_mantissa x) < 1)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Ex:(Rabs x < bpow ex)%R
He:(ex <= fexp ex)%Z
Zx:x <> 0%R

(scaled_mantissa (Rabs x) < 1)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Ex:(Rabs x < bpow ex)%R
He:(ex <= fexp ex)%Z
Zx:x <> 0%R

(Rabs x * bpow (- cexp (Rabs x)) < 1)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Ex:(Rabs x < bpow ex)%R
He:(ex <= fexp ex)%Z
Zx:x <> 0%R

(Rabs x * bpow (- cexp x) < 1)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Ex:(Rabs x < bpow ex)%R
He:(ex <= fexp ex)%Z
Zx:x <> 0%R

(Rabs x * bpow (- fexp (mag beta x)) < 1)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Ex:(Rabs x < bpow ex)%R
He:(ex <= fexp ex)%Z
Zx:x <> 0%R
ex':Z
Ex':x <> 0%R -> (bpow (ex' - 1) <= Rabs x < bpow ex')%R

(Rabs x * bpow (- fexp (Build_mag_prop beta x ex' Ex')) < 1)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Ex:(Rabs x < bpow ex)%R
He:(ex <= fexp ex)%Z
Zx:x <> 0%R
ex':Z
Ex':x <> 0%R -> (bpow (ex' - 1) <= Rabs x < bpow ex')%R

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

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

(ex' <= fexp ex')%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Ex:(Rabs x < bpow ex)%R
He:(ex <= fexp ex)%Z
Zx:x <> 0%R
ex':Z
Ex':(bpow (ex' - 1) <= Rabs x < bpow ex')%R

(ex' <= fexp ex)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Ex:(Rabs x < bpow ex)%R
He:(ex <= fexp ex)%Z
Zx:x <> 0%R
ex':Z
Ex':(bpow (ex' - 1) <= Rabs x < bpow ex')%R
H:(ex' <= fexp ex)%Z
(ex' <= fexp ex')%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Ex:(Rabs x < bpow ex)%R
He:(ex <= fexp ex)%Z
Zx:x <> 0%R
ex':Z
Ex':(bpow (ex' - 1) <= Rabs x < bpow ex')%R

(ex' <= ex)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Ex:(Rabs x < bpow ex)%R
He:(ex <= fexp ex)%Z
Zx:x <> 0%R
ex':Z
Ex':(bpow (ex' - 1) <= Rabs x < bpow ex')%R
H:(ex' <= fexp ex)%Z
(ex' <= fexp ex')%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Ex:(Rabs x < bpow ex)%R
He:(ex <= fexp ex)%Z
Zx:x <> 0%R
ex':Z
Ex':(bpow (ex' - 1) <= Rabs x < bpow ex')%R

(bpow (ex' - 1) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Ex:(Rabs x < bpow ex)%R
He:(ex <= fexp ex)%Z
Zx:x <> 0%R
ex':Z
Ex':(bpow (ex' - 1) <= Rabs x < bpow ex')%R
H:(ex' <= fexp ex)%Z
(ex' <= fexp ex')%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Ex:(Rabs x < bpow ex)%R
He:(ex <= fexp ex)%Z
Zx:x <> 0%R
ex':Z
Ex':(bpow (ex' - 1) <= Rabs x < bpow ex')%R
H:(ex' <= fexp ex)%Z

(ex' <= fexp ex')%Z
now rewrite (proj2 (proj2 (valid_exp _) He)). Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, (Rabs (scaled_mantissa x) < bpow (mag beta x - cexp x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, (Rabs (scaled_mantissa x) < bpow (mag beta x - cexp x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

(Rabs (scaled_mantissa x) < bpow (mag beta x - cexp x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zx:x = 0%R

(Rabs (scaled_mantissa x) < bpow (mag beta x - cexp x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zx:x <> 0%R
(Rabs (scaled_mantissa x) < bpow (mag beta x - cexp x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zx:x = 0%R

(0 < bpow (mag beta 0 - cexp 0))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zx:x <> 0%R
(Rabs (scaled_mantissa x) < bpow (mag beta x - cexp x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zx:x <> 0%R

(Rabs (scaled_mantissa x) < bpow (mag beta x - cexp x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zx:x <> 0%R

(bpow (mag beta (scaled_mantissa x)) <= bpow (mag beta x - cexp x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zx:x <> 0%R

(mag beta (scaled_mantissa x) <= mag beta x - cexp x)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zx:x <> 0%R

(mag beta (x * bpow (- cexp x)) <= mag beta x - cexp x)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zx:x <> 0%R

(mag beta x + - cexp x <= mag beta x - cexp x)%Z
apply Z.le_refl. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, x <> 0%R -> generic_format x -> (cexp x < mag beta x)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, x <> 0%R -> generic_format x -> (cexp x < mag beta x)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zx:x <> 0%R
Gx:generic_format x

(cexp x < mag beta x)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zx:x <> 0%R
Gx:generic_format x

~ (cexp x >= mag beta x)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zx:x <> 0%R
Gx:generic_format x

~ (fexp (mag beta x) >= mag beta x)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zx:x <> 0%R
Gx:generic_format x
ex:Z
Ex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R

~ (fexp ex >= ex)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zx:x <> 0%R
Gx:generic_format x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R

~ (fexp ex >= ex)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zx:x <> 0%R
Gx:generic_format x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H:(fexp ex >= ex)%Z

False
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zx:x <> 0%R
Gx:generic_format x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H:(ex <= fexp ex)%Z

False
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zx:x <> 0%R
Gx:generic_format x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H:(ex <= fexp ex)%Z

(Rabs (scaled_mantissa x) < 1)%R -> False
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Gx:generic_format x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H:(ex <= fexp ex)%Z
Zx:(Rabs (scaled_mantissa x) < 1)%R

x = 0%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Gx:generic_format x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H:(ex <= fexp ex)%Z
Zx:(Rabs (scaled_mantissa x) < 1)%R

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

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

0%Z = Ztrunc (scaled_mantissa x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Gx:generic_format x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H:(ex <= fexp ex)%Z
Zx:(Rabs (scaled_mantissa x) < 1)%R

(Z.abs (Ztrunc (scaled_mantissa x)) < 1)%Z -> 0%Z = Ztrunc (scaled_mantissa x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Gx:generic_format x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H:(ex <= fexp ex)%Z
Zx:(Rabs (scaled_mantissa x) < 1)%R
(Z.abs (Ztrunc (scaled_mantissa x)) < 1)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Gx:generic_format x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H:(ex <= fexp ex)%Z
Zx:(Rabs (scaled_mantissa x) < 1)%R

(Z.abs (Ztrunc (scaled_mantissa x)) < 1)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Gx:generic_format x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H:(ex <= fexp ex)%Z
Zx:(Rabs (scaled_mantissa x) < 1)%R

(IZR (Z.abs (Ztrunc (scaled_mantissa x))) < 1)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Gx:generic_format x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H:(ex <= fexp ex)%Z
Zx:(Rabs (scaled_mantissa x) < 1)%R

(Rabs (IZR (Ztrunc (scaled_mantissa x))) < 1)%R
now rewrite <- scaled_mantissa_generic. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall (x : R) (ex : Z), (bpow (ex - 1) <= x < bpow ex)%R -> (ex <= fexp ex)%Z -> Zfloor (x * bpow (- fexp ex)) = 0%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall (x : R) (ex : Z), (bpow (ex - 1) <= x < bpow ex)%R -> (ex <= fexp ex)%Z -> Zfloor (x * bpow (- fexp ex)) = 0%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
He:(ex <= fexp ex)%Z

Zfloor (x * bpow (- fexp ex)) = 0%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
He:(ex <= fexp ex)%Z

(0 <= x * bpow (- fexp ex) < IZR (0 + 1))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
He:(ex <= fexp ex)%Z

(0 <= x * bpow (- fexp ex) < 1)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
He:(ex <= fexp ex)%Z
H:(0 < x * bpow (- fexp ex) < 1)%R

(0 <= x * bpow (- fexp ex) < 1)%R
split ; try apply Rlt_le ; apply H. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall (x : R) (ex : Z), (bpow (ex - 1) <= x < bpow ex)%R -> (ex <= fexp ex)%Z -> Zceil (x * bpow (- fexp ex)) = 1%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall (x : R) (ex : Z), (bpow (ex - 1) <= x < bpow ex)%R -> (ex <= fexp ex)%Z -> Zceil (x * bpow (- fexp ex)) = 1%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
He:(ex <= fexp ex)%Z

Zceil (x * bpow (- fexp ex)) = 1%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
He:(ex <= fexp ex)%Z

(IZR (1 - 1) < x * bpow (- fexp ex) <= 1)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
He:(ex <= fexp ex)%Z

(0 < x * bpow (- fexp ex) <= 1)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
He:(ex <= fexp ex)%Z
H:(0 < x * bpow (- fexp ex) < 1)%R

(0 < x * bpow (- fexp ex) <= 1)%R
split ; try apply Rlt_le ; apply H. Qed.
Generic facts about any format
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall (x : R) (m : Z), let e := cexp x in (F2R {| Fnum := m; Fexp := e |} < x < F2R {| Fnum := m + 1; Fexp := e |})%R -> ~ generic_format x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall (x : R) (m : Z), let e := cexp x in (F2R {| Fnum := m; Fexp := e |} < x < F2R {| Fnum := m + 1; Fexp := e |})%R -> ~ generic_format x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
m:Z
e:=cexp x:Z
Hx:(F2R {| Fnum := m; Fexp := e |} < x)%R
Hx2:(x < F2R {| Fnum := m + 1; Fexp := e |})%R
Hf:generic_format x

False
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
m:Z
e:=cexp x:Z
Hx:(F2R {| Fnum := m; Fexp := e |} < x)%R
Hx2:(x < F2R {| Fnum := m + 1; Fexp := e |})%R
Hf:generic_format x

(F2R {| Fnum := m + 1; Fexp := e |} <= x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
m:Z
e:=cexp x:Z
Hx:(F2R {| Fnum := m; Fexp := e |} < x)%R
Hf:generic_format x

(F2R {| Fnum := m + 1; Fexp := e |} <= x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
m:Z
e:=cexp x:Z
Hx:(F2R {| Fnum := m; Fexp := e |} < x)%R
Hf:generic_format x

(F2R {| Fnum := m + 1; Fexp := e |} <= F2R {| Fnum := Ztrunc (scaled_mantissa x); Fexp := cexp x |})%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
m:Z
e:=cexp x:Z
Hx:(F2R {| Fnum := m; Fexp := e |} < x)%R
Hf:generic_format x

(F2R {| Fnum := m + 1; Fexp := e |} <= F2R {| Fnum := Ztrunc (scaled_mantissa x); Fexp := e |})%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
m:Z
e:=cexp x:Z
Hx:(F2R {| Fnum := m; Fexp := e |} < x)%R
Hf:generic_format x

(m + 1 <= Ztrunc (scaled_mantissa x))%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
m:Z
e:=cexp x:Z
Hx:(F2R {| Fnum := m; Fexp := e |} < x)%R
Hf:generic_format x

(m < Ztrunc (scaled_mantissa x))%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
m:Z
e:=cexp x:Z
Hx:(F2R {| Fnum := m; Fexp := e |} < x)%R
Hf:generic_format x

(IZR m < IZR (Ztrunc (scaled_mantissa x)))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
m:Z
e:=cexp x:Z
Hx:(F2R {| Fnum := m; Fexp := e |} < x)%R
Hf:generic_format x

(IZR m < scaled_mantissa x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
m:Z
e:=cexp x:Z
Hx:(F2R {| Fnum := m; Fexp := e |} < x)%R
Hf:generic_format x

(0 < bpow e)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
m:Z
e:=cexp x:Z
Hx:(F2R {| Fnum := m; Fexp := e |} < x)%R
Hf:generic_format x
(IZR m * bpow e < scaled_mantissa x * bpow e)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
m:Z
e:=cexp x:Z
Hx:(F2R {| Fnum := m; Fexp := e |} < x)%R
Hf:generic_format x

(IZR m * bpow e < scaled_mantissa x * bpow e)%R
now rewrite scaled_mantissa_mult_bpow. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall f : float beta, canonical f -> generic_format (F2R f)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall f : float beta, canonical f -> generic_format (F2R f)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
m, e:Z
Hf:canonical {| Fnum := m; Fexp := e |}

generic_format (F2R {| Fnum := m; Fexp := e |})
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
m, e:Z
Hf:Fexp {| Fnum := m; Fexp := e |} = cexp (F2R {| Fnum := m; Fexp := e |})

generic_format (F2R {| Fnum := m; Fexp := e |})
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
m, e:Z
Hf:e = cexp (F2R {| Fnum := m; Fexp := e |})

generic_format (F2R {| Fnum := m; Fexp := e |})
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
m, e:Z
Hf:e = cexp (F2R {| Fnum := m; Fexp := e |})

F2R {| Fnum := m; Fexp := e |} = F2R {| Fnum := Ztrunc (F2R {| Fnum := m; Fexp := e |} * bpow (- cexp (F2R {| Fnum := m; Fexp := e |}))); Fexp := cexp (F2R {| Fnum := m; Fexp := e |}) |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
m, e:Z
Hf:e = cexp (F2R {| Fnum := m; Fexp := e |})

F2R {| Fnum := m; Fexp := e |} = F2R {| Fnum := Ztrunc (F2R {| Fnum := m; Fexp := e |} * bpow (- e)); Fexp := e |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
m, e:Z
Hf:e = cexp (F2R {| Fnum := m; Fexp := e |})

m = Ztrunc (F2R {| Fnum := m; Fexp := e |} * bpow (- e))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
m, e:Z
Hf:e = cexp (F2R {| Fnum := m; Fexp := e |})

m = Ztrunc (IZR (Fnum {| Fnum := m; Fexp := e |}) * bpow (Fexp {| Fnum := m; Fexp := e |}) * bpow (- e))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
m, e:Z
Hf:e = cexp (F2R {| Fnum := m; Fexp := e |})

m = Ztrunc (IZR m * bpow e * bpow (- e))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
m, e:Z
Hf:e = cexp (F2R {| Fnum := m; Fexp := e |})

m = Ztrunc (IZR m)
now rewrite Ztrunc_IZR. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall emin : Z, (forall e : Z, (emin <= fexp e)%Z) -> forall x : R, (0 < x)%R -> generic_format x -> (bpow emin <= x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall emin : Z, (forall e : Z, (emin <= fexp e)%Z) -> forall x : R, (0 < x)%R -> generic_format x -> (bpow emin <= x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
emin:Z
Emin:forall e : Z, (emin <= fexp e)%Z
x:R
Hx:(0 < x)%R
Fx:generic_format x

(bpow emin <= x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
emin:Z
Emin:forall e : Z, (emin <= fexp e)%Z
x:R
Hx:(0 < x)%R
Fx:generic_format x

(bpow emin <= F2R {| Fnum := Ztrunc (scaled_mantissa x); Fexp := cexp x |})%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
emin:Z
Emin:forall e : Z, (emin <= fexp e)%Z
x:R
Hx:(0 < x)%R
Fx:generic_format x

(bpow emin <= bpow (fexp (mag beta x)))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
emin:Z
Emin:forall e : Z, (emin <= fexp e)%Z
x:R
Hx:(0 < x)%R
Fx:generic_format x
(bpow (fexp (mag beta x)) <= F2R {| Fnum := Ztrunc (scaled_mantissa x); Fexp := cexp x |})%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
emin:Z
Emin:forall e : Z, (emin <= fexp e)%Z
x:R
Hx:(0 < x)%R
Fx:generic_format x

(bpow (fexp (mag beta x)) <= F2R {| Fnum := Ztrunc (scaled_mantissa x); Fexp := cexp x |})%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
emin:Z
Emin:forall e : Z, (emin <= fexp e)%Z
x:R
Hx:(0 < x)%R
Fx:generic_format x

(0 < Ztrunc (scaled_mantissa x))%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
emin:Z
Emin:forall e : Z, (emin <= fexp e)%Z
x:R
Hx:(0 < x)%R
Fx:generic_format x

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

forall prec : Z, (forall e : Z, (e - prec <= fexp e)%Z) -> forall x : R, (Rabs x < bpow (prec + cexp x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
prec:Z
Hp:forall e : Z, (e - prec <= fexp e)%Z
x:R

(Rabs x < bpow (prec + cexp x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
prec:Z
Hp:forall e : Z, (e - prec <= fexp e)%Z
x:R
Hxz:x = 0%R

(Rabs x < bpow (prec + cexp x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
prec:Z
Hp:forall e : Z, (e - prec <= fexp e)%Z
x:R
Hxz:x <> 0%R
(Rabs x < bpow (prec + cexp x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
prec:Z
Hp:forall e : Z, (e - prec <= fexp e)%Z
x:R
Hxz:x = 0%R

(0 < bpow (prec + cexp 0))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
prec:Z
Hp:forall e : Z, (e - prec <= fexp e)%Z
x:R
Hxz:x <> 0%R
(Rabs x < bpow (prec + cexp x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
prec:Z
Hp:forall e : Z, (e - prec <= fexp e)%Z
x:R
Hxz:x <> 0%R

(Rabs x < bpow (prec + cexp x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
prec:Z
Hp:forall e : Z, (e - prec <= fexp e)%Z
x:R
Hxz:x <> 0%R

(Rabs x < bpow (prec + fexp (mag beta x)))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
prec:Z
Hp:forall e : Z, (e - prec <= fexp e)%Z
x:R
Hxz:x <> 0%R
ex:Z
Ex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R

(Rabs x < bpow (prec + fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
prec:Z
Hp:forall e : Z, (e - prec <= fexp e)%Z
x:R
Hxz:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R

(Rabs x < bpow (prec + fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
prec:Z
Hp:forall e : Z, (e - prec <= fexp e)%Z
x:R
Hxz:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R

(bpow ex <= bpow (prec + fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
prec:Z
Hp:forall e : Z, (e - prec <= fexp e)%Z
x:R
Hxz:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R

(ex <= prec + fexp ex)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
prec, ex:Z
Hp:(ex - prec <= fexp ex)%Z
x:R
Hxz:x <> 0%R
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R

(ex <= prec + fexp ex)%Z
omega. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall e : Z, generic_format (bpow e) -> (fexp (e + 1) <= e)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall e : Z, generic_format (bpow e) -> (fexp (e + 1) <= e)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
e:Z
He:generic_format (bpow e)

(fexp (e + 1) <= e)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
e:Z
He:generic_format (bpow e)

~ (fexp (e + 1) > e)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
e:Z
He:(fexp (e + 1) > e)%Z

~ generic_format (bpow e)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
e:Z
He:(fexp (e + 1) > e)%Z

bpow e <> (IZR (Fnum {| Fnum := Ztrunc (bpow e * bpow (- fexp (mag beta (bpow e)))); Fexp := fexp (mag beta (bpow e)) |}) * bpow (Fexp {| Fnum := Ztrunc (bpow e * bpow (- fexp (mag beta (bpow e)))); Fexp := fexp (mag beta (bpow e)) |}))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
e:Z
He:(fexp (e + 1) > e)%Z

bpow e <> (IZR (Ztrunc (bpow e * bpow (- fexp (mag beta (bpow e))))) * bpow (fexp (mag beta (bpow e))))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
e:Z
He:(fexp (e + 1) > e)%Z

bpow e <> (IZR (Ztrunc (bpow (e + - fexp (e + 1)))) * bpow (fexp (e + 1)))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
e:Z
He:(fexp (e + 1) > e)%Z

(bpow e > IZR (Ztrunc (bpow (e + - fexp (e + 1)))) * bpow (fexp (e + 1)))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
e:Z
He:(fexp (e + 1) > e)%Z

(bpow e > IZR (Zfloor (bpow (e + - fexp (e + 1)))) * bpow (fexp (e + 1)))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
e:Z
He:(fexp (e + 1) > e)%Z
(0 <= bpow (e + - fexp (e + 1)))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
e:Z
He:(fexp (e + 1) > e)%Z

(bpow e > IZR (Zfloor (bpow (e + - fexp (e + 1)))) * bpow (fexp (e + 1)))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
e:Z
He:(fexp (e + 1) > e)%Z

(bpow e > 0 * bpow (fexp (e + 1)))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
e:Z
He:(fexp (e + 1) > e)%Z
(0 <= bpow (e + - fexp (e + 1)) < IZR (0 + 1))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
e:Z
He:(fexp (e + 1) > e)%Z

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

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

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

(bpow (e + - fexp (e + 1)) < IZR (0 + 1))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
e:Z
He:(fexp (e + 1) > e)%Z

(e + - fexp (e + 1) < 0)%Z
clear -He ; omega. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall e : Z, generic_format (bpow e) -> (fexp e <= e)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall e : Z, generic_format (bpow e) -> (fexp e <= e)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
e:Z
He:generic_format (bpow e)

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

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

(fexp e <= e)%Z
omega. Qed. Section Fcore_generic_round_pos.
Rounding functions: R -> Z
Variable rnd : R -> Z.

Class Valid_rnd := {
  Zrnd_le : forall x y, (x <= y)%R -> (rnd x <= rnd y)%Z ;
  Zrnd_IZR : forall n, rnd (IZR n) = n
}.

Context { valid_rnd : Valid_rnd }.

beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd

forall x : R, rnd x = Zfloor x \/ rnd x = Zceil x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd

forall x : R, rnd x = Zfloor x \/ rnd x = Zceil x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R

rnd x = Zfloor x \/ rnd x = Zceil x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx:(rnd x <= Zfloor x)%Z

rnd x = Zfloor x \/ rnd x = Zceil x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx:(Zfloor x < rnd x)%Z
rnd x = Zfloor x \/ rnd x = Zceil x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx:(rnd x <= Zfloor x)%Z

rnd x = Zfloor x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx:(Zfloor x < rnd x)%Z
rnd x = Zfloor x \/ rnd x = Zceil x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx:(rnd x <= Zfloor x)%Z

(Zfloor x <= rnd x)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx:(Zfloor x < rnd x)%Z
rnd x = Zfloor x \/ rnd x = Zceil x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx:(rnd x <= Zfloor x)%Z

(rnd (IZR (Zfloor x)) <= rnd x)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx:(Zfloor x < rnd x)%Z
rnd x = Zfloor x \/ rnd x = Zceil x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx:(rnd x <= Zfloor x)%Z

(IZR (Zfloor x) <= x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx:(Zfloor x < rnd x)%Z
rnd x = Zfloor x \/ rnd x = Zceil x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx:(Zfloor x < rnd x)%Z

rnd x = Zfloor x \/ rnd x = Zceil x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx:(Zfloor x < rnd x)%Z

rnd x = Zceil x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx:(Zfloor x < rnd x)%Z

(rnd x <= Zceil x)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx:(Zfloor x < rnd x)%Z
(Zceil x <= rnd x)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx:(Zfloor x < rnd x)%Z

(rnd x <= rnd (IZR (Zceil x)))%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx:(Zfloor x < rnd x)%Z
(Zceil x <= rnd x)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx:(Zfloor x < rnd x)%Z

(x <= IZR (Zceil x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx:(Zfloor x < rnd x)%Z
(Zceil x <= rnd x)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx:(Zfloor x < rnd x)%Z

(Zceil x <= rnd x)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx:(Zfloor x < rnd x)%Z

(Zfloor x + 1 <= rnd x)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx:(Zfloor x < rnd x)%Z
IZR (Zfloor x) <> x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx:(Zfloor x < rnd x)%Z

IZR (Zfloor x) <> x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx:(Zfloor x < rnd x)%Z
H:IZR (Zfloor x) = x

False
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx:(Zfloor (IZR (Zfloor x)) < rnd (IZR (Zfloor x)))%Z
H:IZR (Zfloor x) = x

False
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx:(Zfloor x < Zfloor x)%Z
H:IZR (Zfloor x) = x

False
apply Z.lt_irrefl with (1 := Hx). Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd

forall x : R, rnd x = Ztrunc x \/ rnd x = Zaway x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd

forall x : R, rnd x = Ztrunc x \/ rnd x = Zaway x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R

rnd x = Ztrunc x \/ rnd x = Zaway x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R

rnd x = (if Rlt_bool x 0 then Zceil x else Zfloor x) \/ rnd x = (if Rlt_bool x 0 then Zfloor x else Zceil x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx:rnd x = Zfloor x

rnd x = Zceil x \/ rnd x = Zfloor x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx:rnd x = Zfloor x
rnd x = Zfloor x \/ rnd x = Zceil x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx:rnd x = Zceil x
rnd x = Zceil x \/ rnd x = Zfloor x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx:rnd x = Zceil x
rnd x = Zfloor x \/ rnd x = Zceil x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx:rnd x = Zfloor x

rnd x = Zfloor x \/ rnd x = Zceil x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx:rnd x = Zceil x
rnd x = Zceil x \/ rnd x = Zfloor x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx:rnd x = Zceil x
rnd x = Zfloor x \/ rnd x = Zceil x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx:rnd x = Zceil x

rnd x = Zceil x \/ rnd x = Zfloor x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx:rnd x = Zceil x
rnd x = Zfloor x \/ rnd x = Zceil x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx:rnd x = Zceil x

rnd x = Zfloor x \/ rnd x = Zceil x
now right. Qed.
the most useful one: R -> F
Definition round x :=
  F2R (Float beta (rnd (scaled_mantissa x)) (cexp x)).

beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd

forall (x : R) (ex : Z), (fexp ex < ex)%Z -> (bpow (ex - 1) <= x < bpow ex)%R -> (bpow (ex - 1) <= round x <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd

forall (x : R) (ex : Z), (fexp ex < ex)%Z -> (bpow (ex - 1) <= x < bpow ex)%R -> (bpow (ex - 1) <= round x <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R

(bpow (ex - 1) <= round x <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R

(bpow (ex - 1) <= F2R {| Fnum := rnd (x * bpow (- cexp x)); Fexp := cexp x |} <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R

(bpow (ex - 1) <= F2R {| Fnum := rnd (x * bpow (- fexp ex)); Fexp := fexp ex |} <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R

(bpow (ex - 1) <= IZR (Fnum {| Fnum := rnd (x * bpow (- fexp ex)); Fexp := fexp ex |}) * bpow (Fexp {| Fnum := rnd (x * bpow (- fexp ex)); Fexp := fexp ex |}) <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R

(bpow (ex - 1) <= IZR (rnd (x * bpow (- fexp ex))) * bpow (fexp ex) <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zfloor (x * bpow (- fexp ex))

(bpow (ex - 1) <= IZR (Zfloor (x * bpow (- fexp ex))) * bpow (fexp ex) <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))
(bpow (ex - 1) <= IZR (Zceil (x * bpow (- fexp ex))) * bpow (fexp ex) <= bpow ex)%R
(* DN *)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zfloor (x * bpow (- fexp ex))

(bpow (ex - 1) <= IZR (Zfloor (x * bpow (- fexp ex))) * bpow (fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zfloor (x * bpow (- fexp ex))
(IZR (Zfloor (x * bpow (- fexp ex))) * bpow (fexp ex) <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))
(bpow (ex - 1) <= IZR (Zceil (x * bpow (- fexp ex))) * bpow (fexp ex) <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zfloor (x * bpow (- fexp ex))

(bpow (ex - 1 + - fexp ex + fexp ex) <= IZR (Zfloor (x * bpow (- fexp ex))) * bpow (fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zfloor (x * bpow (- fexp ex))
(IZR (Zfloor (x * bpow (- fexp ex))) * bpow (fexp ex) <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))
(bpow (ex - 1) <= IZR (Zceil (x * bpow (- fexp ex))) * bpow (fexp ex) <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zfloor (x * bpow (- fexp ex))

(bpow (ex - 1 + - fexp ex) * bpow (fexp ex) <= IZR (Zfloor (x * bpow (- fexp ex))) * bpow (fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zfloor (x * bpow (- fexp ex))
(IZR (Zfloor (x * bpow (- fexp ex))) * bpow (fexp ex) <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))
(bpow (ex - 1) <= IZR (Zceil (x * bpow (- fexp ex))) * bpow (fexp ex) <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zfloor (x * bpow (- fexp ex))

(0 <= bpow (fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zfloor (x * bpow (- fexp ex))
(bpow (ex - 1 + - fexp ex) <= IZR (Zfloor (x * bpow (- fexp ex))))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zfloor (x * bpow (- fexp ex))
(IZR (Zfloor (x * bpow (- fexp ex))) * bpow (fexp ex) <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))
(bpow (ex - 1) <= IZR (Zceil (x * bpow (- fexp ex))) * bpow (fexp ex) <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zfloor (x * bpow (- fexp ex))

(bpow (ex - 1 + - fexp ex) <= IZR (Zfloor (x * bpow (- fexp ex))))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zfloor (x * bpow (- fexp ex))
(IZR (Zfloor (x * bpow (- fexp ex))) * bpow (fexp ex) <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))
(bpow (ex - 1) <= IZR (Zceil (x * bpow (- fexp ex))) * bpow (fexp ex) <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zfloor (x * bpow (- fexp ex))

IZR (beta ^ (ex - 1 - fexp ex)) = bpow (ex - 1 + - fexp ex)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zfloor (x * bpow (- fexp ex))
Hf:IZR (beta ^ (ex - 1 - fexp ex)) = bpow (ex - 1 + - fexp ex)
(bpow (ex - 1 + - fexp ex) <= IZR (Zfloor (x * bpow (- fexp ex))))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zfloor (x * bpow (- fexp ex))
(IZR (Zfloor (x * bpow (- fexp ex))) * bpow (fexp ex) <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))
(bpow (ex - 1) <= IZR (Zceil (x * bpow (- fexp ex))) * bpow (fexp ex) <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zfloor (x * bpow (- fexp ex))

(0 <= ex - 1 + - fexp ex)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zfloor (x * bpow (- fexp ex))
Hf:IZR (beta ^ (ex - 1 - fexp ex)) = bpow (ex - 1 + - fexp ex)
(bpow (ex - 1 + - fexp ex) <= IZR (Zfloor (x * bpow (- fexp ex))))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zfloor (x * bpow (- fexp ex))
(IZR (Zfloor (x * bpow (- fexp ex))) * bpow (fexp ex) <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))
(bpow (ex - 1) <= IZR (Zceil (x * bpow (- fexp ex))) * bpow (fexp ex) <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zfloor (x * bpow (- fexp ex))
Hf:IZR (beta ^ (ex - 1 - fexp ex)) = bpow (ex - 1 + - fexp ex)

(bpow (ex - 1 + - fexp ex) <= IZR (Zfloor (x * bpow (- fexp ex))))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zfloor (x * bpow (- fexp ex))
(IZR (Zfloor (x * bpow (- fexp ex))) * bpow (fexp ex) <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))
(bpow (ex - 1) <= IZR (Zceil (x * bpow (- fexp ex))) * bpow (fexp ex) <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zfloor (x * bpow (- fexp ex))
Hf:IZR (beta ^ (ex - 1 - fexp ex)) = bpow (ex - 1 + - fexp ex)

(IZR (beta ^ (ex - 1 - fexp ex)) <= IZR (Zfloor (x * bpow (- fexp ex))))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zfloor (x * bpow (- fexp ex))
(IZR (Zfloor (x * bpow (- fexp ex))) * bpow (fexp ex) <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))
(bpow (ex - 1) <= IZR (Zceil (x * bpow (- fexp ex))) * bpow (fexp ex) <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zfloor (x * bpow (- fexp ex))
Hf:IZR (beta ^ (ex - 1 - fexp ex)) = bpow (ex - 1 + - fexp ex)

(beta ^ (ex - 1 - fexp ex) <= Zfloor (x * bpow (- fexp ex)))%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zfloor (x * bpow (- fexp ex))
(IZR (Zfloor (x * bpow (- fexp ex))) * bpow (fexp ex) <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))
(bpow (ex - 1) <= IZR (Zceil (x * bpow (- fexp ex))) * bpow (fexp ex) <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zfloor (x * bpow (- fexp ex))
Hf:IZR (beta ^ (ex - 1 - fexp ex)) = bpow (ex - 1 + - fexp ex)

(IZR (beta ^ (ex - 1 - fexp ex)) <= x * bpow (- fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zfloor (x * bpow (- fexp ex))
(IZR (Zfloor (x * bpow (- fexp ex))) * bpow (fexp ex) <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))
(bpow (ex - 1) <= IZR (Zceil (x * bpow (- fexp ex))) * bpow (fexp ex) <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zfloor (x * bpow (- fexp ex))
Hf:IZR (beta ^ (ex - 1 - fexp ex)) = bpow (ex - 1 + - fexp ex)

(bpow (ex - 1 + - fexp ex) <= x * bpow (- fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zfloor (x * bpow (- fexp ex))
(IZR (Zfloor (x * bpow (- fexp ex))) * bpow (fexp ex) <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))
(bpow (ex - 1) <= IZR (Zceil (x * bpow (- fexp ex))) * bpow (fexp ex) <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zfloor (x * bpow (- fexp ex))
Hf:IZR (beta ^ (ex - 1 - fexp ex)) = bpow (ex - 1 + - fexp ex)

(bpow (ex - 1) * bpow (- fexp ex) <= x * bpow (- fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zfloor (x * bpow (- fexp ex))
(IZR (Zfloor (x * bpow (- fexp ex))) * bpow (fexp ex) <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))
(bpow (ex - 1) <= IZR (Zceil (x * bpow (- fexp ex))) * bpow (fexp ex) <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zfloor (x * bpow (- fexp ex))
Hf:IZR (beta ^ (ex - 1 - fexp ex)) = bpow (ex - 1 + - fexp ex)

(0 <= bpow (- fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zfloor (x * bpow (- fexp ex))
Hf:IZR (beta ^ (ex - 1 - fexp ex)) = bpow (ex - 1 + - fexp ex)
(bpow (ex - 1) <= x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zfloor (x * bpow (- fexp ex))
(IZR (Zfloor (x * bpow (- fexp ex))) * bpow (fexp ex) <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))
(bpow (ex - 1) <= IZR (Zceil (x * bpow (- fexp ex))) * bpow (fexp ex) <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zfloor (x * bpow (- fexp ex))
Hf:IZR (beta ^ (ex - 1 - fexp ex)) = bpow (ex - 1 + - fexp ex)

(bpow (ex - 1) <= x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zfloor (x * bpow (- fexp ex))
(IZR (Zfloor (x * bpow (- fexp ex))) * bpow (fexp ex) <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))
(bpow (ex - 1) <= IZR (Zceil (x * bpow (- fexp ex))) * bpow (fexp ex) <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zfloor (x * bpow (- fexp ex))

(IZR (Zfloor (x * bpow (- fexp ex))) * bpow (fexp ex) <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))
(bpow (ex - 1) <= IZR (Zceil (x * bpow (- fexp ex))) * bpow (fexp ex) <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zfloor (x * bpow (- fexp ex))

(IZR (Zfloor (x * bpow (- fexp ex))) * bpow (fexp ex) <= x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))
(bpow (ex - 1) <= IZR (Zceil (x * bpow (- fexp ex))) * bpow (fexp ex) <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zfloor (x * bpow (- fexp ex))

(0 < bpow (- fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zfloor (x * bpow (- fexp ex))
(IZR (Zfloor (x * bpow (- fexp ex))) * bpow (fexp ex) * bpow (- fexp ex) <= x * bpow (- fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))
(bpow (ex - 1) <= IZR (Zceil (x * bpow (- fexp ex))) * bpow (fexp ex) <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zfloor (x * bpow (- fexp ex))

(IZR (Zfloor (x * bpow (- fexp ex))) * bpow (fexp ex) * bpow (- fexp ex) <= x * bpow (- fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))
(bpow (ex - 1) <= IZR (Zceil (x * bpow (- fexp ex))) * bpow (fexp ex) <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zfloor (x * bpow (- fexp ex))

(IZR (Zfloor (x * bpow (- fexp ex))) <= x * bpow (- fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))
(bpow (ex - 1) <= IZR (Zceil (x * bpow (- fexp ex))) * bpow (fexp ex) <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))

(bpow (ex - 1) <= IZR (Zceil (x * bpow (- fexp ex))) * bpow (fexp ex) <= bpow ex)%R
(* UP *)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))

(bpow (ex - 1) <= IZR (Zceil (x * bpow (- fexp ex))) * bpow (fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))
(IZR (Zceil (x * bpow (- fexp ex))) * bpow (fexp ex) <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))

(x <= IZR (Zceil (x * bpow (- fexp ex))) * bpow (fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))
(IZR (Zceil (x * bpow (- fexp ex))) * bpow (fexp ex) <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))

(0 < bpow (- fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))
(x * bpow (- fexp ex) <= IZR (Zceil (x * bpow (- fexp ex))) * bpow (fexp ex) * bpow (- fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))
(IZR (Zceil (x * bpow (- fexp ex))) * bpow (fexp ex) <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))

(x * bpow (- fexp ex) <= IZR (Zceil (x * bpow (- fexp ex))) * bpow (fexp ex) * bpow (- fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))
(IZR (Zceil (x * bpow (- fexp ex))) * bpow (fexp ex) <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))

(x * bpow (- fexp ex) <= IZR (Zceil (x * bpow (- fexp ex))))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))
(IZR (Zceil (x * bpow (- fexp ex))) * bpow (fexp ex) <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))

(IZR (Zceil (x * bpow (- fexp ex))) * bpow (fexp ex) <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))

(IZR (Zceil (x * bpow (- fexp ex))) * bpow (fexp ex) <= bpow (ex - fexp ex + fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))

(IZR (Zceil (x * bpow (- fexp ex))) * bpow (fexp ex) <= bpow (ex - fexp ex) * bpow (fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))

(0 <= bpow (fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))
(IZR (Zceil (x * bpow (- fexp ex))) <= bpow (ex - fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))

(IZR (Zceil (x * bpow (- fexp ex))) <= bpow (ex - fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))

IZR (beta ^ (ex - fexp ex)) = bpow (ex - fexp ex)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))
Hf:IZR (beta ^ (ex - fexp ex)) = bpow (ex - fexp ex)
(IZR (Zceil (x * bpow (- fexp ex))) <= bpow (ex - fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))

(0 <= ex - fexp ex)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))
Hf:IZR (beta ^ (ex - fexp ex)) = bpow (ex - fexp ex)
(IZR (Zceil (x * bpow (- fexp ex))) <= bpow (ex - fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))
Hf:IZR (beta ^ (ex - fexp ex)) = bpow (ex - fexp ex)

(IZR (Zceil (x * bpow (- fexp ex))) <= bpow (ex - fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))
Hf:IZR (beta ^ (ex - fexp ex)) = bpow (ex - fexp ex)

(IZR (Zceil (x * bpow (- fexp ex))) <= IZR (beta ^ (ex - fexp ex)))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))
Hf:IZR (beta ^ (ex - fexp ex)) = bpow (ex - fexp ex)

(Zceil (x * bpow (- fexp ex)) <= beta ^ (ex - fexp ex))%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))
Hf:IZR (beta ^ (ex - fexp ex)) = bpow (ex - fexp ex)

(x * bpow (- fexp ex) <= IZR (beta ^ (ex - fexp ex)))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))
Hf:IZR (beta ^ (ex - fexp ex)) = bpow (ex - fexp ex)

(x * bpow (- fexp ex) <= bpow (ex - fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))
Hf:IZR (beta ^ (ex - fexp ex)) = bpow (ex - fexp ex)

(x * bpow (- fexp ex) <= bpow (ex + - fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))
Hf:IZR (beta ^ (ex - fexp ex)) = bpow (ex - fexp ex)

(x * bpow (- fexp ex) <= bpow ex * bpow (- fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))
Hf:IZR (beta ^ (ex - fexp ex)) = bpow (ex - fexp ex)

(0 <= bpow (- fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))
Hf:IZR (beta ^ (ex - fexp ex)) = bpow (ex - fexp ex)
(x <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))
Hf:IZR (beta ^ (ex - fexp ex)) = bpow (ex - fexp ex)

(x <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(fexp ex < ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))
Hf:IZR (beta ^ (ex - fexp ex)) = bpow (ex - fexp ex)

(x < bpow ex)%R
apply Hx. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd

forall (x : R) (ex : Z), (ex <= fexp ex)%Z -> (bpow (ex - 1) <= x < bpow ex)%R -> round x = 0%R \/ round x = bpow (fexp ex)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd

forall (x : R) (ex : Z), (ex <= fexp ex)%Z -> (bpow (ex - 1) <= x < bpow ex)%R -> round x = 0%R \/ round x = bpow (fexp ex)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(ex <= fexp ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R

round x = 0%R \/ round x = bpow (fexp ex)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(ex <= fexp ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R

F2R {| Fnum := rnd (x * bpow (- cexp x)); Fexp := cexp x |} = 0%R \/ F2R {| Fnum := rnd (x * bpow (- cexp x)); Fexp := cexp x |} = bpow (fexp ex)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(ex <= fexp ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R

F2R {| Fnum := rnd (x * bpow (- fexp ex)); Fexp := fexp ex |} = 0%R \/ F2R {| Fnum := rnd (x * bpow (- fexp ex)); Fexp := fexp ex |} = bpow (fexp ex)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(ex <= fexp ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R

(IZR (Fnum {| Fnum := rnd (x * bpow (- fexp ex)); Fexp := fexp ex |}) * bpow (Fexp {| Fnum := rnd (x * bpow (- fexp ex)); Fexp := fexp ex |}))%R = 0%R \/ (IZR (Fnum {| Fnum := rnd (x * bpow (- fexp ex)); Fexp := fexp ex |}) * bpow (Fexp {| Fnum := rnd (x * bpow (- fexp ex)); Fexp := fexp ex |}))%R = bpow (fexp ex)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(ex <= fexp ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R

(IZR (rnd (x * bpow (- fexp ex))) * bpow (fexp ex))%R = 0%R \/ (IZR (rnd (x * bpow (- fexp ex))) * bpow (fexp ex))%R = bpow (fexp ex)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(ex <= fexp ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zfloor (x * bpow (- fexp ex))

(IZR (Zfloor (x * bpow (- fexp ex))) * bpow (fexp ex))%R = 0%R \/ (IZR (Zfloor (x * bpow (- fexp ex))) * bpow (fexp ex))%R = bpow (fexp ex)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(ex <= fexp ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))
(IZR (Zceil (x * bpow (- fexp ex))) * bpow (fexp ex))%R = 0%R \/ (IZR (Zceil (x * bpow (- fexp ex))) * bpow (fexp ex))%R = bpow (fexp ex)
(* DN *)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(ex <= fexp ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zfloor (x * bpow (- fexp ex))

(IZR (Zfloor (x * bpow (- fexp ex))) * bpow (fexp ex))%R = 0%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(ex <= fexp ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))
(IZR (Zceil (x * bpow (- fexp ex))) * bpow (fexp ex))%R = 0%R \/ (IZR (Zceil (x * bpow (- fexp ex))) * bpow (fexp ex))%R = bpow (fexp ex)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(ex <= fexp ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zfloor (x * bpow (- fexp ex))

IZR (Zfloor (x * bpow (- fexp ex))) = 0%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(ex <= fexp ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))
(IZR (Zceil (x * bpow (- fexp ex))) * bpow (fexp ex))%R = 0%R \/ (IZR (Zceil (x * bpow (- fexp ex))) * bpow (fexp ex))%R = bpow (fexp ex)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(ex <= fexp ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zfloor (x * bpow (- fexp ex))

Zfloor (x * bpow (- fexp ex)) = 0%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(ex <= fexp ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))
(IZR (Zceil (x * bpow (- fexp ex))) * bpow (fexp ex))%R = 0%R \/ (IZR (Zceil (x * bpow (- fexp ex))) * bpow (fexp ex))%R = bpow (fexp ex)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(ex <= fexp ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zfloor (x * bpow (- fexp ex))

(0 <= x * bpow (- fexp ex) < IZR (0 + 1))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(ex <= fexp ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))
(IZR (Zceil (x * bpow (- fexp ex))) * bpow (fexp ex))%R = 0%R \/ (IZR (Zceil (x * bpow (- fexp ex))) * bpow (fexp ex))%R = bpow (fexp ex)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(ex <= fexp ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zfloor (x * bpow (- fexp ex))

(0 < x * bpow (- fexp ex) < IZR (0 + 1))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(ex <= fexp ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))
(IZR (Zceil (x * bpow (- fexp ex))) * bpow (fexp ex))%R = 0%R \/ (IZR (Zceil (x * bpow (- fexp ex))) * bpow (fexp ex))%R = bpow (fexp ex)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(ex <= fexp ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))

(IZR (Zceil (x * bpow (- fexp ex))) * bpow (fexp ex))%R = 0%R \/ (IZR (Zceil (x * bpow (- fexp ex))) * bpow (fexp ex))%R = bpow (fexp ex)
(* UP *)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(ex <= fexp ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))

(IZR (Zceil (x * bpow (- fexp ex))) * bpow (fexp ex))%R = bpow (fexp ex)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(ex <= fexp ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))

(IZR (Zceil (x * bpow (- fexp ex))) * bpow (fexp ex))%R = (1 * bpow (fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(ex <= fexp ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))

IZR (Zceil (x * bpow (- fexp ex))) = 1%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(ex <= fexp ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))

Zceil (x * bpow (- fexp ex)) = 1%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(ex <= fexp ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))

(IZR (1 - 1) < x * bpow (- fexp ex) <= 1)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
He:(ex <= fexp ex)%Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
Hr:rnd (x * bpow (- fexp ex)) = Zceil (x * bpow (- fexp ex))

(IZR (1 - 1) < x * bpow (- fexp ex) < 1)%R
now apply mantissa_small_pos. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd

forall x y : R, (0 < x)%R -> (x <= y)%R -> (round x <= round y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd

forall x y : R, (0 < x)%R -> (x <= y)%R -> (round x <= round y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R

(round x <= round y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R

(round x <= round y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R

(round x <= round y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R

(round x <= round y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= Rabs y < bpow ey)%R

(round x <= round y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= Rabs y < bpow ey)%R

(round x <= round y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= Rabs y < bpow ey)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= Rabs y < bpow ey)%R

(round x <= round y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R

(round x <= round y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= Rabs y < bpow ey)%R
(0 <= y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R

(round x <= round y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R

(ex <= ey)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R
He:(ex <= ey)%Z
(round x <= round y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R

(bpow (ex - 1) < bpow ey)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R
He:(ex <= ey)%Z
(round x <= round y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R

(x < bpow ey)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R
He:(ex <= ey)%Z
(round x <= round y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R
He:(ex <= ey)%Z

(round x <= round y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R
He:(ex <= ey)%Z

fexp ex = fexp ey -> (round x <= round y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R
He:(ex <= ey)%Z
Heq:fexp ex = fexp ey -> (round x <= round y)%R
(round x <= round y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R
He:(ex <= ey)%Z
H:fexp ex = fexp ey

(round x <= round y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R
He:(ex <= ey)%Z
Heq:fexp ex = fexp ey -> (round x <= round y)%R
(round x <= round y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R
He:(ex <= ey)%Z
H:fexp ex = fexp ey

(F2R {| Fnum := rnd (x * bpow (- fexp (mag beta x))); Fexp := fexp (mag beta x) |} <= F2R {| Fnum := rnd (y * bpow (- fexp (mag beta y))); Fexp := fexp (mag beta y) |})%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R
He:(ex <= ey)%Z
Heq:fexp ex = fexp ey -> (round x <= round y)%R
(round x <= round y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R
He:(ex <= ey)%Z
H:fexp ex = fexp ey

(F2R {| Fnum := rnd (x * bpow (- fexp ex)); Fexp := fexp ex |} <= F2R {| Fnum := rnd (y * bpow (- fexp (mag beta y))); Fexp := fexp (mag beta y) |})%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R
He:(ex <= ey)%Z
Heq:fexp ex = fexp ey -> (round x <= round y)%R
(round x <= round y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R
He:(ex <= ey)%Z
H:fexp ex = fexp ey

(F2R {| Fnum := rnd (x * bpow (- fexp ex)); Fexp := fexp ex |} <= F2R {| Fnum := rnd (y * bpow (- fexp ey)); Fexp := fexp ey |})%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R
He:(ex <= ey)%Z
Heq:fexp ex = fexp ey -> (round x <= round y)%R
(round x <= round y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R
He:(ex <= ey)%Z
H:fexp ex = fexp ey

(F2R {| Fnum := rnd (x * bpow (- fexp ey)); Fexp := fexp ey |} <= F2R {| Fnum := rnd (y * bpow (- fexp ey)); Fexp := fexp ey |})%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R
He:(ex <= ey)%Z
Heq:fexp ex = fexp ey -> (round x <= round y)%R
(round x <= round y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R
He:(ex <= ey)%Z
H:fexp ex = fexp ey

(rnd (x * bpow (- fexp ey)) <= rnd (y * bpow (- fexp ey)))%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R
He:(ex <= ey)%Z
Heq:fexp ex = fexp ey -> (round x <= round y)%R
(round x <= round y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R
He:(ex <= ey)%Z
H:fexp ex = fexp ey

(x * bpow (- fexp ey) <= y * bpow (- fexp ey))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R
He:(ex <= ey)%Z
Heq:fexp ex = fexp ey -> (round x <= round y)%R
(round x <= round y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R
He:(ex <= ey)%Z
H:fexp ex = fexp ey

(0 <= bpow (- fexp ey))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R
He:(ex <= ey)%Z
Heq:fexp ex = fexp ey -> (round x <= round y)%R
(round x <= round y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R
He:(ex <= ey)%Z
Heq:fexp ex = fexp ey -> (round x <= round y)%R

(round x <= round y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R
He:(ex <= ey)%Z
Heq:fexp ex = fexp ey -> (round x <= round y)%R
Hy1:(ey <= fexp ey)%Z

(round x <= round y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R
He:(ex <= ey)%Z
Heq:fexp ex = fexp ey -> (round x <= round y)%R
Hy1:(fexp ey < ey)%Z
(round x <= round y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R
He:(ex <= ey)%Z
Heq:fexp ex = fexp ey -> (round x <= round y)%R
Hy1:(ey <= fexp ey)%Z

fexp ex = fexp ey
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R
He:(ex <= ey)%Z
Heq:fexp ex = fexp ey -> (round x <= round y)%R
Hy1:(fexp ey < ey)%Z
(round x <= round y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R
He:(ex <= ey)%Z
Heq:fexp ex = fexp ey -> (round x <= round y)%R
Hy1:(ey <= fexp ey)%Z

(ex <= fexp ey)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R
He:(ex <= ey)%Z
Heq:fexp ex = fexp ey -> (round x <= round y)%R
Hy1:(fexp ey < ey)%Z
(round x <= round y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R
He:(ex <= ey)%Z
Heq:fexp ex = fexp ey -> (round x <= round y)%R
Hy1:(fexp ey < ey)%Z

(round x <= round y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R
He:(ex <= ey)%Z
Heq:fexp ex = fexp ey -> (round x <= round y)%R
Hy1:(fexp ey < ey)%Z
He':(ex < ey)%Z

(round x <= round y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R
He:(ex <= ey)%Z
Heq:fexp ex = fexp ey -> (round x <= round y)%R
Hy1:(fexp ey < ey)%Z
He':ex = ey
(round x <= round y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R
He:(ex <= ey)%Z
Heq:fexp ex = fexp ey -> (round x <= round y)%R
Hy1:(fexp ey < ey)%Z
He':(ex < ey)%Z

(round x <= round y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R
He:(ex <= ey)%Z
Heq:fexp ex = fexp ey -> (round x <= round y)%R
Hy1:(fexp ey < ey)%Z
He':(ex < ey)%Z

(round x <= bpow (ey - 1))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R
He:(ex <= ey)%Z
Heq:fexp ex = fexp ey -> (round x <= round y)%R
Hy1:(fexp ey < ey)%Z
He':(ex < ey)%Z
(bpow (ey - 1) <= round y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R
He:(ex <= ey)%Z
Heq:fexp ex = fexp ey -> (round x <= round y)%R
Hy1:(fexp ey < ey)%Z
He':(ex < ey)%Z

(round x <= bpow (ey - 1))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R
He:(ex <= ey)%Z
Heq:fexp ex = fexp ey -> (round x <= round y)%R
Hy1:(fexp ey < ey)%Z
He':(ex < ey)%Z
Hx1:(ex <= fexp ex)%Z

(round x <= bpow (ey - 1))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R
He:(ex <= ey)%Z
Heq:fexp ex = fexp ey -> (round x <= round y)%R
Hy1:(fexp ey < ey)%Z
He':(ex < ey)%Z
Hx1:(fexp ex < ex)%Z
(round x <= bpow (ey - 1))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R
He:(ex <= ey)%Z
Heq:fexp ex = fexp ey -> (round x <= round y)%R
Hy1:(fexp ey < ey)%Z
He':(ex < ey)%Z
Hx1:(ex <= fexp ex)%Z

(0 <= bpow (ey - 1))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R
He:(ex <= ey)%Z
Heq:fexp ex = fexp ey -> (round x <= round y)%R
Hy1:(fexp ey < ey)%Z
He':(ex < ey)%Z
Hx1:(ex <= fexp ex)%Z
(bpow (fexp ex) <= bpow (ey - 1))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R
He:(ex <= ey)%Z
Heq:fexp ex = fexp ey -> (round x <= round y)%R
Hy1:(fexp ey < ey)%Z
He':(ex < ey)%Z
Hx1:(fexp ex < ex)%Z
(round x <= bpow (ey - 1))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R
He:(ex <= ey)%Z
Heq:fexp ex = fexp ey -> (round x <= round y)%R
Hy1:(fexp ey < ey)%Z
He':(ex < ey)%Z
Hx1:(ex <= fexp ex)%Z

(bpow (fexp ex) <= bpow (ey - 1))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R
He:(ex <= ey)%Z
Heq:fexp ex = fexp ey -> (round x <= round y)%R
Hy1:(fexp ey < ey)%Z
He':(ex < ey)%Z
Hx1:(fexp ex < ex)%Z
(round x <= bpow (ey - 1))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R
He:(ex <= ey)%Z
Heq:fexp ex = fexp ey -> (round x <= round y)%R
Hy1:(fexp ey < ey)%Z
He':(ex < ey)%Z
Hx1:(ex <= fexp ex)%Z

(fexp ex <= ey - 1)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R
He:(ex <= ey)%Z
Heq:fexp ex = fexp ey -> (round x <= round y)%R
Hy1:(fexp ey < ey)%Z
He':(ex < ey)%Z
Hx1:(fexp ex < ex)%Z
(round x <= bpow (ey - 1))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R
He:(ex <= ey)%Z
Heq:fexp ex = fexp ey -> (round x <= round y)%R
Hy1:(fexp ey < ey)%Z
He':(ex < ey)%Z
Hx1:forall l : Z, (l <= fexp ex)%Z -> fexp l = fexp ex

(fexp ex <= ey - 1)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R
He:(ex <= ey)%Z
Heq:fexp ex = fexp ey -> (round x <= round y)%R
Hy1:(fexp ey < ey)%Z
He':(ex < ey)%Z
Hx1:(fexp ex < ex)%Z
(round x <= bpow (ey - 1))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R
He:(ex <= ey)%Z
Heq:fexp ex = fexp ey -> (round x <= round y)%R
Hy1:(fexp ey < ey)%Z
He':(ex < ey)%Z
Hx1:(ey <= fexp ex)%Z -> fexp ey = fexp ex

(fexp ex <= ey - 1)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R
He:(ex <= ey)%Z
Heq:fexp ex = fexp ey -> (round x <= round y)%R
Hy1:(fexp ey < ey)%Z
He':(ex < ey)%Z
Hx1:(fexp ex < ex)%Z
(round x <= bpow (ey - 1))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R
He:(ex <= ey)%Z
Heq:fexp ex = fexp ey -> (round x <= round y)%R
Hy1:(fexp ey < ey)%Z
He':(ex < ey)%Z
Hx1:(fexp ex < ex)%Z

(round x <= bpow (ey - 1))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R
He:(ex <= ey)%Z
Heq:fexp ex = fexp ey -> (round x <= round y)%R
Hy1:(fexp ey < ey)%Z
He':(ex < ey)%Z
Hx1:(fexp ex < ex)%Z

(round x <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R
He:(ex <= ey)%Z
Heq:fexp ex = fexp ey -> (round x <= round y)%R
Hy1:(fexp ey < ey)%Z
He':(ex < ey)%Z
Hx1:(fexp ex < ex)%Z
(bpow ex <= bpow (ey - 1))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R
He:(ex <= ey)%Z
Heq:fexp ex = fexp ey -> (round x <= round y)%R
Hy1:(fexp ey < ey)%Z
He':(ex < ey)%Z
Hx1:(fexp ex < ex)%Z

(bpow ex <= bpow (ey - 1))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
ey:Z
Hey:(bpow (ey - 1) <= y < bpow ey)%R
He:(ex <= ey)%Z
Heq:fexp ex = fexp ey -> (round x <= round y)%R
Hy1:(fexp ey < ey)%Z
He':(ex < ey)%Z
Hx1:(fexp ex < ex)%Z

(ex <= ey - 1)%Z
now apply Z.lt_le_pred. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd

forall x : R, generic_format x -> round x = x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd

forall x : R, generic_format x -> round x = x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx:generic_format x

round x = x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx:generic_format x

F2R {| Fnum := rnd (scaled_mantissa x); Fexp := cexp x |} = x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx:generic_format x

F2R {| Fnum := rnd (IZR (Ztrunc (scaled_mantissa x))); Fexp := cexp x |} = x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx:generic_format x

F2R {| Fnum := Ztrunc (scaled_mantissa x); Fexp := cexp x |} = x
now apply sym_eq. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd

round 0 = 0%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd

round 0 = 0%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd

F2R {| Fnum := rnd (0 * bpow (- cexp 0)); Fexp := cexp 0 |} = 0%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd

F2R {| Fnum := rnd 0; Fexp := cexp 0 |} = 0%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd

F2R {| Fnum := 0; Fexp := cexp 0 |} = 0%R
apply F2R_0. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd

forall (x : R) (ex : Z), (bpow (ex - 1) <= x < bpow ex)%R -> round x = 0%R -> (ex <= fexp ex)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd

forall (x : R) (ex : Z), (bpow (ex - 1) <= x < bpow ex)%R -> round x = 0%R -> (ex <= fexp ex)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
H:(bpow (ex - 1) <= x < bpow ex)%R
H1:round x = 0%R

(ex <= fexp ex)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
H:(bpow (ex - 1) <= x < bpow ex)%R
H1:round x = 0%R
V:(fexp ex < ex)%Z

(ex <= fexp ex)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
H:(bpow (ex - 1) <= x < bpow ex)%R
V:(fexp ex < ex)%Z

round x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
H:(bpow (ex - 1) <= x < bpow ex)%R
V:(fexp ex < ex)%Z

(round x > 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
H:(bpow (ex - 1) <= x < bpow ex)%R
V:(fexp ex < ex)%Z

(0 < bpow (ex - 1))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
H:(bpow (ex - 1) <= x < bpow ex)%R
V:(fexp ex < ex)%Z
(bpow (ex - 1) <= round x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
ex:Z
H:(bpow (ex - 1) <= x < bpow ex)%R
V:(fexp ex < ex)%Z

(bpow (ex - 1) <= round x)%R
apply (round_bounded_large_pos); assumption. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd

forall x : R, (0 < x)%R -> generic_format (round x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd

forall x : R, (0 < x)%R -> generic_format (round x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx0:(0 < x)%R

generic_format (round x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx0:(0 < x)%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R

generic_format (round x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx0:(0 < x)%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R

generic_format (round x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx0:(0 < x)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R

generic_format (round x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx0:(0 < x)%R
ex:Z
Hex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx0:(0 < x)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R

generic_format (round x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx0:(0 < x)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
He:(ex <= fexp ex)%Z

generic_format (round x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx0:(0 < x)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
He:(fexp ex < ex)%Z
generic_format (round x)
(* small *)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx0:(0 < x)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
He:(ex <= fexp ex)%Z
Hr:round x = 0%R

generic_format 0
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx0:(0 < x)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
He:(ex <= fexp ex)%Z
Hr:round x = bpow (fexp ex)
generic_format (bpow (fexp ex))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx0:(0 < x)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
He:(fexp ex < ex)%Z
generic_format (round x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx0:(0 < x)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
He:(ex <= fexp ex)%Z
Hr:round x = bpow (fexp ex)

generic_format (bpow (fexp ex))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx0:(0 < x)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
He:(fexp ex < ex)%Z
generic_format (round x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx0:(0 < x)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
He:(ex <= fexp ex)%Z
Hr:round x = bpow (fexp ex)

(fexp (fexp ex + 1) <= fexp ex)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx0:(0 < x)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
He:(fexp ex < ex)%Z
generic_format (round x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx0:(0 < x)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
He:(fexp ex < ex)%Z

generic_format (round x)
(* large *)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx0:(0 < x)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
He:(fexp ex < ex)%Z

(bpow (ex - 1) <= round x <= bpow ex)%R -> generic_format (round x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx0:(0 < x)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
He:(fexp ex < ex)%Z
Hr1:(bpow (ex - 1) <= round x)%R
Hr2:(round x <= bpow ex)%R

generic_format (round x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx0:(0 < x)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
He:(fexp ex < ex)%Z
Hr1:(bpow (ex - 1) <= round x)%R
Hr2:(round x <= bpow ex)%R
Hr:(bpow ex <= round x)%R

generic_format (round x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx0:(0 < x)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
He:(fexp ex < ex)%Z
Hr1:(bpow (ex - 1) <= round x)%R
Hr2:(round x <= bpow ex)%R
Hr:(round x < bpow ex)%R
generic_format (round x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx0:(0 < x)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
He:(fexp ex < ex)%Z
Hr1:(bpow (ex - 1) <= round x)%R
Hr2:(round x <= bpow ex)%R
Hr:(bpow ex <= round x)%R

generic_format (bpow ex)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx0:(0 < x)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
He:(fexp ex < ex)%Z
Hr1:(bpow (ex - 1) <= round x)%R
Hr2:(round x <= bpow ex)%R
Hr:(round x < bpow ex)%R
generic_format (round x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx0:(0 < x)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
He:(fexp ex < ex)%Z
Hr1:(bpow (ex - 1) <= round x)%R
Hr2:(round x <= bpow ex)%R
Hr:(bpow ex <= round x)%R

(fexp (ex + 1) <= ex)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx0:(0 < x)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
He:(fexp ex < ex)%Z
Hr1:(bpow (ex - 1) <= round x)%R
Hr2:(round x <= bpow ex)%R
Hr:(round x < bpow ex)%R
generic_format (round x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx0:(0 < x)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
He:(fexp ex < ex)%Z
Hr1:(bpow (ex - 1) <= round x)%R
Hr2:(round x <= bpow ex)%R
Hr:(round x < bpow ex)%R

generic_format (round x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx0:(0 < x)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
He:(fexp ex < ex)%Z
Hr1:(bpow (ex - 1) <= round x)%R
Hr2:(round x <= bpow ex)%R
Hr:(round x < bpow ex)%R

rnd (scaled_mantissa x) <> 0%Z -> (cexp (F2R {| Fnum := rnd (scaled_mantissa x); Fexp := cexp x |}) <= cexp x)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx0:(0 < x)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
He:(fexp ex < ex)%Z
Hr1:(bpow (ex - 1) <= round x)%R
Hr2:(round x <= bpow ex)%R
Hr:(round x < bpow ex)%R

(cexp (F2R {| Fnum := rnd (scaled_mantissa x); Fexp := cexp x |}) <= cexp x)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx0:(0 < x)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
He:(fexp ex < ex)%Z
Hr1:(bpow (ex - 1) <= round x)%R
Hr2:(round x <= bpow ex)%R
Hr:(round x < bpow ex)%R

(fexp ex <= cexp x)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd
x:R
Hx0:(0 < x)%R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
He:(fexp ex < ex)%Z
Hr1:(bpow (ex - 1) <= round x)%R
Hr2:(round x <= bpow ex)%R
Hr:(round x < bpow ex)%R

(fexp ex <= fexp ex)%Z
now apply Zeq_le. Qed. End Fcore_generic_round_pos.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall rnd1 rnd2 : R -> Z, (forall x : R, rnd1 x = rnd2 x) -> forall x : R, round rnd1 x = round rnd2 x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall rnd1 rnd2 : R -> Z, (forall x : R, rnd1 x = rnd2 x) -> forall x : R, round rnd1 x = round rnd2 x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd1, rnd2:R -> Z
Hext:forall x0 : R, rnd1 x0 = rnd2 x0
x:R

round rnd1 x = round rnd2 x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd1, rnd2:R -> Z
Hext:forall x0 : R, rnd1 x0 = rnd2 x0
x:R

F2R {| Fnum := rnd1 (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := rnd2 (scaled_mantissa x); Fexp := cexp x |}
now rewrite Hext. Qed. Section Zround_opp. Variable rnd : R -> Z. Context { valid_rnd : Valid_rnd rnd }. Definition Zrnd_opp x := Z.opp (rnd (-x)).
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd

Valid_rnd Zrnd_opp
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd

Valid_rnd Zrnd_opp
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd

forall x y : R, (x <= y)%R -> (Zrnd_opp x <= Zrnd_opp y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
forall n : Z, Zrnd_opp (IZR n) = n
(* *)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R

(Zrnd_opp x <= Zrnd_opp y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
forall n : Z, Zrnd_opp (IZR n) = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R

(- rnd (- x) <= - rnd (- y))%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
forall n : Z, Zrnd_opp (IZR n) = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R

(- - rnd (- y) <= - - rnd (- x))%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
forall n : Z, Zrnd_opp (IZR n) = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R

(rnd (- y) <= rnd (- x))%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
forall n : Z, Zrnd_opp (IZR n) = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R

(- y <= - x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
forall n : Z, Zrnd_opp (IZR n) = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd

forall n : Z, Zrnd_opp (IZR n) = n
(* *)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
n:Z

Zrnd_opp (IZR n) = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
n:Z

(- rnd (- IZR n))%Z = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
n:Z

(- - n)%Z = n
apply Z.opp_involutive. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd

forall x : R, round rnd (- x) = (- round Zrnd_opp x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd

forall x : R, round rnd (- x) = (- round Zrnd_opp x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R

round rnd (- x) = (- round Zrnd_opp x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R

F2R {| Fnum := rnd (scaled_mantissa (- x)); Fexp := cexp (- x) |} = (- F2R {| Fnum := Zrnd_opp (scaled_mantissa x); Fexp := cexp x |})%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R

F2R {| Fnum := rnd (- scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := - Zrnd_opp (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R

rnd (- scaled_mantissa x) = (- Zrnd_opp (scaled_mantissa x))%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R

(- Zrnd_opp (scaled_mantissa x))%Z = rnd (- scaled_mantissa x)
exact (Z.opp_involutive _). Qed. End Zround_opp.
IEEE-754 roundings: up, down and to zero
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

Valid_rnd Zfloor
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

Valid_rnd Zfloor
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x y : R, (x <= y)%R -> (Zfloor x <= Zfloor y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
forall n : Z, Zfloor (IZR n) = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall n : Z, Zfloor (IZR n) = n
apply Zfloor_IZR. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

Valid_rnd Zceil
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

Valid_rnd Zceil
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x y : R, (x <= y)%R -> (Zceil x <= Zceil y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
forall n : Z, Zceil (IZR n) = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall n : Z, Zceil (IZR n) = n
apply Zceil_IZR. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

Valid_rnd Ztrunc
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

Valid_rnd Ztrunc
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x y : R, (x <= y)%R -> (Ztrunc x <= Ztrunc y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
forall n : Z, Ztrunc (IZR n) = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall n : Z, Ztrunc (IZR n) = n
apply Ztrunc_IZR. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

Valid_rnd Zaway
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

Valid_rnd Zaway
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x y : R, (x <= y)%R -> (Zaway x <= Zaway y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
forall n : Z, Zaway (IZR n) = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall n : Z, Zaway (IZR n) = n
apply Zaway_IZR. Qed. Section monotone. Variable rnd : R -> Z. Context { valid_rnd : Valid_rnd rnd }.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd

forall x : R, round rnd x = round Zfloor x \/ round rnd x = round Zceil x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd

forall x : R, round rnd x = round Zfloor x \/ round rnd x = round Zceil x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R

round rnd x = round Zfloor x \/ round rnd x = round Zceil x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R

F2R {| Fnum := rnd (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zfloor (scaled_mantissa x); Fexp := cexp x |} \/ F2R {| Fnum := rnd (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zceil (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:rnd (scaled_mantissa x) = Zfloor (scaled_mantissa x)

F2R {| Fnum := rnd (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zfloor (scaled_mantissa x); Fexp := cexp x |} \/ F2R {| Fnum := rnd (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zceil (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:rnd (scaled_mantissa x) = Zceil (scaled_mantissa x)
F2R {| Fnum := rnd (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zfloor (scaled_mantissa x); Fexp := cexp x |} \/ F2R {| Fnum := rnd (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zceil (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:rnd (scaled_mantissa x) = Zfloor (scaled_mantissa x)

F2R {| Fnum := rnd (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zfloor (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:rnd (scaled_mantissa x) = Zceil (scaled_mantissa x)
F2R {| Fnum := rnd (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zfloor (scaled_mantissa x); Fexp := cexp x |} \/ F2R {| Fnum := rnd (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zceil (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:rnd (scaled_mantissa x) = Zceil (scaled_mantissa x)

F2R {| Fnum := rnd (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zfloor (scaled_mantissa x); Fexp := cexp x |} \/ F2R {| Fnum := rnd (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zceil (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:rnd (scaled_mantissa x) = Zceil (scaled_mantissa x)

F2R {| Fnum := rnd (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zceil (scaled_mantissa x); Fexp := cexp x |}
now rewrite Hx. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd

forall x : R, round rnd x = round Ztrunc x \/ round rnd x = round Zaway x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd

forall x : R, round rnd x = round Ztrunc x \/ round rnd x = round Zaway x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R

round rnd x = round Ztrunc x \/ round rnd x = round Zaway x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R

F2R {| Fnum := rnd (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Ztrunc (scaled_mantissa x); Fexp := cexp x |} \/ F2R {| Fnum := rnd (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zaway (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:rnd (scaled_mantissa x) = Ztrunc (scaled_mantissa x)

F2R {| Fnum := rnd (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Ztrunc (scaled_mantissa x); Fexp := cexp x |} \/ F2R {| Fnum := rnd (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zaway (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:rnd (scaled_mantissa x) = Zaway (scaled_mantissa x)
F2R {| Fnum := rnd (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Ztrunc (scaled_mantissa x); Fexp := cexp x |} \/ F2R {| Fnum := rnd (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zaway (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:rnd (scaled_mantissa x) = Ztrunc (scaled_mantissa x)

F2R {| Fnum := rnd (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Ztrunc (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:rnd (scaled_mantissa x) = Zaway (scaled_mantissa x)
F2R {| Fnum := rnd (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Ztrunc (scaled_mantissa x); Fexp := cexp x |} \/ F2R {| Fnum := rnd (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zaway (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:rnd (scaled_mantissa x) = Zaway (scaled_mantissa x)

F2R {| Fnum := rnd (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Ztrunc (scaled_mantissa x); Fexp := cexp x |} \/ F2R {| Fnum := rnd (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zaway (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:rnd (scaled_mantissa x) = Zaway (scaled_mantissa x)

F2R {| Fnum := rnd (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zaway (scaled_mantissa x); Fexp := cexp x |}
now rewrite Hx. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd

forall x y : R, (x <= y)%R -> (round rnd x <= round rnd y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd

forall x y : R, (x <= y)%R -> (round rnd x <= round rnd y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R

(round rnd x <= round rnd y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R

(round rnd x <= round rnd y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:x = 0%R
(round rnd x <= round rnd y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:(x > 0)%R
(round rnd x <= round rnd y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R

(round rnd x <= round rnd y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:x = 0%R
(round rnd x <= round rnd y)%R
(* x < 0 *)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R

(F2R {| Fnum := rnd (scaled_mantissa x); Fexp := cexp x |} <= F2R {| Fnum := rnd (scaled_mantissa y); Fexp := cexp y |})%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:x = 0%R
(round rnd x <= round rnd y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(y < 0)%R

(F2R {| Fnum := rnd (scaled_mantissa x); Fexp := cexp x |} <= F2R {| Fnum := rnd (scaled_mantissa y); Fexp := cexp y |})%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R
(F2R {| Fnum := rnd (scaled_mantissa x); Fexp := cexp x |} <= F2R {| Fnum := rnd (scaled_mantissa y); Fexp := cexp y |})%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:x = 0%R
(round rnd x <= round rnd y)%R
(* . y < 0 *)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(y < 0)%R

(F2R {| Fnum := rnd (scaled_mantissa (- - x)); Fexp := cexp (- - x) |} <= F2R {| Fnum := rnd (scaled_mantissa (- - y)); Fexp := cexp (- - y) |})%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R
(F2R {| Fnum := rnd (scaled_mantissa x); Fexp := cexp x |} <= F2R {| Fnum := rnd (scaled_mantissa y); Fexp := cexp y |})%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:x = 0%R
(round rnd x <= round rnd y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(y < 0)%R

(F2R {| Fnum := rnd (- scaled_mantissa (- x)); Fexp := cexp (- - x) |} <= F2R {| Fnum := rnd (- scaled_mantissa (- y)); Fexp := cexp (- - y) |})%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R
(F2R {| Fnum := rnd (scaled_mantissa x); Fexp := cexp x |} <= F2R {| Fnum := rnd (scaled_mantissa y); Fexp := cexp y |})%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:x = 0%R
(round rnd x <= round rnd y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(y < 0)%R

(F2R {| Fnum := rnd (- scaled_mantissa (- x)); Fexp := cexp (- x) |} <= F2R {| Fnum := rnd (- scaled_mantissa (- y)); Fexp := cexp (- y) |})%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R
(F2R {| Fnum := rnd (scaled_mantissa x); Fexp := cexp x |} <= F2R {| Fnum := rnd (scaled_mantissa y); Fexp := cexp y |})%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:x = 0%R
(round rnd x <= round rnd y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(y < 0)%R

(- F2R {| Fnum := rnd (- scaled_mantissa (- y)); Fexp := cexp (- y) |} <= - F2R {| Fnum := rnd (- scaled_mantissa (- x)); Fexp := cexp (- x) |})%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R
(F2R {| Fnum := rnd (scaled_mantissa x); Fexp := cexp x |} <= F2R {| Fnum := rnd (scaled_mantissa y); Fexp := cexp y |})%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:x = 0%R
(round rnd x <= round rnd y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(y < 0)%R

(F2R {| Fnum := - rnd (- scaled_mantissa (- y)); Fexp := cexp (- y) |} <= F2R {| Fnum := - rnd (- scaled_mantissa (- x)); Fexp := cexp (- x) |})%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R
(F2R {| Fnum := rnd (scaled_mantissa x); Fexp := cexp x |} <= F2R {| Fnum := rnd (scaled_mantissa y); Fexp := cexp y |})%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:x = 0%R
(round rnd x <= round rnd y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(y < 0)%R

(0 < - y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(y < 0)%R
(- y <= - x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R
(F2R {| Fnum := rnd (scaled_mantissa x); Fexp := cexp x |} <= F2R {| Fnum := rnd (scaled_mantissa y); Fexp := cexp y |})%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:x = 0%R
(round rnd x <= round rnd y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(y < 0)%R

(- 0 < - y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(y < 0)%R
(- y <= - x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R
(F2R {| Fnum := rnd (scaled_mantissa x); Fexp := cexp x |} <= F2R {| Fnum := rnd (scaled_mantissa y); Fexp := cexp y |})%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:x = 0%R
(round rnd x <= round rnd y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(y < 0)%R

(- y <= - x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R
(F2R {| Fnum := rnd (scaled_mantissa x); Fexp := cexp x |} <= F2R {| Fnum := rnd (scaled_mantissa y); Fexp := cexp y |})%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:x = 0%R
(round rnd x <= round rnd y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R

(F2R {| Fnum := rnd (scaled_mantissa x); Fexp := cexp x |} <= F2R {| Fnum := rnd (scaled_mantissa y); Fexp := cexp y |})%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:x = 0%R
(round rnd x <= round rnd y)%R
(* . 0 <= y *)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R

(F2R {| Fnum := rnd (scaled_mantissa x); Fexp := cexp x |} <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R
(0 <= F2R {| Fnum := rnd (scaled_mantissa y); Fexp := cexp y |})%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:x = 0%R
(round rnd x <= round rnd y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R

(Fnum {| Fnum := rnd (scaled_mantissa x); Fexp := cexp x |} <= 0)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R
(0 <= F2R {| Fnum := rnd (scaled_mantissa y); Fexp := cexp y |})%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:x = 0%R
(round rnd x <= round rnd y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R

(rnd (scaled_mantissa x) <= 0)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R
(0 <= F2R {| Fnum := rnd (scaled_mantissa y); Fexp := cexp y |})%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:x = 0%R
(round rnd x <= round rnd y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R

(rnd (scaled_mantissa x) <= rnd 0)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R
(0 <= F2R {| Fnum := rnd (scaled_mantissa y); Fexp := cexp y |})%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:x = 0%R
(round rnd x <= round rnd y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R

(scaled_mantissa x <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R
(0 <= F2R {| Fnum := rnd (scaled_mantissa y); Fexp := cexp y |})%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:x = 0%R
(round rnd x <= round rnd y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R

(scaled_mantissa x <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R
(0 <= F2R {| Fnum := rnd (scaled_mantissa y); Fexp := cexp y |})%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:x = 0%R
(round rnd x <= round rnd y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R

(scaled_mantissa x <= 0 * bpow (- fexp (mag beta x)))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R
(0 <= F2R {| Fnum := rnd (scaled_mantissa y); Fexp := cexp y |})%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:x = 0%R
(round rnd x <= round rnd y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R

(0 <= bpow (- fexp (mag beta x)))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R
(x <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R
(0 <= F2R {| Fnum := rnd (scaled_mantissa y); Fexp := cexp y |})%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:x = 0%R
(round rnd x <= round rnd y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R

(x <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R
(0 <= F2R {| Fnum := rnd (scaled_mantissa y); Fexp := cexp y |})%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:x = 0%R
(round rnd x <= round rnd y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R

(0 <= F2R {| Fnum := rnd (scaled_mantissa y); Fexp := cexp y |})%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:x = 0%R
(round rnd x <= round rnd y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R

(0 <= Fnum {| Fnum := rnd (scaled_mantissa y); Fexp := cexp y |})%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:x = 0%R
(round rnd x <= round rnd y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R

(0 <= rnd (scaled_mantissa y))%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:x = 0%R
(round rnd x <= round rnd y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R

(rnd 0 <= rnd (scaled_mantissa y))%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:x = 0%R
(round rnd x <= round rnd y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R

(0 <= scaled_mantissa y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:x = 0%R
(round rnd x <= round rnd y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R

(0 <= y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R
(0 <= bpow (- cexp y))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:x = 0%R
(round rnd x <= round rnd y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R

(0 <= bpow (- cexp y))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:x = 0%R
(round rnd x <= round rnd y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:x = 0%R

(round rnd x <= round rnd y)%R
(* x = 0 *)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:x = 0%R

(round rnd 0 <= round rnd y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:x = 0%R

(0 <= round rnd y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:x = 0%R

(0 <= Fnum {| Fnum := rnd (scaled_mantissa y); Fexp := cexp y |})%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:x = 0%R

(0 <= rnd (scaled_mantissa y))%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:x = 0%R

(rnd 0 <= rnd (scaled_mantissa y))%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:x = 0%R

(0 <= scaled_mantissa y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:x = 0%R

(0 <= y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:x = 0%R
(0 <= bpow (- cexp y))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hxy:(x <= y)%R
Hx:x = 0%R

(0 <= bpow (- cexp y))%R
apply bpow_ge_0. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd

forall x y : R, generic_format x -> (x <= y)%R -> (x <= round rnd y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd

forall x y : R, generic_format x -> (x <= y)%R -> (x <= round rnd y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hx:generic_format x
Hxy:(x <= y)%R

(x <= round rnd y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hx:generic_format x
Hxy:(x <= y)%R

(round rnd x <= round rnd y)%R
now apply round_le. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd

forall x y : R, generic_format y -> (x <= y)%R -> (round rnd x <= y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd

forall x y : R, generic_format y -> (x <= y)%R -> (round rnd x <= y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hy:generic_format y
Hxy:(x <= y)%R

(round rnd x <= y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R
Hy:generic_format y
Hxy:(x <= y)%R

(round rnd x <= round rnd y)%R
now apply round_le. Qed. End monotone.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall P : R -> R -> Prop, (forall rnd : R -> Z, Valid_rnd rnd -> forall x : R, (0 <= x)%R -> P x (round rnd x)) -> forall rnd : R -> Z, Valid_rnd rnd -> forall x : R, P (Rabs x) (Rabs (round rnd x))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall P : R -> R -> Prop, (forall rnd : R -> Z, Valid_rnd rnd -> forall x : R, (0 <= x)%R -> P x (round rnd x)) -> forall rnd : R -> Z, Valid_rnd rnd -> forall x : R, P (Rabs x) (Rabs (round rnd x))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
P:R -> R -> Prop
HP:forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 <= x0)%R -> P x0 (round rnd0 x0)
rnd:R -> Z
Hr:Valid_rnd rnd
x:R

P (Rabs x) (Rabs (round rnd x))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
P:R -> R -> Prop
HP:forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 <= x0)%R -> P x0 (round rnd0 x0)
rnd:R -> Z
Hr:Valid_rnd rnd
x:R
Hx:(0 <= x)%R

P (Rabs x) (Rabs (round rnd x))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
P:R -> R -> Prop
HP:forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 <= x0)%R -> P x0 (round rnd0 x0)
rnd:R -> Z
Hr:Valid_rnd rnd
x:R
Hx:(x < 0)%R
P (Rabs x) (Rabs (round rnd x))
(* . *)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
P:R -> R -> Prop
HP:forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 <= x0)%R -> P x0 (round rnd0 x0)
rnd:R -> Z
Hr:Valid_rnd rnd
x:R
Hx:(0 <= x)%R

P x (round rnd x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
P:R -> R -> Prop
HP:forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 <= x0)%R -> P x0 (round rnd0 x0)
rnd:R -> Z
Hr:Valid_rnd rnd
x:R
Hx:(0 <= x)%R
(0 <= round rnd x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
P:R -> R -> Prop
HP:forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 <= x0)%R -> P x0 (round rnd0 x0)
rnd:R -> Z
Hr:Valid_rnd rnd
x:R
Hx:(0 <= x)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
P:R -> R -> Prop
HP:forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 <= x0)%R -> P x0 (round rnd0 x0)
rnd:R -> Z
Hr:Valid_rnd rnd
x:R
Hx:(x < 0)%R
P (Rabs x) (Rabs (round rnd x))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
P:R -> R -> Prop
HP:forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 <= x0)%R -> P x0 (round rnd0 x0)
rnd:R -> Z
Hr:Valid_rnd rnd
x:R
Hx:(0 <= x)%R

(0 <= round rnd x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
P:R -> R -> Prop
HP:forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 <= x0)%R -> P x0 (round rnd0 x0)
rnd:R -> Z
Hr:Valid_rnd rnd
x:R
Hx:(0 <= x)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
P:R -> R -> Prop
HP:forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 <= x0)%R -> P x0 (round rnd0 x0)
rnd:R -> Z
Hr:Valid_rnd rnd
x:R
Hx:(x < 0)%R
P (Rabs x) (Rabs (round rnd x))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
P:R -> R -> Prop
HP:forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 <= x0)%R -> P x0 (round rnd0 x0)
rnd:R -> Z
Hr:Valid_rnd rnd
x:R
Hx:(0 <= x)%R

(round rnd 0 <= round rnd x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
P:R -> R -> Prop
HP:forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 <= x0)%R -> P x0 (round rnd0 x0)
rnd:R -> Z
Hr:Valid_rnd rnd
x:R
Hx:(0 <= x)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
P:R -> R -> Prop
HP:forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 <= x0)%R -> P x0 (round rnd0 x0)
rnd:R -> Z
Hr:Valid_rnd rnd
x:R
Hx:(x < 0)%R
P (Rabs x) (Rabs (round rnd x))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
P:R -> R -> Prop
HP:forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 <= x0)%R -> P x0 (round rnd0 x0)
rnd:R -> Z
Hr:Valid_rnd rnd
x:R
Hx:(0 <= x)%R

(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
P:R -> R -> Prop
HP:forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 <= x0)%R -> P x0 (round rnd0 x0)
rnd:R -> Z
Hr:Valid_rnd rnd
x:R
Hx:(x < 0)%R
P (Rabs x) (Rabs (round rnd x))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
P:R -> R -> Prop
HP:forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 <= x0)%R -> P x0 (round rnd0 x0)
rnd:R -> Z
Hr:Valid_rnd rnd
x:R
Hx:(x < 0)%R

P (Rabs x) (Rabs (round rnd x))
(* . *)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
P:R -> R -> Prop
HP:forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 <= x0)%R -> P x0 (round rnd0 x0)
rnd:R -> Z
Hr:Valid_rnd rnd
x:R
Hx:(x < 0)%R

P (- x)%R (Rabs (round rnd x))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
P:R -> R -> Prop
HP:forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 <= x0)%R -> P x0 (round rnd0 x0)
rnd:R -> Z
Hr:Valid_rnd rnd
x:R
Hx:(x < 0)%R

P (- x)%R (- round rnd x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
P:R -> R -> Prop
HP:forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 <= x0)%R -> P x0 (round rnd0 x0)
rnd:R -> Z
Hr:Valid_rnd rnd
x:R
Hx:(x < 0)%R
(round rnd x <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
P:R -> R -> Prop
HP:forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 <= x0)%R -> P x0 (round rnd0 x0)
rnd:R -> Z
Hr:Valid_rnd rnd
x:R
Hx:(x < 0)%R

P (- x)%R (- round rnd (- - x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
P:R -> R -> Prop
HP:forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 <= x0)%R -> P x0 (round rnd0 x0)
rnd:R -> Z
Hr:Valid_rnd rnd
x:R
Hx:(x < 0)%R
(round rnd x <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
P:R -> R -> Prop
HP:forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 <= x0)%R -> P x0 (round rnd0 x0)
rnd:R -> Z
Hr:Valid_rnd rnd
x:R
Hx:(x < 0)%R

P (- x)%R (- - round (Zrnd_opp rnd) (- x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
P:R -> R -> Prop
HP:forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 <= x0)%R -> P x0 (round rnd0 x0)
rnd:R -> Z
Hr:Valid_rnd rnd
x:R
Hx:(x < 0)%R
(round rnd x <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
P:R -> R -> Prop
HP:forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 <= x0)%R -> P x0 (round rnd0 x0)
rnd:R -> Z
Hr:Valid_rnd rnd
x:R
Hx:(x < 0)%R

P (- x)%R (round (Zrnd_opp rnd) (- x))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
P:R -> R -> Prop
HP:forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 <= x0)%R -> P x0 (round rnd0 x0)
rnd:R -> Z
Hr:Valid_rnd rnd
x:R
Hx:(x < 0)%R
(round rnd x <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
P:R -> R -> Prop
HP:forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 <= x0)%R -> P x0 (round rnd0 x0)
rnd:R -> Z
Hr:Valid_rnd rnd
x:R
Hx:(x < 0)%R

(0 <= - x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
P:R -> R -> Prop
HP:forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 <= x0)%R -> P x0 (round rnd0 x0)
rnd:R -> Z
Hr:Valid_rnd rnd
x:R
Hx:(x < 0)%R
(round rnd x <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
P:R -> R -> Prop
HP:forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 <= x0)%R -> P x0 (round rnd0 x0)
rnd:R -> Z
Hr:Valid_rnd rnd
x:R
Hx:(x < 0)%R

(- 0 <= - x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
P:R -> R -> Prop
HP:forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 <= x0)%R -> P x0 (round rnd0 x0)
rnd:R -> Z
Hr:Valid_rnd rnd
x:R
Hx:(x < 0)%R
(round rnd x <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
P:R -> R -> Prop
HP:forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 <= x0)%R -> P x0 (round rnd0 x0)
rnd:R -> Z
Hr:Valid_rnd rnd
x:R
Hx:(x < 0)%R

(x <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
P:R -> R -> Prop
HP:forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 <= x0)%R -> P x0 (round rnd0 x0)
rnd:R -> Z
Hr:Valid_rnd rnd
x:R
Hx:(x < 0)%R
(round rnd x <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
P:R -> R -> Prop
HP:forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 <= x0)%R -> P x0 (round rnd0 x0)
rnd:R -> Z
Hr:Valid_rnd rnd
x:R
Hx:(x < 0)%R

(round rnd x <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
P:R -> R -> Prop
HP:forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 <= x0)%R -> P x0 (round rnd0 x0)
rnd:R -> Z
Hr:Valid_rnd rnd
x:R
Hx:(x < 0)%R

(round rnd x <= round rnd 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
P:R -> R -> Prop
HP:forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 <= x0)%R -> P x0 (round rnd0 x0)
rnd:R -> Z
Hr:Valid_rnd rnd
x:R
Hx:(x < 0)%R

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

forall rnd : R -> Z, Valid_rnd rnd -> forall (x : R) (ex : Z), (fexp ex < ex)%Z -> (bpow (ex - 1) <= Rabs x < bpow ex)%R -> (bpow (ex - 1) <= Rabs (round rnd x) <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall rnd : R -> Z, Valid_rnd rnd -> forall (x : R) (ex : Z), (fexp ex < ex)%Z -> (bpow (ex - 1) <= Rabs x < bpow ex)%R -> (bpow (ex - 1) <= Rabs (round rnd x) <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
Hr:Valid_rnd rnd
x:R
ex:Z
He:(fexp ex < ex)%Z

(bpow (ex - 1) <= Rabs x < bpow ex)%R -> (bpow (ex - 1) <= Rabs (round rnd x) <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
Hr:Valid_rnd rnd
x:R
ex:Z
He:(fexp ex < ex)%Z

forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 <= x0)%R -> (bpow (ex - 1) <= x0 < bpow ex)%R -> (bpow (ex - 1) <= round rnd0 x0 <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
ex:Z
He:(fexp ex < ex)%Z

forall rnd : R -> Z, Valid_rnd rnd -> forall x : R, (0 <= x)%R -> (bpow (ex - 1) <= x < bpow ex)%R -> (bpow (ex - 1) <= round rnd x <= bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
ex:Z
He:(fexp ex < ex)%Z
rnd':R -> Z
Hr:Valid_rnd rnd'
x:R

(bpow (ex - 1) <= x < bpow ex)%R -> (bpow (ex - 1) <= round rnd' x <= bpow ex)%R
apply round_bounded_large_pos... Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall rnd : R -> Z, Valid_rnd rnd -> forall (x : R) (ex : Z), (bpow (ex - 1) <= Rabs x < bpow ex)%R -> round rnd x = 0%R -> (ex <= fexp ex)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall rnd : R -> Z, Valid_rnd rnd -> forall (x : R) (ex : Z), (bpow (ex - 1) <= Rabs x < bpow ex)%R -> round rnd x = 0%R -> (ex <= fexp ex)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
Hr:Valid_rnd rnd
x:R
ex:Z
H1:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:round rnd x = 0%R

(ex <= fexp ex)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
Hr:Valid_rnd rnd
x:R
ex:Z
H1:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:round rnd x = 0%R

Rabs 0 = 0%R -> (ex <= fexp ex)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
Hr:Valid_rnd rnd
x:R
ex:Z
H1:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:round rnd x = 0%R

Rabs (round rnd x) = 0%R -> (ex <= fexp ex)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
Hr:Valid_rnd rnd
x:R
ex:Z
H1:(bpow (ex - 1) <= Rabs x < bpow ex)%R
H2:round rnd x = 0%R

forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 <= x0)%R -> forall ex0 : Z, (bpow (ex0 - 1) <= x0 < bpow ex0)%R -> round rnd0 x0 = 0%R -> (ex0 <= fexp ex0)%Z
beta:radix
fexp:Z -> Z
rnd:R -> Z
Hr:Valid_rnd rnd
x:R
Hx:(0 <= x)%R

forall ex : Z, (bpow (ex - 1) <= x < bpow ex)%R -> round rnd x = 0%R -> (ex <= fexp ex)%Z
now apply exp_small_round_0_pos. Qed. Section monotone_abs. Variable rnd : R -> Z. Context { valid_rnd : Valid_rnd rnd }.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd

forall x y : R, generic_format x -> (x <= Rabs y)%R -> (x <= Rabs (round rnd y))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd

forall x y : R, generic_format x -> (x <= Rabs y)%R -> (x <= Rabs (round rnd y))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R

generic_format x -> (x <= Rabs y)%R -> (x <= Rabs (round rnd y))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R

forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 <= x0)%R -> generic_format x -> (x <= x0)%R -> (x <= round rnd0 x0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

forall rnd : R -> Z, Valid_rnd rnd -> forall x0 : R, (0 <= x0)%R -> generic_format x -> (x <= x0)%R -> (x <= round rnd x0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
rnd':R -> Z
Hrnd:Valid_rnd rnd'
y:R
Hy:(0 <= y)%R

generic_format x -> (x <= y)%R -> (x <= round rnd' y)%R
apply round_ge_generic... Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd

forall x y : R, generic_format y -> (Rabs x <= y)%R -> (Rabs (round rnd x) <= y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd

forall x y : R, generic_format y -> (Rabs x <= y)%R -> (Rabs (round rnd x) <= y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R

generic_format y -> (Rabs x <= y)%R -> (Rabs (round rnd x) <= y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x, y:R

forall rnd0 : R -> Z, Valid_rnd rnd0 -> forall x0 : R, (0 <= x0)%R -> generic_format y -> (x0 <= y)%R -> (round rnd0 x0 <= y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
y:R

forall rnd : R -> Z, Valid_rnd rnd -> forall x : R, (0 <= x)%R -> generic_format y -> (x <= y)%R -> (round rnd x <= y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
y:R
rnd':R -> Z
Hrnd:Valid_rnd rnd'
x:R
Hx:(0 <= x)%R

generic_format y -> (x <= y)%R -> (round rnd' x <= y)%R
apply round_le_generic... Qed. End monotone_abs.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, round Zfloor (- x) = (- round Zceil x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, round Zfloor (- x) = (- round Zceil x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

round Zfloor (- x) = (- round Zceil x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

F2R {| Fnum := Zfloor (scaled_mantissa (- x)); Fexp := cexp (- x) |} = (- F2R {| Fnum := Zceil (scaled_mantissa x); Fexp := cexp x |})%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

F2R {| Fnum := Zfloor (- scaled_mantissa x); Fexp := cexp (- x) |} = (- F2R {| Fnum := Zceil (scaled_mantissa x); Fexp := cexp x |})%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

F2R {| Fnum := Zfloor (- scaled_mantissa x); Fexp := cexp (- x) |} = F2R {| Fnum := - Zceil (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

F2R {| Fnum := Zfloor (- scaled_mantissa x); Fexp := cexp (- x) |} = F2R {| Fnum := - - Zfloor (- scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

F2R {| Fnum := Zfloor (- scaled_mantissa x); Fexp := cexp (- x) |} = F2R {| Fnum := Zfloor (- scaled_mantissa x); Fexp := cexp x |}
now rewrite cexp_opp. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, round Zceil (- x) = (- round Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, round Zceil (- x) = (- round Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

round Zceil (- x) = (- round Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

F2R {| Fnum := Zceil (scaled_mantissa (- x)); Fexp := cexp (- x) |} = (- F2R {| Fnum := Zfloor (scaled_mantissa x); Fexp := cexp x |})%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

F2R {| Fnum := Zceil (- scaled_mantissa x); Fexp := cexp (- x) |} = (- F2R {| Fnum := Zfloor (scaled_mantissa x); Fexp := cexp x |})%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

F2R {| Fnum := Zceil (- scaled_mantissa x); Fexp := cexp (- x) |} = F2R {| Fnum := - Zfloor (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

F2R {| Fnum := - Zfloor (- - scaled_mantissa x); Fexp := cexp (- x) |} = F2R {| Fnum := - Zfloor (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

F2R {| Fnum := - Zfloor (scaled_mantissa x); Fexp := cexp (- x) |} = F2R {| Fnum := - Zfloor (scaled_mantissa x); Fexp := cexp x |}
now rewrite cexp_opp. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, round Ztrunc (- x) = (- round Ztrunc x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, round Ztrunc (- x) = (- round Ztrunc x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

round Ztrunc (- x) = (- round Ztrunc x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

F2R {| Fnum := Ztrunc (scaled_mantissa (- x)); Fexp := cexp (- x) |} = (- F2R {| Fnum := Ztrunc (scaled_mantissa x); Fexp := cexp x |})%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

F2R {| Fnum := - Ztrunc (scaled_mantissa x); Fexp := cexp x |} = (- F2R {| Fnum := Ztrunc (scaled_mantissa x); Fexp := cexp x |})%R
apply F2R_Zopp. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, round Ztrunc (Rabs x) = Rabs (round Ztrunc x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, round Ztrunc (Rabs x) = Rabs (round Ztrunc x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

round Ztrunc (Rabs x) = Rabs (round Ztrunc x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

Rabs (round Ztrunc x) = round Ztrunc (Rabs x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

Rabs (round Ztrunc x) = round Ztrunc (if Rcase_abs x then (- x)%R else x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x < 0)%R

Rabs (round Ztrunc x) = round Ztrunc (- x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x >= 0)%R
Rabs (round Ztrunc x) = round Ztrunc x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x < 0)%R

Rabs (round Ztrunc x) = (- round Ztrunc x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x >= 0)%R
Rabs (round Ztrunc x) = round Ztrunc x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x < 0)%R

(round Ztrunc x <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x >= 0)%R
Rabs (round Ztrunc x) = round Ztrunc x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x < 0)%R

(round Ztrunc x <= round Ztrunc 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x >= 0)%R
Rabs (round Ztrunc x) = round Ztrunc x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x < 0)%R

(x <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x >= 0)%R
Rabs (round Ztrunc x) = round Ztrunc x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x >= 0)%R

Rabs (round Ztrunc x) = round Ztrunc x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x >= 0)%R

(0 <= round Ztrunc x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x >= 0)%R

(round Ztrunc 0 <= round Ztrunc x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x >= 0)%R

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

forall x : R, round Zaway (- x) = (- round Zaway x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, round Zaway (- x) = (- round Zaway x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

round Zaway (- x) = (- round Zaway x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

F2R {| Fnum := Zaway (scaled_mantissa (- x)); Fexp := cexp (- x) |} = (- F2R {| Fnum := Zaway (scaled_mantissa x); Fexp := cexp x |})%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

F2R {| Fnum := - Zaway (scaled_mantissa x); Fexp := cexp x |} = (- F2R {| Fnum := Zaway (scaled_mantissa x); Fexp := cexp x |})%R
apply F2R_Zopp. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, round Zaway (Rabs x) = Rabs (round Zaway x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, round Zaway (Rabs x) = Rabs (round Zaway x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

round Zaway (Rabs x) = Rabs (round Zaway x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

Rabs (round Zaway x) = round Zaway (Rabs x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

Rabs (round Zaway x) = round Zaway (if Rcase_abs x then (- x)%R else x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x < 0)%R

Rabs (round Zaway x) = round Zaway (- x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x >= 0)%R
Rabs (round Zaway x) = round Zaway x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x < 0)%R

Rabs (round Zaway x) = (- round Zaway x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x >= 0)%R
Rabs (round Zaway x) = round Zaway x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x < 0)%R

(round Zaway x <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x >= 0)%R
Rabs (round Zaway x) = round Zaway x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x < 0)%R

(round Zaway x <= round Zaway 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x >= 0)%R
Rabs (round Zaway x) = round Zaway x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x < 0)%R

(x <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x >= 0)%R
Rabs (round Zaway x) = round Zaway x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x >= 0)%R

Rabs (round Zaway x) = round Zaway x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x >= 0)%R

(0 <= round Zaway x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x >= 0)%R

(round Zaway 0 <= round Zaway x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x >= 0)%R

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

forall x : R, (0 <= x)%R -> round Ztrunc x = round Zfloor x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, (0 <= x)%R -> round Ztrunc x = round Zfloor x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(0 <= x)%R

round Ztrunc x = round Zfloor x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(0 <= x)%R

F2R {| Fnum := if Rlt_bool (scaled_mantissa x) 0 then Zceil (scaled_mantissa x) else Zfloor (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zfloor (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(0 <= x)%R

(scaled_mantissa x < 0)%R -> F2R {| Fnum := Zceil (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zfloor (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(0 <= x)%R
(0 <= scaled_mantissa x)%R -> F2R {| Fnum := Zfloor (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zfloor (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(0 <= x)%R
H:(scaled_mantissa x < 0)%R

F2R {| Fnum := Zceil (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zfloor (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(0 <= x)%R
(0 <= scaled_mantissa x)%R -> F2R {| Fnum := Zfloor (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zfloor (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(0 <= x)%R
H:(scaled_mantissa x < 0)%R

(0 <= scaled_mantissa x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(0 <= x)%R
(0 <= scaled_mantissa x)%R -> F2R {| Fnum := Zfloor (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zfloor (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(0 <= x)%R
H:(scaled_mantissa x < 0)%R

(0 * bpow (- cexp x) <= scaled_mantissa x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(0 <= x)%R
(0 <= scaled_mantissa x)%R -> F2R {| Fnum := Zfloor (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zfloor (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(0 <= x)%R
H:(scaled_mantissa x < 0)%R

(0 <= bpow (- cexp x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(0 <= x)%R
(0 <= scaled_mantissa x)%R -> F2R {| Fnum := Zfloor (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zfloor (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(0 <= x)%R

(0 <= scaled_mantissa x)%R -> F2R {| Fnum := Zfloor (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zfloor (scaled_mantissa x); Fexp := cexp x |}
easy. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, (x <= 0)%R -> round Ztrunc x = round Zceil x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, (x <= 0)%R -> round Ztrunc x = round Zceil x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x <= 0)%R

round Ztrunc x = round Zceil x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x <= 0)%R

F2R {| Fnum := if Rlt_bool (scaled_mantissa x) 0 then Zceil (scaled_mantissa x) else Zfloor (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zceil (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x <= 0)%R

(scaled_mantissa x < 0)%R -> F2R {| Fnum := Zceil (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zceil (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x <= 0)%R
(0 <= scaled_mantissa x)%R -> F2R {| Fnum := Zfloor (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zceil (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x <= 0)%R

(0 <= scaled_mantissa x)%R -> F2R {| Fnum := Zfloor (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zceil (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x <= 0)%R
H:(0 < scaled_mantissa x)%R

F2R {| Fnum := Zfloor (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zceil (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x <= 0)%R
H:0%R = scaled_mantissa x
F2R {| Fnum := Zfloor (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zceil (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x <= 0)%R
H:(0 < scaled_mantissa x)%R

(scaled_mantissa x <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x <= 0)%R
H:0%R = scaled_mantissa x
F2R {| Fnum := Zfloor (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zceil (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x <= 0)%R
H:(0 < scaled_mantissa x)%R

(scaled_mantissa x <= 0 * bpow (- cexp x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x <= 0)%R
H:0%R = scaled_mantissa x
F2R {| Fnum := Zfloor (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zceil (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x <= 0)%R
H:(0 < scaled_mantissa x)%R

(0 <= bpow (- cexp x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x <= 0)%R
H:0%R = scaled_mantissa x
F2R {| Fnum := Zfloor (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zceil (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x <= 0)%R
H:0%R = scaled_mantissa x

F2R {| Fnum := Zfloor (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zceil (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x <= 0)%R
H:0%R = scaled_mantissa x

F2R {| Fnum := Zfloor 0; Fexp := cexp x |} = F2R {| Fnum := Zceil 0; Fexp := cexp x |}
now rewrite Zfloor_IZR, Zceil_IZR. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, (0 <= x)%R -> round Zaway x = round Zceil x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, (0 <= x)%R -> round Zaway x = round Zceil x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(0 <= x)%R

round Zaway x = round Zceil x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(0 <= x)%R

F2R {| Fnum := if Rlt_bool (scaled_mantissa x) 0 then Zfloor (scaled_mantissa x) else Zceil (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zceil (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(0 <= x)%R

(scaled_mantissa x < 0)%R -> F2R {| Fnum := Zfloor (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zceil (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(0 <= x)%R
(0 <= scaled_mantissa x)%R -> F2R {| Fnum := Zceil (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zceil (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(0 <= x)%R
H:(scaled_mantissa x < 0)%R

F2R {| Fnum := Zfloor (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zceil (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(0 <= x)%R
(0 <= scaled_mantissa x)%R -> F2R {| Fnum := Zceil (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zceil (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(0 <= x)%R
H:(scaled_mantissa x < 0)%R

(0 <= scaled_mantissa x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(0 <= x)%R
(0 <= scaled_mantissa x)%R -> F2R {| Fnum := Zceil (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zceil (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(0 <= x)%R
H:(scaled_mantissa x < 0)%R

(0 * bpow (- cexp x) <= scaled_mantissa x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(0 <= x)%R
(0 <= scaled_mantissa x)%R -> F2R {| Fnum := Zceil (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zceil (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(0 <= x)%R
H:(scaled_mantissa x < 0)%R

(0 <= bpow (- cexp x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(0 <= x)%R
(0 <= scaled_mantissa x)%R -> F2R {| Fnum := Zceil (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zceil (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(0 <= x)%R

(0 <= scaled_mantissa x)%R -> F2R {| Fnum := Zceil (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zceil (scaled_mantissa x); Fexp := cexp x |}
easy. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, (x <= 0)%R -> round Zaway x = round Zfloor x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, (x <= 0)%R -> round Zaway x = round Zfloor x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x <= 0)%R

round Zaway x = round Zfloor x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x <= 0)%R

F2R {| Fnum := if Rlt_bool (scaled_mantissa x) 0 then Zfloor (scaled_mantissa x) else Zceil (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zfloor (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x <= 0)%R

(scaled_mantissa x < 0)%R -> F2R {| Fnum := Zfloor (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zfloor (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x <= 0)%R
(0 <= scaled_mantissa x)%R -> F2R {| Fnum := Zceil (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zfloor (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x <= 0)%R

(0 <= scaled_mantissa x)%R -> F2R {| Fnum := Zceil (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zfloor (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x <= 0)%R
H:(0 < scaled_mantissa x)%R

F2R {| Fnum := Zceil (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zfloor (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x <= 0)%R
H:0%R = scaled_mantissa x
F2R {| Fnum := Zceil (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zfloor (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x <= 0)%R
H:(0 < scaled_mantissa x)%R

(scaled_mantissa x <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x <= 0)%R
H:0%R = scaled_mantissa x
F2R {| Fnum := Zceil (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zfloor (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x <= 0)%R
H:(0 < scaled_mantissa x)%R

(scaled_mantissa x <= 0 * bpow (- cexp x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x <= 0)%R
H:0%R = scaled_mantissa x
F2R {| Fnum := Zceil (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zfloor (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x <= 0)%R
H:(0 < scaled_mantissa x)%R

(0 <= bpow (- cexp x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x <= 0)%R
H:0%R = scaled_mantissa x
F2R {| Fnum := Zceil (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zfloor (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x <= 0)%R
H:0%R = scaled_mantissa x

F2R {| Fnum := Zceil (scaled_mantissa x); Fexp := cexp x |} = F2R {| Fnum := Zfloor (scaled_mantissa x); Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x <= 0)%R
H:0%R = scaled_mantissa x

F2R {| Fnum := Zceil 0; Fexp := cexp x |} = F2R {| Fnum := Zfloor 0; Fexp := cexp x |}
now rewrite Zfloor_IZR, Zceil_IZR. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall rnd : R -> Z, Valid_rnd rnd -> forall x : R, generic_format (round rnd x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall rnd : R -> Z, Valid_rnd rnd -> forall x : R, generic_format (round rnd x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R

generic_format (round rnd x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(x < 0)%R

generic_format (round rnd x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:x = 0%R
generic_format (round rnd x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(x > 0)%R
generic_format (round rnd x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(x < 0)%R

generic_format (round rnd (- - x))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:x = 0%R
generic_format (round rnd x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(x > 0)%R
generic_format (round rnd x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(x < 0)%R
Hr:round rnd (- - x) = round Zfloor (- - x)

generic_format (round Zfloor (- - x))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(x < 0)%R
Hr:round rnd (- - x) = round Zceil (- - x)
generic_format (round Zceil (- - x))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:x = 0%R
generic_format (round rnd x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(x > 0)%R
generic_format (round rnd x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(x < 0)%R
Hr:round rnd (- - x) = round Zfloor (- - x)

generic_format (- round Zceil (- x))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(x < 0)%R
Hr:round rnd (- - x) = round Zceil (- - x)
generic_format (round Zceil (- - x))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:x = 0%R
generic_format (round rnd x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(x > 0)%R
generic_format (round rnd x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(x < 0)%R
Hr:round rnd (- - x) = round Zfloor (- - x)

generic_format (round Zceil (- x))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(x < 0)%R
Hr:round rnd (- - x) = round Zceil (- - x)
generic_format (round Zceil (- - x))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:x = 0%R
generic_format (round rnd x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(x > 0)%R
generic_format (round rnd x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(x < 0)%R
Hr:round rnd (- - x) = round Zfloor (- - x)

(0 < - x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(x < 0)%R
Hr:round rnd (- - x) = round Zceil (- - x)
generic_format (round Zceil (- - x))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:x = 0%R
generic_format (round rnd x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(x > 0)%R
generic_format (round rnd x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(x < 0)%R
Hr:round rnd (- - x) = round Zceil (- - x)

generic_format (round Zceil (- - x))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:x = 0%R
generic_format (round rnd x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(x > 0)%R
generic_format (round rnd x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(x < 0)%R
Hr:round rnd (- - x) = round Zceil (- - x)

generic_format (- round Zfloor (- x))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:x = 0%R
generic_format (round rnd x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(x > 0)%R
generic_format (round rnd x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(x < 0)%R
Hr:round rnd (- - x) = round Zceil (- - x)

generic_format (round Zfloor (- x))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:x = 0%R
generic_format (round rnd x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(x > 0)%R
generic_format (round rnd x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(x < 0)%R
Hr:round rnd (- - x) = round Zceil (- - x)

(0 < - x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:x = 0%R
generic_format (round rnd x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(x > 0)%R
generic_format (round rnd x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:x = 0%R

generic_format (round rnd x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(x > 0)%R
generic_format (round rnd x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:x = 0%R

generic_format (round rnd 0)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(x > 0)%R
generic_format (round rnd x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:x = 0%R

generic_format 0
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(x > 0)%R
generic_format (round rnd x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
Zrnd:Valid_rnd rnd
x:R
Hx:(x > 0)%R

generic_format (round rnd x)
now apply generic_format_round_pos. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, Rnd_DN_pt generic_format x (round Zfloor x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, Rnd_DN_pt generic_format x (round Zfloor x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

Rnd_DN_pt generic_format x (round Zfloor x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

generic_format (round Zfloor x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
(round Zfloor x <= x)%R /\ (forall g : R, generic_format g -> (g <= x)%R -> (g <= round Zfloor x)%R)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

(round Zfloor x <= x)%R /\ (forall g : R, generic_format g -> (g <= x)%R -> (g <= round Zfloor x)%R)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

(round Zfloor x <= x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
forall g : R, generic_format g -> (g <= x)%R -> (g <= round Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

(round Zfloor x <= scaled_mantissa x * bpow (cexp x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
forall g : R, generic_format g -> (g <= x)%R -> (g <= round Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

(IZR (Fnum {| Fnum := Zfloor (scaled_mantissa x); Fexp := cexp x |}) * bpow (Fexp {| Fnum := Zfloor (scaled_mantissa x); Fexp := cexp x |}) <= scaled_mantissa x * bpow (cexp x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
forall g : R, generic_format g -> (g <= x)%R -> (g <= round Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

(IZR (Zfloor (scaled_mantissa x)) * bpow (cexp x) <= scaled_mantissa x * bpow (cexp x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
forall g : R, generic_format g -> (g <= x)%R -> (g <= round Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

(0 <= bpow (cexp x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
(IZR (Zfloor (scaled_mantissa x)) <= scaled_mantissa x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
forall g : R, generic_format g -> (g <= x)%R -> (g <= round Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

(IZR (Zfloor (scaled_mantissa x)) <= scaled_mantissa x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
forall g : R, generic_format g -> (g <= x)%R -> (g <= round Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

forall g : R, generic_format g -> (g <= x)%R -> (g <= round Zfloor x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x, g:R
Hg:generic_format g
Hgx:(g <= x)%R

(g <= round Zfloor x)%R
apply round_ge_generic... Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

satisfies_any generic_format
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

satisfies_any generic_format
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

generic_format 0
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
forall x : R, generic_format x -> generic_format (- x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
round_pred_total (Rnd_DN_pt generic_format)
(* symmetric set *)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, generic_format x -> generic_format (- x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
round_pred_total (Rnd_DN_pt generic_format)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

round_pred_total (Rnd_DN_pt generic_format)
(* round down *)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

exists f : R, Rnd_DN_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

Rnd_DN_pt generic_format x ?f
apply round_DN_pt. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, Rnd_UP_pt generic_format x (round Zceil x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, Rnd_UP_pt generic_format x (round Zceil x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

Rnd_UP_pt generic_format x (round Zceil x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

Rnd_UP_pt generic_format (- - x) (round Zceil (- - x))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

Rnd_UP_pt generic_format (- - x) (- round Zfloor (- x))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

forall x0 : R, generic_format x0 -> generic_format (- x0)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Rnd_DN_pt generic_format (- x) (round Zfloor (- x))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

Rnd_DN_pt generic_format (- x) (round Zfloor (- x))
apply round_DN_pt. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, Rnd_ZR_pt generic_format x (round Ztrunc x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, Rnd_ZR_pt generic_format x (round Ztrunc x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

Rnd_ZR_pt generic_format x (round Ztrunc x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(0 <= x)%R

Rnd_DN_pt generic_format x (round Ztrunc x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x <= 0)%R
Rnd_UP_pt generic_format x (round Ztrunc x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(0 <= x)%R

Rnd_DN_pt generic_format x (round Zfloor x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x <= 0)%R
Rnd_UP_pt generic_format x (round Ztrunc x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x <= 0)%R

Rnd_UP_pt generic_format x (round Ztrunc x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x <= 0)%R

Rnd_UP_pt generic_format x (round Zceil x)
apply round_UP_pt. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall (x : R) (ex : Z), (bpow (ex - 1) <= x < bpow ex)%R -> (ex <= fexp ex)%Z -> round Zfloor x = 0%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall (x : R) (ex : Z), (bpow (ex - 1) <= x < bpow ex)%R -> (ex <= fexp ex)%Z -> round Zfloor x = 0%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
He:(ex <= fexp ex)%Z

round Zfloor x = 0%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
He:(ex <= fexp ex)%Z

round Zfloor x = F2R {| Fnum := 0; Fexp := cexp x |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
He:(ex <= fexp ex)%Z

round Zfloor x = F2R {| Fnum := Zfloor (x * bpow (- fexp ex)); Fexp := cexp x |}
now rewrite <- cexp_fexp_pos with (1 := Hx). Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, ~ generic_format x -> (round Zfloor x < x < round Zceil x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, ~ generic_format x -> (round Zfloor x < x < round Zceil x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Fx:~ generic_format x

(round Zfloor x < x < round Zceil x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Fx:~ generic_format x

(round Zfloor x <= x <= round Zceil x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Fx:~ generic_format x
Hx:(round Zfloor x <= x <= round Zceil x)%R
(round Zfloor x < x < round Zceil x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Fx:~ generic_format x

(round Zfloor x <= x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Fx:~ generic_format x
(x <= round Zceil x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Fx:~ generic_format x
Hx:(round Zfloor x <= x <= round Zceil x)%R
(round Zfloor x < x < round Zceil x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Fx:~ generic_format x

(x <= round Zceil x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Fx:~ generic_format x
Hx:(round Zfloor x <= x <= round Zceil x)%R
(round Zfloor x < x < round Zceil x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Fx:~ generic_format x
Hx:(round Zfloor x <= x <= round Zceil x)%R

(round Zfloor x < x < round Zceil x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Fx:~ generic_format x
Hx:(round Zfloor x <= x <= round Zceil x)%R

(round Zfloor x < x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Fx:~ generic_format x
Hx:(round Zfloor x <= x <= round Zceil x)%R
(x < round Zceil x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Fx:~ generic_format x
Hx:(round Zfloor x <= x)%R

(round Zfloor x < x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Fx:~ generic_format x
Hx:(round Zfloor x <= x <= round Zceil x)%R
(x < round Zceil x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Fx:~ generic_format x
Hx:(round Zfloor x <= x)%R
Hle:(x <= round Zfloor x)%R

False
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Fx:~ generic_format x
Hx:(round Zfloor x <= x <= round Zceil x)%R
(x < round Zceil x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Fx:~ generic_format x
Hx:(round Zfloor x <= x)%R
Hle:(x <= round Zfloor x)%R
H:x = round Zfloor x

False
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Fx:~ generic_format x
Hx:(round Zfloor x <= x <= round Zceil x)%R
(x < round Zceil x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Fx:~ generic_format (round Zfloor x)
Hx:(round Zfloor x <= x)%R
Hle:(x <= round Zfloor x)%R
H:x = round Zfloor x

False
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Fx:~ generic_format x
Hx:(round Zfloor x <= x <= round Zceil x)%R
(x < round Zceil x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(round Zfloor x <= x)%R
Hle:(x <= round Zfloor x)%R
H:x = round Zfloor x

generic_format (round Zfloor x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Fx:~ generic_format x
Hx:(round Zfloor x <= x <= round Zceil x)%R
(x < round Zceil x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Fx:~ generic_format x
Hx:(round Zfloor x <= x <= round Zceil x)%R

(x < round Zceil x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Fx:~ generic_format x
Hx:(x <= round Zceil x)%R

(x < round Zceil x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Fx:~ generic_format x
Hx:(x <= round Zceil x)%R
Hle:(round Zceil x <= x)%R

False
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Fx:~ generic_format x
Hx:(x <= round Zceil x)%R
Hle:(round Zceil x <= x)%R
H:x = round Zceil x

False
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Fx:~ generic_format (round Zceil x)
Hx:(x <= round Zceil x)%R
Hle:(round Zceil x <= x)%R
H:x = round Zceil x

False
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:(x <= round Zceil x)%R
Hle:(round Zceil x <= x)%R
H:x = round Zceil x

generic_format (round Zceil x)
apply generic_format_round... Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall (x : R) (ex : Z), (bpow (ex - 1) <= x < bpow ex)%R -> (ex <= fexp ex)%Z -> round Zceil x = bpow (fexp ex)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall (x : R) (ex : Z), (bpow (ex - 1) <= x < bpow ex)%R -> (ex <= fexp ex)%Z -> round Zceil x = bpow (fexp ex)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
He:(ex <= fexp ex)%Z

round Zceil x = bpow (fexp ex)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
He:(ex <= fexp ex)%Z

round Zceil x = F2R {| Fnum := 1; Fexp := fexp ex |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
ex:Z
Hx:(bpow (ex - 1) <= x < bpow ex)%R
He:(ex <= fexp ex)%Z

round Zceil x = F2R {| Fnum := Zceil (x * bpow (- fexp ex)); Fexp := fexp ex |}
now rewrite <- cexp_fexp_pos with (1 := Hx). Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, generic_format x \/ ~ generic_format x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, generic_format x \/ ~ generic_format x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

generic_format x \/ ~ generic_format x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:round Zfloor x = x

generic_format x \/ ~ generic_format x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:round Zfloor x <> x
generic_format x \/ ~ generic_format x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:round Zfloor x = x

generic_format x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:round Zfloor x <> x
generic_format x \/ ~ generic_format x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:round Zfloor x = x

generic_format (round Zfloor x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:round Zfloor x <> x
generic_format x \/ ~ generic_format x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:round Zfloor x <> x

generic_format x \/ ~ generic_format x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:round Zfloor x <> x

~ generic_format x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:round Zfloor x <> x
H:generic_format x

False
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hx:round Zfloor x <> x
H:generic_format x

round Zfloor x = x
apply round_generic... Qed. Section round_large. Variable rnd : R -> Z. Context { valid_rnd : Valid_rnd rnd }.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd

forall (x : R) (e : Z), (0 < round rnd x)%R -> (bpow e <= x)%R -> (bpow e <= round rnd x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd

forall (x : R) (e : Z), (0 < round rnd x)%R -> (bpow e <= x)%R -> (bpow e <= round rnd x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
e:Z
Hd:(0 < round rnd x)%R
Hex:(bpow e <= x)%R

(bpow e <= round rnd x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
e:Z
Hd:(0 < round rnd x)%R
Hex:(bpow e <= x)%R
ex:Z
He:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R

(bpow e <= round rnd x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
e:Z
Hd:(0 < round rnd x)%R
Hex:(bpow e <= x)%R
ex:Z
He:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R

(0 < x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
e:Z
Hd:(0 < round rnd x)%R
Hex:(bpow e <= x)%R
ex:Z
He:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:(0 < x)%R
(bpow e <= round rnd x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
e:Z
Hd:(0 < round rnd x)%R
Hex:(bpow e <= x)%R
ex:Z
He:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R

(0 < bpow e)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
e:Z
Hd:(0 < round rnd x)%R
Hex:(bpow e <= x)%R
ex:Z
He:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:(0 < x)%R
(bpow e <= round rnd x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
e:Z
Hd:(0 < round rnd x)%R
Hex:(bpow e <= x)%R
ex:Z
He:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:(0 < x)%R

(bpow e <= round rnd x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
e:Z
Hd:(0 < round rnd x)%R
Hex:(bpow e <= x)%R
ex:Z
He:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:(0 < x)%R

(bpow e <= round rnd x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
e:Z
Hd:(0 < round rnd x)%R
Hex:(bpow e <= x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R
Hx:(0 < x)%R

(bpow e <= round rnd x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
e:Z
Hd:(0 < round rnd x)%R
Hex:(bpow e <= x)%R
ex:Z
He:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hx:(0 < x)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
e:Z
Hd:(0 < round rnd x)%R
Hex:(bpow e <= x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R
Hx:(0 < x)%R

(bpow e <= round rnd x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
e:Z
Hd:(0 < round rnd x)%R
Hex:(bpow e <= x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R
Hx:(0 < x)%R

(bpow e <= bpow (ex - 1))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
e:Z
Hd:(0 < round rnd x)%R
Hex:(bpow e <= x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R
Hx:(0 < x)%R
(bpow (ex - 1) <= round rnd x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
e:Z
Hd:(0 < round rnd x)%R
Hex:(bpow e <= x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R
Hx:(0 < x)%R

(e <= ex - 1)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
e:Z
Hd:(0 < round rnd x)%R
Hex:(bpow e <= x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R
Hx:(0 < x)%R
(bpow (ex - 1) <= round rnd x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
e:Z
Hd:(0 < round rnd x)%R
Hex:(bpow e <= x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R
Hx:(0 < x)%R

(e < ex)%Z -> (e <= ex - 1)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
e:Z
Hd:(0 < round rnd x)%R
Hex:(bpow e <= x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R
Hx:(0 < x)%R
(e < ex)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
e:Z
Hd:(0 < round rnd x)%R
Hex:(bpow e <= x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R
Hx:(0 < x)%R
(bpow (ex - 1) <= round rnd x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
e:Z
Hd:(0 < round rnd x)%R
Hex:(bpow e <= x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R
Hx:(0 < x)%R

(e < ex)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
e:Z
Hd:(0 < round rnd x)%R
Hex:(bpow e <= x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R
Hx:(0 < x)%R
(bpow (ex - 1) <= round rnd x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
e:Z
Hd:(0 < round rnd x)%R
Hex:(bpow e <= x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R
Hx:(0 < x)%R

(bpow e < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
e:Z
Hd:(0 < round rnd x)%R
Hex:(bpow e <= x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R
Hx:(0 < x)%R
(bpow (ex - 1) <= round rnd x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
e:Z
Hd:(0 < round rnd x)%R
Hex:(bpow e <= x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R
Hx:(0 < x)%R

(bpow (ex - 1) <= round rnd x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
e:Z
Hd:(0 < round rnd x)%R
Hex:(bpow e <= x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R
Hx:(0 < x)%R
H:(ex <= fexp ex)%Z

(bpow (ex - 1) <= round rnd x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
e:Z
Hd:(0 < round rnd x)%R
Hex:(bpow e <= x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R
Hx:(0 < x)%R
H:(fexp ex < ex)%Z
(bpow (ex - 1) <= round rnd x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
e:Z
Hd:(0 < round rnd x)%R
Hex:(bpow e <= x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R
Hx:(0 < x)%R
H:(ex <= fexp ex)%Z
Hr:round rnd x = 0%R

(bpow (ex - 1) <= round rnd x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
e:Z
Hd:(0 < round rnd x)%R
Hex:(bpow e <= x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R
Hx:(0 < x)%R
H:(ex <= fexp ex)%Z
Hr:round rnd x = bpow (fexp ex)
(bpow (ex - 1) <= round rnd x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
e:Z
Hd:(0 < round rnd x)%R
Hex:(bpow e <= x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R
Hx:(0 < x)%R
H:(fexp ex < ex)%Z
(bpow (ex - 1) <= round rnd x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
e:Z
Hd:(0 < 0)%R
Hex:(bpow e <= x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R
Hx:(0 < x)%R
H:(ex <= fexp ex)%Z
Hr:round rnd x = 0%R

(bpow (ex - 1) <= round rnd x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
e:Z
Hd:(0 < round rnd x)%R
Hex:(bpow e <= x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R
Hx:(0 < x)%R
H:(ex <= fexp ex)%Z
Hr:round rnd x = bpow (fexp ex)
(bpow (ex - 1) <= round rnd x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
e:Z
Hd:(0 < round rnd x)%R
Hex:(bpow e <= x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R
Hx:(0 < x)%R
H:(fexp ex < ex)%Z
(bpow (ex - 1) <= round rnd x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
e:Z
Hd:(0 < round rnd x)%R
Hex:(bpow e <= x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R
Hx:(0 < x)%R
H:(ex <= fexp ex)%Z
Hr:round rnd x = bpow (fexp ex)

(bpow (ex - 1) <= round rnd x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
e:Z
Hd:(0 < round rnd x)%R
Hex:(bpow e <= x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R
Hx:(0 < x)%R
H:(fexp ex < ex)%Z
(bpow (ex - 1) <= round rnd x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
e:Z
Hd:(0 < round rnd x)%R
Hex:(bpow e <= x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R
Hx:(0 < x)%R
H:(ex <= fexp ex)%Z
Hr:round rnd x = bpow (fexp ex)

(bpow (ex - 1) <= bpow (fexp ex))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
e:Z
Hd:(0 < round rnd x)%R
Hex:(bpow e <= x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R
Hx:(0 < x)%R
H:(fexp ex < ex)%Z
(bpow (ex - 1) <= round rnd x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
e:Z
Hd:(0 < round rnd x)%R
Hex:(bpow e <= x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R
Hx:(0 < x)%R
H:(ex <= fexp ex)%Z
Hr:round rnd x = bpow (fexp ex)

(ex - 1 <= fexp ex)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
e:Z
Hd:(0 < round rnd x)%R
Hex:(bpow e <= x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R
Hx:(0 < x)%R
H:(fexp ex < ex)%Z
(bpow (ex - 1) <= round rnd x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
e:Z
Hd:(0 < round rnd x)%R
Hex:(bpow e <= x)%R
ex:Z
He:(bpow (ex - 1) <= x < bpow ex)%R
Hx:(0 < x)%R
H:(fexp ex < ex)%Z

(bpow (ex - 1) <= round rnd x)%R
apply (round_bounded_large_pos rnd x ex H He). Qed. End round_large.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, round Ztrunc x <> 0%R -> mag beta (round Ztrunc x) = mag beta x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, round Ztrunc x <> 0%R -> mag beta (round Ztrunc x) = mag beta x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zr:round Ztrunc x <> 0%R

mag beta (round Ztrunc x) = mag beta x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zr:round Ztrunc x <> 0%R
Zx:x = 0%R

mag beta (round Ztrunc x) = mag beta x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zr:round Ztrunc x <> 0%R
Zx:x <> 0%R
mag beta (round Ztrunc x) = mag beta x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zr:round Ztrunc x <> 0%R
Zx:x <> 0%R

mag beta (round Ztrunc x) = mag beta x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zr:round Ztrunc x <> 0%R
Zx:x <> 0%R

(bpow (mag beta x - 1) <= Rabs (round Ztrunc x) < bpow (mag beta x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zr:round Ztrunc x <> 0%R
Zx:x <> 0%R
ex:Z
Ex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R

(bpow (ex - 1) <= Rabs (round Ztrunc x) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zr:round Ztrunc x <> 0%R
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R

(bpow (ex - 1) <= Rabs (round Ztrunc x) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zr:round Ztrunc x <> 0%R
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R

(bpow (ex - 1) <= round Ztrunc (Rabs x) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zr:round Ztrunc x <> 0%R
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R

(bpow (ex - 1) <= round Ztrunc (Rabs x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zr:round Ztrunc x <> 0%R
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(round Ztrunc (Rabs x) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zr:round Ztrunc x <> 0%R
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R

(0 < round Ztrunc (Rabs x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zr:round Ztrunc x <> 0%R
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(bpow (ex - 1) <= Rabs x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zr:round Ztrunc x <> 0%R
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(round Ztrunc (Rabs x) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zr:round Ztrunc x <> 0%R
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R

(0 < Rabs (round Ztrunc x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zr:round Ztrunc x <> 0%R
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(bpow (ex - 1) <= Rabs x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zr:round Ztrunc x <> 0%R
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(round Ztrunc (Rabs x) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zr:round Ztrunc x <> 0%R
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R

(bpow (ex - 1) <= Rabs x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zr:round Ztrunc x <> 0%R
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(round Ztrunc (Rabs x) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zr:round Ztrunc x <> 0%R
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R

(round Ztrunc (Rabs x) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zr:round Ztrunc x <> 0%R
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R

(round Ztrunc (Rabs x) <= Rabs x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zr:round Ztrunc x <> 0%R
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R

(round Zfloor (Rabs x) <= Rabs x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zr:round Ztrunc x <> 0%R
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(0 <= Rabs x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zr:round Ztrunc x <> 0%R
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R

(0 <= Rabs x)%R
apply Rabs_pos. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall rnd : R -> Z, Valid_rnd rnd -> forall x : R, round rnd x <> 0%R -> mag beta (round rnd x) = mag beta x \/ Rabs (round rnd x) = bpow (Z.max (mag beta x) (fexp (mag beta x)))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall rnd : R -> Z, Valid_rnd rnd -> forall x : R, round rnd x <> 0%R -> mag beta (round rnd x) = mag beta x \/ Rabs (round rnd x) = bpow (Z.max (mag beta x) (fexp (mag beta x)))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
rnd:R -> Z
Hrnd:Valid_rnd rnd
x:R

round rnd x <> 0%R -> mag beta (round rnd x) = mag beta x \/ Rabs (round rnd x) = bpow (Z.max (mag beta x) (fexp (mag beta x)))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

round Ztrunc x <> 0%R -> mag beta (round Ztrunc x) = mag beta x \/ Rabs (round Ztrunc x) = bpow (Z.max (mag beta x) (fexp (mag beta x)))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
round Zaway x <> 0%R -> mag beta (round Zaway x) = mag beta x \/ Rabs (round Zaway x) = bpow (Z.max (mag beta x) (fexp (mag beta x)))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
H:round Ztrunc x <> 0%R

mag beta (round Ztrunc x) = mag beta x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
round Zaway x <> 0%R -> mag beta (round Zaway x) = mag beta x \/ Rabs (round Zaway x) = bpow (Z.max (mag beta x) (fexp (mag beta x)))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

round Zaway x <> 0%R -> mag beta (round Zaway x) = mag beta x \/ Rabs (round Zaway x) = bpow (Z.max (mag beta x) (fexp (mag beta x)))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zr:round Zaway x <> 0%R

mag beta (round Zaway x) = mag beta x \/ Rabs (round Zaway x) = bpow (Z.max (mag beta x) (fexp (mag beta x)))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zr:round Zaway x <> 0%R
Zx:x = 0%R

mag beta (round Zaway x) = mag beta x \/ Rabs (round Zaway x) = bpow (Z.max (mag beta x) (fexp (mag beta x)))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zr:round Zaway x <> 0%R
Zx:x <> 0%R
mag beta (round Zaway x) = mag beta x \/ Rabs (round Zaway x) = bpow (Z.max (mag beta x) (fexp (mag beta x)))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zr:round Zaway x <> 0%R
Zx:x <> 0%R

mag beta (round Zaway x) = mag beta x \/ Rabs (round Zaway x) = bpow (Z.max (mag beta x) (fexp (mag beta x)))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zr:round Zaway x <> 0%R
Zx:x <> 0%R
ex:Z
Ex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R

mag beta (round Zaway x) = ex \/ Rabs (round Zaway x) = bpow (Z.max ex (fexp ex))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zr:round Zaway x <> 0%R
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R

mag beta (round Zaway x) = ex \/ Rabs (round Zaway x) = bpow (Z.max ex (fexp ex))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zr:round Zaway x <> 0%R
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R

mag beta (Rabs (round Zaway x)) = ex \/ Rabs (round Zaway x) = bpow (Z.max ex (fexp ex))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zr:round Zaway x <> 0%R
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R

mag beta (round Zaway (Rabs x)) = ex \/ round Zaway (Rabs x) = bpow (Z.max ex (fexp ex))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zr:round Zaway x <> 0%R
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
He:(ex <= fexp ex)%Z

mag beta (round Zaway (Rabs x)) = ex \/ round Zaway (Rabs x) = bpow (Z.max ex (fexp ex))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zr:round Zaway x <> 0%R
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
He:(fexp ex < ex)%Z
mag beta (round Zaway (Rabs x)) = ex \/ round Zaway (Rabs x) = bpow (Z.max ex (fexp ex))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zr:round Zaway x <> 0%R
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
He:(ex <= fexp ex)%Z

round Zaway (Rabs x) = bpow (Z.max ex (fexp ex))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zr:round Zaway x <> 0%R
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
He:(fexp ex < ex)%Z
mag beta (round Zaway (Rabs x)) = ex \/ round Zaway (Rabs x) = bpow (Z.max ex (fexp ex))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zr:round Zaway x <> 0%R
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
He:(ex <= fexp ex)%Z

round Zaway (Rabs x) = bpow (fexp ex)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zr:round Zaway x <> 0%R
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
He:(fexp ex < ex)%Z
mag beta (round Zaway (Rabs x)) = ex \/ round Zaway (Rabs x) = bpow (Z.max ex (fexp ex))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zr:round Zaway x <> 0%R
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
He:(ex <= fexp ex)%Z

round Zceil (Rabs x) = bpow (fexp ex)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zr:round Zaway x <> 0%R
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
He:(fexp ex < ex)%Z
mag beta (round Zaway (Rabs x)) = ex \/ round Zaway (Rabs x) = bpow (Z.max ex (fexp ex))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zr:round Zaway x <> 0%R
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
He:(fexp ex < ex)%Z

mag beta (round Zaway (Rabs x)) = ex \/ round Zaway (Rabs x) = bpow (Z.max ex (fexp ex))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zr:round Zaway x <> 0%R
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
He:(fexp ex < ex)%Z
H1:(bpow (ex - 1) <= round Zaway (Rabs x))%R
H2:(round Zaway (Rabs x) < bpow ex)%R

mag beta (round Zaway (Rabs x)) = ex \/ round Zaway (Rabs x) = bpow (Z.max ex (fexp ex))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zr:round Zaway x <> 0%R
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
He:(fexp ex < ex)%Z
H1:(bpow (ex - 1) <= round Zaway (Rabs x))%R
H2:round Zaway (Rabs x) = bpow ex
mag beta (round Zaway (Rabs x)) = ex \/ round Zaway (Rabs x) = bpow (Z.max ex (fexp ex))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zr:round Zaway x <> 0%R
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
He:(fexp ex < ex)%Z
H1:(bpow (ex - 1) <= round Zaway (Rabs x))%R
H2:(round Zaway (Rabs x) < bpow ex)%R

mag beta (round Zaway (Rabs x)) = ex
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zr:round Zaway x <> 0%R
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
He:(fexp ex < ex)%Z
H1:(bpow (ex - 1) <= round Zaway (Rabs x))%R
H2:round Zaway (Rabs x) = bpow ex
mag beta (round Zaway (Rabs x)) = ex \/ round Zaway (Rabs x) = bpow (Z.max ex (fexp ex))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zr:round Zaway x <> 0%R
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
He:(fexp ex < ex)%Z
H1:(bpow (ex - 1) <= round Zaway (Rabs x))%R
H2:(round Zaway (Rabs x) < bpow ex)%R

(bpow (ex - 1) <= Rabs (round Zaway (Rabs x)) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zr:round Zaway x <> 0%R
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
He:(fexp ex < ex)%Z
H1:(bpow (ex - 1) <= round Zaway (Rabs x))%R
H2:round Zaway (Rabs x) = bpow ex
mag beta (round Zaway (Rabs x)) = ex \/ round Zaway (Rabs x) = bpow (Z.max ex (fexp ex))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zr:round Zaway x <> 0%R
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
He:(fexp ex < ex)%Z
H1:(bpow (ex - 1) <= round Zaway (Rabs x))%R
H2:(round Zaway (Rabs x) < bpow ex)%R

(bpow (ex - 1) <= round Zaway (Rabs x) < bpow ex)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zr:round Zaway x <> 0%R
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
He:(fexp ex < ex)%Z
H1:(bpow (ex - 1) <= round Zaway (Rabs x))%R
H2:round Zaway (Rabs x) = bpow ex
mag beta (round Zaway (Rabs x)) = ex \/ round Zaway (Rabs x) = bpow (Z.max ex (fexp ex))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zr:round Zaway x <> 0%R
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
He:(fexp ex < ex)%Z
H1:(bpow (ex - 1) <= round Zaway (Rabs x))%R
H2:round Zaway (Rabs x) = bpow ex

mag beta (round Zaway (Rabs x)) = ex \/ round Zaway (Rabs x) = bpow (Z.max ex (fexp ex))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Zr:round Zaway x <> 0%R
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
He:(fexp ex < ex)%Z
H1:(bpow (ex - 1) <= round Zaway (Rabs x))%R
H2:round Zaway (Rabs x) = bpow ex

round Zaway (Rabs x) = bpow (Z.max ex (fexp ex))
now rewrite Z.max_l with (1 := Zlt_le_weak _ _ He). Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, (0 < round Zfloor x)%R -> mag beta (round Zfloor x) = mag beta x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, (0 < round Zfloor x)%R -> mag beta (round Zfloor x) = mag beta x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hd:(0 < round Zfloor x)%R

mag beta (round Zfloor x) = mag beta x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hd:(0 < round Zfloor x)%R

(0 < x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hd:(0 < round Zfloor x)%R
H:(0 < x)%R
mag beta (round Zfloor x) = mag beta x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hd:(0 < round Zfloor x)%R

(round Zfloor x <= x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hd:(0 < round Zfloor x)%R
H:(0 < x)%R
mag beta (round Zfloor x) = mag beta x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hd:(0 < round Zfloor x)%R
H:(0 < x)%R

mag beta (round Zfloor x) = mag beta x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
H:(0 < x)%R

(0 < round Zfloor x)%R -> mag beta (round Zfloor x) = mag beta x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
H:(0 < x)%R

(0 < round Ztrunc x)%R -> mag beta (round Ztrunc x) = mag beta x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
H:(0 < x)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
H:(0 < x)%R
Hd:(0 < round Ztrunc x)%R

mag beta (round Ztrunc x) = mag beta x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
H:(0 < x)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
H:(0 < x)%R
Hd:(0 < round Ztrunc x)%R

round Ztrunc x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
H:(0 < x)%R
(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
H:(0 < x)%R

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

forall x : R, (0 < round Zfloor x)%R -> cexp (round Zfloor x) = cexp x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, (0 < round Zfloor x)%R -> cexp (round Zfloor x) = cexp x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hd:(0 < round Zfloor x)%R

cexp (round Zfloor x) = cexp x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hd:(0 < round Zfloor x)%R

mag beta (round Zfloor x) = mag beta x
now apply mag_DN. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, (0 < round Zfloor x)%R -> scaled_mantissa (round Zfloor x) = IZR (Zfloor (scaled_mantissa x))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, (0 < round Zfloor x)%R -> scaled_mantissa (round Zfloor x) = IZR (Zfloor (scaled_mantissa x))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hd:(0 < round Zfloor x)%R

scaled_mantissa (round Zfloor x) = IZR (Zfloor (scaled_mantissa x))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hd:(0 < round Zfloor x)%R

(round Zfloor x * bpow (- cexp (round Zfloor x)))%R = IZR (Zfloor (x * bpow (- cexp x)))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hd:(0 < round Zfloor x)%R

(round Zfloor x * bpow (- cexp x))%R = IZR (Zfloor (x * bpow (- cexp x)))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hd:(0 < round Zfloor x)%R

(IZR (Fnum {| Fnum := Zfloor (scaled_mantissa x); Fexp := cexp x |}) * bpow (Fexp {| Fnum := Zfloor (scaled_mantissa x); Fexp := cexp x |}) * bpow (- cexp x))%R = IZR (Zfloor (x * bpow (- cexp x)))
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
Hd:(0 < round Zfloor x)%R

(IZR (Zfloor (scaled_mantissa x)) * bpow (cexp x) * bpow (- cexp x))%R = IZR (Zfloor (x * bpow (- cexp x)))
now rewrite Rmult_assoc, <- bpow_plus, Zplus_opp_r, Rmult_1_r. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x f : R, Rnd_N_pt generic_format x f -> f = round Zfloor x \/ f = round Zceil x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x f : R, Rnd_N_pt generic_format x f -> f = round Zfloor x \/ f = round Zceil x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x, f:R
Hxf:Rnd_N_pt generic_format x f

f = round Zfloor x \/ f = round Zceil x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x, f:R
Hxf:Rnd_N_pt generic_format x f
H:Rnd_DN_pt generic_format x f

f = round Zfloor x \/ f = round Zceil x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x, f:R
Hxf:Rnd_N_pt generic_format x f
H:Rnd_UP_pt generic_format x f
f = round Zfloor x \/ f = round Zceil x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x, f:R
Hxf:Rnd_N_pt generic_format x f
H:Rnd_DN_pt generic_format x f

f = round Zfloor x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x, f:R
Hxf:Rnd_N_pt generic_format x f
H:Rnd_UP_pt generic_format x f
f = round Zfloor x \/ f = round Zceil x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x, f:R
Hxf:Rnd_N_pt generic_format x f
H:Rnd_DN_pt generic_format x f

Rnd_DN_pt generic_format x (round Zfloor x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x, f:R
Hxf:Rnd_N_pt generic_format x f
H:Rnd_UP_pt generic_format x f
f = round Zfloor x \/ f = round Zceil x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x, f:R
Hxf:Rnd_N_pt generic_format x f
H:Rnd_UP_pt generic_format x f

f = round Zfloor x \/ f = round Zceil x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x, f:R
Hxf:Rnd_N_pt generic_format x f
H:Rnd_UP_pt generic_format x f

f = round Zceil x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x, f:R
Hxf:Rnd_N_pt generic_format x f
H:Rnd_UP_pt generic_format x f

Rnd_UP_pt generic_format x (round Zceil x)
apply round_UP_pt. Qed. Section not_FTZ. Class Exp_not_FTZ := exp_not_FTZ : forall e, (fexp (fexp e + 1) <= fexp e)%Z. Context { exp_not_FTZ_ : Exp_not_FTZ }.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
exp_not_FTZ_:Exp_not_FTZ

forall (e : Z) (x : R), (e <= fexp e)%Z -> generic_format x -> x = F2R {| Fnum := Ztrunc (x * bpow (- fexp e)); Fexp := fexp e |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
exp_not_FTZ_:Exp_not_FTZ

forall (e : Z) (x : R), (e <= fexp e)%Z -> generic_format x -> x = F2R {| Fnum := Ztrunc (x * bpow (- fexp e)); Fexp := fexp e |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
exp_not_FTZ_:Exp_not_FTZ
e:Z
x:R
He:(e <= fexp e)%Z
Hx:generic_format x

x = F2R {| Fnum := Ztrunc (x * bpow (- fexp e)); Fexp := fexp e |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
exp_not_FTZ_:Exp_not_FTZ
e:Z
x:R
He:(e <= fexp e)%Z
Hx:generic_format x

x = F2R {| Fnum := Ztrunc (F2R {| Fnum := Ztrunc (scaled_mantissa x); Fexp := cexp x |} * bpow (- fexp e)); Fexp := fexp e |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
exp_not_FTZ_:Exp_not_FTZ
e:Z
x:R
He:(e <= fexp e)%Z
Hx:generic_format x

x = F2R {| Fnum := Ztrunc (IZR (Fnum {| Fnum := Ztrunc (scaled_mantissa x); Fexp := cexp x |}) * bpow (Fexp {| Fnum := Ztrunc (scaled_mantissa x); Fexp := cexp x |}) * bpow (- fexp e)); Fexp := fexp e |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
exp_not_FTZ_:Exp_not_FTZ
e:Z
x:R
He:(e <= fexp e)%Z
Hx:generic_format x

x = F2R {| Fnum := Ztrunc (IZR (Ztrunc (scaled_mantissa x)) * bpow (cexp x) * bpow (- fexp e)); Fexp := fexp e |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
exp_not_FTZ_:Exp_not_FTZ
e:Z
x:R
He:(e <= fexp e)%Z
Hx:generic_format x

x = F2R {| Fnum := Ztrunc (IZR (Ztrunc (scaled_mantissa x)) * bpow (cexp x + - fexp e)); Fexp := fexp e |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
exp_not_FTZ_:Exp_not_FTZ
e:Z
x:R
He:(e <= fexp e)%Z
Hx:generic_format x

IZR (beta ^ (cexp x + - fexp e)) = bpow (cexp x + - fexp e)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
exp_not_FTZ_:Exp_not_FTZ
e:Z
x:R
He:(e <= fexp e)%Z
Hx:generic_format x
H:IZR (beta ^ (cexp x + - fexp e)) = bpow (cexp x + - fexp e)
x = F2R {| Fnum := Ztrunc (IZR (Ztrunc (scaled_mantissa x)) * bpow (cexp x + - fexp e)); Fexp := fexp e |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
exp_not_FTZ_:Exp_not_FTZ
e:Z
x:R
He:(e <= fexp e)%Z
Hx:generic_format x

(0 <= cexp x + - fexp e)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
exp_not_FTZ_:Exp_not_FTZ
e:Z
x:R
He:(e <= fexp e)%Z
Hx:generic_format x
H:IZR (beta ^ (cexp x + - fexp e)) = bpow (cexp x + - fexp e)
x = F2R {| Fnum := Ztrunc (IZR (Ztrunc (scaled_mantissa x)) * bpow (cexp x + - fexp e)); Fexp := fexp e |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
exp_not_FTZ_:Exp_not_FTZ
e:Z
x:R
He:(e <= fexp e)%Z
Hx:generic_format x

(0 <= fexp (mag beta x) + - fexp e)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
exp_not_FTZ_:Exp_not_FTZ
e:Z
x:R
He:(e <= fexp e)%Z
Hx:generic_format x
H:IZR (beta ^ (cexp x + - fexp e)) = bpow (cexp x + - fexp e)
x = F2R {| Fnum := Ztrunc (IZR (Ztrunc (scaled_mantissa x)) * bpow (cexp x + - fexp e)); Fexp := fexp e |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
exp_not_FTZ_:Exp_not_FTZ
e:Z
x:R
He:(e <= fexp e)%Z
Hx:generic_format x
ex:=mag beta x:mag_prop beta x

(0 <= fexp ex + - fexp e)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
exp_not_FTZ_:Exp_not_FTZ
e:Z
x:R
He:(e <= fexp e)%Z
Hx:generic_format x
H:IZR (beta ^ (cexp x + - fexp e)) = bpow (cexp x + - fexp e)
x = F2R {| Fnum := Ztrunc (IZR (Ztrunc (scaled_mantissa x)) * bpow (cexp x + - fexp e)); Fexp := fexp e |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
exp_not_FTZ_:Exp_not_FTZ
e:Z
x:R
He:(e <= fexp e)%Z
Hx:generic_format x
ex:=mag beta x:mag_prop beta x

(fexp (fexp ex + 1) <= fexp ex)%Z -> (0 <= fexp ex + - fexp e)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
exp_not_FTZ_:Exp_not_FTZ
e:Z
x:R
He:(e <= fexp e)%Z
Hx:generic_format x
H:IZR (beta ^ (cexp x + - fexp e)) = bpow (cexp x + - fexp e)
x = F2R {| Fnum := Ztrunc (IZR (Ztrunc (scaled_mantissa x)) * bpow (cexp x + - fexp e)); Fexp := fexp e |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
exp_not_FTZ_:Exp_not_FTZ
e:Z
x:R
He:(e <= fexp e)%Z
Hx:generic_format x
ex:=mag beta x:mag_prop beta x

((fexp ex + 1 <= fexp e)%Z -> fexp (fexp ex + 1) = fexp e) -> (fexp (fexp ex + 1) <= fexp ex)%Z -> (0 <= fexp ex + - fexp e)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
exp_not_FTZ_:Exp_not_FTZ
e:Z
x:R
He:(e <= fexp e)%Z
Hx:generic_format x
H:IZR (beta ^ (cexp x + - fexp e)) = bpow (cexp x + - fexp e)
x = F2R {| Fnum := Ztrunc (IZR (Ztrunc (scaled_mantissa x)) * bpow (cexp x + - fexp e)); Fexp := fexp e |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
exp_not_FTZ_:Exp_not_FTZ
e:Z
x:R
He:(e <= fexp e)%Z
Hx:generic_format x
H:IZR (beta ^ (cexp x + - fexp e)) = bpow (cexp x + - fexp e)

x = F2R {| Fnum := Ztrunc (IZR (Ztrunc (scaled_mantissa x)) * bpow (cexp x + - fexp e)); Fexp := fexp e |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
exp_not_FTZ_:Exp_not_FTZ
e:Z
x:R
He:(e <= fexp e)%Z
Hx:generic_format x
H:IZR (beta ^ (cexp x + - fexp e)) = bpow (cexp x + - fexp e)

x = F2R {| Fnum := Ztrunc (IZR (Ztrunc (scaled_mantissa x)) * IZR (beta ^ (cexp x + - fexp e))); Fexp := fexp e |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
exp_not_FTZ_:Exp_not_FTZ
e:Z
x:R
He:(e <= fexp e)%Z
Hx:generic_format x
H:IZR (beta ^ (cexp x + - fexp e)) = bpow (cexp x + - fexp e)

x = F2R {| Fnum := Ztrunc (scaled_mantissa x) * beta ^ (cexp x + - fexp e); Fexp := fexp e |}
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
exp_not_FTZ_:Exp_not_FTZ
e:Z
x:R
He:(e <= fexp e)%Z
Hx:generic_format x
H:IZR (beta ^ (cexp x + - fexp e)) = bpow (cexp x + - fexp e)

x = (IZR (Fnum {| Fnum := Ztrunc (scaled_mantissa x) * beta ^ (cexp x + - fexp e); Fexp := fexp e |}) * bpow (Fexp {| Fnum := Ztrunc (scaled_mantissa x) * beta ^ (cexp x + - fexp e); Fexp := fexp e |}))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
exp_not_FTZ_:Exp_not_FTZ
e:Z
x:R
He:(e <= fexp e)%Z
Hx:generic_format x
H:IZR (beta ^ (cexp x + - fexp e)) = bpow (cexp x + - fexp e)

x = (IZR (Ztrunc (scaled_mantissa x) * beta ^ (cexp x + - fexp e)) * bpow (fexp e))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
exp_not_FTZ_:Exp_not_FTZ
e:Z
x:R
He:(e <= fexp e)%Z
Hx:generic_format x
H:IZR (beta ^ (cexp x + - fexp e)) = bpow (cexp x + - fexp e)

x = (IZR (Ztrunc (scaled_mantissa x)) * IZR (beta ^ (cexp x + - fexp e)) * bpow (fexp e))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
exp_not_FTZ_:Exp_not_FTZ
e:Z
x:R
He:(e <= fexp e)%Z
Hx:generic_format x
H:IZR (beta ^ (cexp x + - fexp e)) = bpow (cexp x + - fexp e)

x = (IZR (Ztrunc (scaled_mantissa x)) * bpow (cexp x + - fexp e) * bpow (fexp e))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
exp_not_FTZ_:Exp_not_FTZ
e:Z
x:R
He:(e <= fexp e)%Z
Hx:generic_format x
H:IZR (beta ^ (cexp x + - fexp e)) = bpow (cexp x + - fexp e)

x = (IZR (Ztrunc (scaled_mantissa x)) * bpow (cexp x + - fexp e + fexp e))%R
now ring_simplify (cexp x + - fexp e + fexp e)%Z. Qed. End not_FTZ. Section monotone_exp. Class Monotone_exp := monotone_exp : forall ex ey, (ex <= ey)%Z -> (fexp ex <= fexp ey)%Z. Context { monotone_exp_ : Monotone_exp }.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
monotone_exp_:Monotone_exp

Exp_not_FTZ
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
monotone_exp_:Monotone_exp

Exp_not_FTZ
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
monotone_exp_:Monotone_exp
e:Z

(fexp (fexp e + 1) <= fexp e)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
monotone_exp_:Monotone_exp
e:Z
He:(fexp e < e)%Z

(fexp (fexp e + 1) <= fexp e)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
monotone_exp_:Monotone_exp
e:Z
He:(e <= fexp e)%Z
(fexp (fexp e + 1) <= fexp e)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
monotone_exp_:Monotone_exp
e:Z
He:(fexp e < e)%Z

(fexp e + 1 <= e)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
monotone_exp_:Monotone_exp
e:Z
He:(e <= fexp e)%Z
(fexp (fexp e + 1) <= fexp e)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
monotone_exp_:Monotone_exp
e:Z
He:(e <= fexp e)%Z

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

forall (x : R) (e : Z), x <> 0%R -> (Rabs x < bpow e)%R -> (cexp x <= fexp e)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
monotone_exp_:Monotone_exp

forall (x : R) (e : Z), x <> 0%R -> (Rabs x < bpow e)%R -> (cexp x <= fexp e)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
monotone_exp_:Monotone_exp
x:R
e:Z
Zx:x <> 0%R
Hx:(Rabs x < bpow e)%R

(cexp x <= fexp e)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
monotone_exp_:Monotone_exp
x:R
e:Z
Zx:x <> 0%R
Hx:(Rabs x < bpow e)%R

(mag beta x <= e)%Z
now apply mag_le_bpow. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
monotone_exp_:Monotone_exp

forall (x : R) (e : Z), (bpow (e - 1) <= Rabs x)%R -> (fexp e <= cexp x)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
monotone_exp_:Monotone_exp

forall (x : R) (e : Z), (bpow (e - 1) <= Rabs x)%R -> (fexp e <= cexp x)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
monotone_exp_:Monotone_exp
x:R
e:Z
Hx:(bpow (e - 1) <= Rabs x)%R

(fexp e <= cexp x)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
monotone_exp_:Monotone_exp
x:R
e:Z
Hx:(bpow (e - 1) <= Rabs x)%R

(e <= mag beta x)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
monotone_exp_:Monotone_exp
x:R
e:Z
Hx:(bpow (e - 1) <= Rabs x)%R

(Z.succ (Z.pred e) <= mag beta x)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
monotone_exp_:Monotone_exp
x:R
e:Z
Hx:(bpow (e - 1) <= Rabs x)%R

(Z.pred e < mag beta x)%Z
now apply mag_gt_bpow. Qed. Variable rnd : R -> Z. Context { valid_rnd : Valid_rnd rnd }.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
monotone_exp_:Monotone_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd

forall x : R, round rnd x <> 0%R -> (mag beta x <= mag beta (round rnd x))%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
monotone_exp_:Monotone_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd

forall x : R, round rnd x <> 0%R -> (mag beta x <= mag beta (round rnd x))%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
monotone_exp_:Monotone_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R

round rnd x <> 0%R -> (mag beta x <= mag beta (round rnd x))%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
monotone_exp_:Monotone_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Zr:round Ztrunc x <> 0%R

(mag beta x <= mag beta (round Ztrunc x))%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
monotone_exp_:Monotone_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Zr:round Zaway x <> 0%R
(mag beta x <= mag beta (round Zaway x))%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
monotone_exp_:Monotone_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Zr:round Ztrunc x <> 0%R

(mag beta x <= mag beta x)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
monotone_exp_:Monotone_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Zr:round Zaway x <> 0%R
(mag beta x <= mag beta (round Zaway x))%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
monotone_exp_:Monotone_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Zr:round Zaway x <> 0%R

(mag beta x <= mag beta (round Zaway x))%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
monotone_exp_:Monotone_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Zr:round Zaway x <> 0%R

x <> 0%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
monotone_exp_:Monotone_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Zr:round Zaway x <> 0%R
(Rabs x <= Rabs (round Zaway x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
monotone_exp_:Monotone_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Zr:x = 0%R

round Zaway x = 0%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
monotone_exp_:Monotone_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Zr:round Zaway x <> 0%R
(Rabs x <= Rabs (round Zaway x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
monotone_exp_:Monotone_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Zr:x = 0%R

round Zaway 0 = 0%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
monotone_exp_:Monotone_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Zr:round Zaway x <> 0%R
(Rabs x <= Rabs (round Zaway x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
monotone_exp_:Monotone_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Zr:round Zaway x <> 0%R

(Rabs x <= Rabs (round Zaway x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
monotone_exp_:Monotone_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Zr:round Zaway x <> 0%R

(Rabs x <= round Zaway (Rabs x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
monotone_exp_:Monotone_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Zr:round Zaway x <> 0%R

(Rabs x <= round Zceil (Rabs x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
monotone_exp_:Monotone_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Zr:round Zaway x <> 0%R
(0 <= Rabs x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
monotone_exp_:Monotone_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Zr:round Zaway x <> 0%R

(0 <= Rabs x)%R
apply Rabs_pos. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
monotone_exp_:Monotone_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd

forall x : R, round rnd x <> 0%R -> (cexp x <= cexp (round rnd x))%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
monotone_exp_:Monotone_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd

forall x : R, round rnd x <> 0%R -> (cexp x <= cexp (round rnd x))%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
monotone_exp_:Monotone_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Zr:round rnd x <> 0%R

(cexp x <= cexp (round rnd x))%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
monotone_exp_:Monotone_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Zr:round rnd x <> 0%R

(fexp (mag beta x) <= fexp (mag beta (round rnd x)))%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
monotone_exp_:Monotone_exp
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Zr:round rnd x <> 0%R

(mag beta x <= mag beta (round rnd x))%Z
now apply mag_round_ge. Qed. End monotone_exp. Section Znearest.
Roundings to nearest: when in the middle, use the choice function
Variable choice : Z -> bool.

Definition Znearest x :=
  match Rcompare (x - IZR (Zfloor x)) (/2) with
  | Lt => Zfloor x
  | Eq => if choice (Zfloor x) then Zceil x else Zfloor x
  | Gt => Zceil x
  end.

beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool

forall x : R, Znearest x = Zfloor x \/ Znearest x = Zceil x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool

forall x : R, Znearest x = Zfloor x \/ Znearest x = Zceil x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R

Znearest x = Zfloor x \/ Znearest x = Zceil x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R

match Rcompare (x - IZR (Zfloor x)) (/ 2) with | Eq => if choice (Zfloor x) then Zceil x else Zfloor x | Lt => Zfloor x | Gt => Zceil x end = Zfloor x \/ match Rcompare (x - IZR (Zfloor x)) (/ 2) with | Eq => if choice (Zfloor x) then Zceil x else Zfloor x | Lt => Zfloor x | Gt => Zceil x end = Zceil x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R

Zfloor x = Zfloor x \/ Zfloor x = Zceil x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
(if choice (Zfloor x) then Zceil x else Zfloor x) = Zfloor x \/ (if choice (Zfloor x) then Zceil x else Zfloor x) = Zceil x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Zceil x = Zfloor x \/ Zceil x = Zceil x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R

(if choice (Zfloor x) then Zceil x else Zfloor x) = Zfloor x \/ (if choice (Zfloor x) then Zceil x else Zfloor x) = Zceil x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Zceil x = Zfloor x \/ Zceil x = Zceil x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R

Zceil x = Zfloor x \/ Zceil x = Zceil x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Zfloor x = Zfloor x \/ Zfloor x = Zceil x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Zceil x = Zfloor x \/ Zceil x = Zceil x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R

Zfloor x = Zfloor x \/ Zfloor x = Zceil x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Zceil x = Zfloor x \/ Zceil x = Zceil x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R

Zceil x = Zfloor x \/ Zceil x = Zceil x
now right. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool

forall x : R, (Zfloor x <= Znearest x)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool

forall x : R, (Zfloor x <= Znearest x)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R

(Zfloor x <= Znearest x)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:Znearest x = Zfloor x

(Zfloor x <= Zfloor x)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:Znearest x = Zceil x
(Zfloor x <= Zceil x)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:Znearest x = Zceil x

(Zfloor x <= Zceil x)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:Znearest x = Zceil x

(IZR (Zfloor x) <= IZR (Zceil x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:Znearest x = Zceil x

(IZR (Zfloor x) <= x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:Znearest x = Zceil x
(x <= IZR (Zceil x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:Znearest x = Zceil x

(x <= IZR (Zceil x))%R
apply Zceil_ub. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool

forall x : R, (Znearest x <= Zceil x)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool

forall x : R, (Znearest x <= Zceil x)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R

(Znearest x <= Zceil x)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:Znearest x = Zfloor x

(Zfloor x <= Zceil x)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:Znearest x = Zceil x
(Zceil x <= Zceil x)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:Znearest x = Zfloor x

(IZR (Zfloor x) <= IZR (Zceil x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:Znearest x = Zceil x
(Zceil x <= Zceil x)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:Znearest x = Zfloor x

(IZR (Zfloor x) <= x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:Znearest x = Zfloor x
(x <= IZR (Zceil x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:Znearest x = Zceil x
(Zceil x <= Zceil x)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:Znearest x = Zfloor x

(x <= IZR (Zceil x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:Znearest x = Zceil x
(Zceil x <= Zceil x)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:Znearest x = Zceil x

(Zceil x <= Zceil x)%Z
apply Z.le_refl. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool

Valid_rnd Znearest
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool

Valid_rnd Znearest
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool

forall x y : R, (x <= y)%R -> (Znearest x <= Znearest y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
forall n : Z, Znearest (IZR n) = n
(* *)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R

(Znearest x <= Znearest y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
forall n : Z, Znearest (IZR n) = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(IZR (Zceil x) <= y)%R

(Znearest x <= Znearest y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
(Znearest x <= Znearest y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
forall n : Z, Znearest (IZR n) = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(IZR (Zceil x) <= y)%R

(Zceil x <= Znearest y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
(Znearest x <= Znearest y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
forall n : Z, Znearest (IZR n) = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(IZR (Zceil x) <= y)%R

(Zceil x <= Zfloor y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
(Znearest x <= Znearest y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
forall n : Z, Znearest (IZR n) = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R

(Znearest x <= Znearest y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
forall n : Z, Znearest (IZR n) = n
(* . *)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R

Zfloor y = Zfloor x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
(Znearest x <= Znearest y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
forall n : Z, Znearest (IZR n) = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R

(IZR (Zfloor x) <= y < IZR (Zfloor x + 1))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
(Znearest x <= Znearest y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
forall n : Z, Znearest (IZR n) = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R

(IZR (Zfloor x) <= y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
(y < IZR (Zfloor x + 1))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
(Znearest x <= Znearest y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
forall n : Z, Znearest (IZR n) = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R

(IZR (Zfloor x) <= IZR (Zfloor y))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
(y < IZR (Zfloor x + 1))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
(Znearest x <= Znearest y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
forall n : Z, Znearest (IZR n) = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R

(Zfloor x <= Zfloor y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
(y < IZR (Zfloor x + 1))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
(Znearest x <= Znearest y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
forall n : Z, Znearest (IZR n) = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R

(y < IZR (Zfloor x + 1))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
(Znearest x <= Znearest y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
forall n : Z, Znearest (IZR n) = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R

(IZR (Zceil x) <= IZR (Zfloor x + 1))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
(Znearest x <= Znearest y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
forall n : Z, Znearest (IZR n) = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R

(Zceil x <= Zfloor x + 1)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
(Znearest x <= Znearest y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
forall n : Z, Znearest (IZR n) = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R

(x <= IZR (Zfloor x + 1))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
(Znearest x <= Znearest y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
forall n : Z, Znearest (IZR n) = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R

(x < IZR (Zfloor x + 1))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
(Znearest x <= Znearest y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
forall n : Z, Znearest (IZR n) = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R

(x < IZR (Zfloor x) + 1)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
(Znearest x <= Znearest y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
forall n : Z, Znearest (IZR n) = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x

(Znearest x <= Znearest y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
forall n : Z, Znearest (IZR n) = n
(* . *)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x

(match Rcompare (x - IZR (Zfloor x)) (/ 2) with | Eq => if choice (Zfloor x) then Zceil x else Zfloor x | Lt => Zfloor x | Gt => Zceil x end <= Znearest y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
forall n : Z, Znearest (IZR n) = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(x - IZR (Zfloor x) < / 2)%R

(Zfloor x <= Znearest y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
((if choice (Zfloor x) then Zceil x else Zfloor x) <= Znearest y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(/ 2 < x - IZR (Zfloor x))%R
(Zceil x <= Znearest y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
forall n : Z, Znearest (IZR n) = n
(* .. *)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(x - IZR (Zfloor x) < / 2)%R

(Zfloor y <= Znearest y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
((if choice (Zfloor x) then Zceil x else Zfloor x) <= Znearest y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(/ 2 < x - IZR (Zfloor x))%R
(Zceil x <= Znearest y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
forall n : Z, Znearest (IZR n) = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R

((if choice (Zfloor x) then Zceil x else Zfloor x) <= Znearest y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(/ 2 < x - IZR (Zfloor x))%R
(Zceil x <= Znearest y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
forall n : Z, Znearest (IZR n) = n
(* .. *)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R

((if choice (Zfloor x) then Zceil x else Zfloor x) <= match Rcompare (y - IZR (Zfloor y)) (/ 2) with | Eq => if choice (Zfloor y) then Zceil y else Zfloor y | Lt => Zfloor y | Gt => Zceil y end)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(/ 2 < x - IZR (Zfloor x))%R
(Zceil x <= Znearest y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
forall n : Z, Znearest (IZR n) = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R

((if choice (Zfloor x) then Zceil x else Zfloor x) <= match Rcompare (y - IZR (Zfloor x)) (/ 2) with | Eq => if choice (Zfloor x) then Zceil y else Zfloor x | Lt => Zfloor x | Gt => Zceil y end)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(/ 2 < x - IZR (Zfloor x))%R
(Zceil x <= Znearest y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
forall n : Z, Znearest (IZR n) = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
Hy:(y - IZR (Zfloor x) < / 2)%R

((if choice (Zfloor x) then Zceil x else Zfloor x) <= Zfloor x)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
Hy:(y - IZR (Zfloor x))%R = (/ 2)%R
((if choice (Zfloor x) then Zceil x else Zfloor x) <= (if choice (Zfloor x) then Zceil y else Zfloor x))%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
Hy:(/ 2 < y - IZR (Zfloor x))%R
((if choice (Zfloor x) then Zceil x else Zfloor x) <= Zceil y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(/ 2 < x - IZR (Zfloor x))%R
(Zceil x <= Znearest y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
forall n : Z, Znearest (IZR n) = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
Hy:(y - IZR (Zfloor x) < / 2)%R

(/ 2 <= y - IZR (Zfloor x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
Hy:(y - IZR (Zfloor x))%R = (/ 2)%R
((if choice (Zfloor x) then Zceil x else Zfloor x) <= (if choice (Zfloor x) then Zceil y else Zfloor x))%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
Hy:(/ 2 < y - IZR (Zfloor x))%R
((if choice (Zfloor x) then Zceil x else Zfloor x) <= Zceil y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(/ 2 < x - IZR (Zfloor x))%R
(Zceil x <= Znearest y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
forall n : Z, Znearest (IZR n) = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
Hy:(y - IZR (Zfloor x) < / 2)%R

(x - IZR (Zfloor x) <= y - IZR (Zfloor x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
Hy:(y - IZR (Zfloor x))%R = (/ 2)%R
((if choice (Zfloor x) then Zceil x else Zfloor x) <= (if choice (Zfloor x) then Zceil y else Zfloor x))%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
Hy:(/ 2 < y - IZR (Zfloor x))%R
((if choice (Zfloor x) then Zceil x else Zfloor x) <= Zceil y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(/ 2 < x - IZR (Zfloor x))%R
(Zceil x <= Znearest y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
forall n : Z, Znearest (IZR n) = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
Hy:(y - IZR (Zfloor x))%R = (/ 2)%R

((if choice (Zfloor x) then Zceil x else Zfloor x) <= (if choice (Zfloor x) then Zceil y else Zfloor x))%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
Hy:(/ 2 < y - IZR (Zfloor x))%R
((if choice (Zfloor x) then Zceil x else Zfloor x) <= Zceil y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(/ 2 < x - IZR (Zfloor x))%R
(Zceil x <= Znearest y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
forall n : Z, Znearest (IZR n) = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
Hy:(y - IZR (Zfloor x))%R = (/ 2)%R

((if choice (Zfloor x) then Zceil x else Zfloor x) <= (if choice (Zfloor x) then Zceil x else Zfloor x))%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
Hy:(y - IZR (Zfloor x))%R = (/ 2)%R
x = y
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
Hy:(/ 2 < y - IZR (Zfloor x))%R
((if choice (Zfloor x) then Zceil x else Zfloor x) <= Zceil y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(/ 2 < x - IZR (Zfloor x))%R
(Zceil x <= Znearest y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
forall n : Z, Znearest (IZR n) = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
Hy:(y - IZR (Zfloor x))%R = (/ 2)%R

x = y
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
Hy:(/ 2 < y - IZR (Zfloor x))%R
((if choice (Zfloor x) then Zceil x else Zfloor x) <= Zceil y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(/ 2 < x - IZR (Zfloor x))%R
(Zceil x <= Znearest y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
forall n : Z, Znearest (IZR n) = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
Hy:(y - IZR (Zfloor x))%R = (/ 2)%R

(- IZR (Zfloor x) + x)%R = (- IZR (Zfloor x) + y)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
Hy:(/ 2 < y - IZR (Zfloor x))%R
((if choice (Zfloor x) then Zceil x else Zfloor x) <= Zceil y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(/ 2 < x - IZR (Zfloor x))%R
(Zceil x <= Znearest y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
forall n : Z, Znearest (IZR n) = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
Hy:(y - IZR (Zfloor x))%R = (/ 2)%R

(x + - IZR (Zfloor x))%R = (y + - IZR (Zfloor x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
Hy:(/ 2 < y - IZR (Zfloor x))%R
((if choice (Zfloor x) then Zceil x else Zfloor x) <= Zceil y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(/ 2 < x - IZR (Zfloor x))%R
(Zceil x <= Znearest y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
forall n : Z, Znearest (IZR n) = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
Hy:(y - IZR (Zfloor x))%R = (/ 2)%R

(x - IZR (Zfloor x))%R = (y - IZR (Zfloor x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
Hy:(/ 2 < y - IZR (Zfloor x))%R
((if choice (Zfloor x) then Zceil x else Zfloor x) <= Zceil y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(/ 2 < x - IZR (Zfloor x))%R
(Zceil x <= Znearest y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
forall n : Z, Znearest (IZR n) = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
Hy:(/ 2 < y - IZR (Zfloor x))%R

((if choice (Zfloor x) then Zceil x else Zfloor x) <= Zceil y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(/ 2 < x - IZR (Zfloor x))%R
(Zceil x <= Znearest y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
forall n : Z, Znearest (IZR n) = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
Hy:(/ 2 < y - IZR (Zfloor x))%R

((if choice (Zfloor x) then Zceil x else Zfloor x) <= Zceil x)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
Hy:(/ 2 < y - IZR (Zfloor x))%R
(Zceil x <= Zceil y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(/ 2 < x - IZR (Zfloor x))%R
(Zceil x <= Znearest y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
forall n : Z, Znearest (IZR n) = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
Hy:(/ 2 < y - IZR (Zfloor x))%R

(Zceil x <= Zceil x)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
Hy:(/ 2 < y - IZR (Zfloor x))%R
(Zfloor x <= Zceil x)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
Hy:(/ 2 < y - IZR (Zfloor x))%R
(Zceil x <= Zceil y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(/ 2 < x - IZR (Zfloor x))%R
(Zceil x <= Znearest y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
forall n : Z, Znearest (IZR n) = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
Hy:(/ 2 < y - IZR (Zfloor x))%R

(Zfloor x <= Zceil x)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
Hy:(/ 2 < y - IZR (Zfloor x))%R
(Zceil x <= Zceil y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(/ 2 < x - IZR (Zfloor x))%R
(Zceil x <= Znearest y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
forall n : Z, Znearest (IZR n) = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
Hy:(/ 2 < y - IZR (Zfloor x))%R

(IZR (Zfloor x) <= IZR (Zceil x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
Hy:(/ 2 < y - IZR (Zfloor x))%R
(Zceil x <= Zceil y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(/ 2 < x - IZR (Zfloor x))%R
(Zceil x <= Znearest y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
forall n : Z, Znearest (IZR n) = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
Hy:(/ 2 < y - IZR (Zfloor x))%R

(IZR (Zfloor x) <= x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
Hy:(/ 2 < y - IZR (Zfloor x))%R
(x <= IZR (Zceil x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
Hy:(/ 2 < y - IZR (Zfloor x))%R
(Zceil x <= Zceil y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(/ 2 < x - IZR (Zfloor x))%R
(Zceil x <= Znearest y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
forall n : Z, Znearest (IZR n) = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
Hy:(/ 2 < y - IZR (Zfloor x))%R

(x <= IZR (Zceil x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
Hy:(/ 2 < y - IZR (Zfloor x))%R
(Zceil x <= Zceil y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(/ 2 < x - IZR (Zfloor x))%R
(Zceil x <= Znearest y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
forall n : Z, Znearest (IZR n) = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
Hy:(/ 2 < y - IZR (Zfloor x))%R

(Zceil x <= Zceil y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(/ 2 < x - IZR (Zfloor x))%R
(Zceil x <= Znearest y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
forall n : Z, Znearest (IZR n) = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(/ 2 < x - IZR (Zfloor x))%R

(Zceil x <= Znearest y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
forall n : Z, Znearest (IZR n) = n
(* .. *)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(/ 2 < x - IZR (Zfloor x))%R

(Zceil x <= match Rcompare (y - IZR (Zfloor y)) (/ 2) with | Eq => if choice (Zfloor y) then Zceil y else Zfloor y | Lt => Zfloor y | Gt => Zceil y end)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
forall n : Z, Znearest (IZR n) = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(/ 2 < x - IZR (Zfloor x))%R

(Zceil x <= match Rcompare (y - IZR (Zfloor x)) (/ 2) with | Eq => if choice (Zfloor x) then Zceil y else Zfloor x | Lt => Zfloor x | Gt => Zceil y end)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
forall n : Z, Znearest (IZR n) = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(/ 2 < x - IZR (Zfloor x))%R

(Zceil x <= Zceil y)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(/ 2 < x - IZR (Zfloor x))%R
(/ 2 < y - IZR (Zfloor x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
forall n : Z, Znearest (IZR n) = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(/ 2 < x - IZR (Zfloor x))%R

(/ 2 < y - IZR (Zfloor x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
forall n : Z, Znearest (IZR n) = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x, y:R
Hxy:(x <= y)%R
H:(y < IZR (Zceil x))%R
Hf:Zfloor y = Zfloor x
Hx:(/ 2 < x - IZR (Zfloor x))%R

(x - IZR (Zfloor x) <= y - IZR (Zfloor x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
forall n : Z, Znearest (IZR n) = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool

forall n : Z, Znearest (IZR n) = n
(* *)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
n:Z

Znearest (IZR n) = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
n:Z

match Rcompare (IZR n - IZR (Zfloor (IZR n))) (/ 2) with | Eq => if choice (Zfloor (IZR n)) then Zceil (IZR n) else Zfloor (IZR n) | Lt => Zfloor (IZR n) | Gt => Zceil (IZR n) end = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
n:Z

match Rcompare (IZR n - IZR n) (/ 2) with | Eq => if choice n then Zceil (IZR n) else n | Lt => n | Gt => Zceil (IZR n) end = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
n:Z

n = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
n:Z
(IZR n - IZR n < / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
n:Z

(IZR n - IZR n < / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
n:Z

(IZR n + - IZR n < / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
n:Z

(0 < / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
n:Z

(0 < 2)%R
now apply IZR_lt. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool

forall x : R, (x - IZR (Zfloor x))%R <> (/ 2)%R -> (Rabs (x - IZR (Znearest x)) < / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool

forall x : R, (x - IZR (Zfloor x))%R <> (/ 2)%R -> (Rabs (x - IZR (Znearest x)) < / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R

(Rabs (x - IZR (Znearest x)) < / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R

(Rabs (x - IZR match Rcompare (x - IZR (Zfloor x)) (/ 2) with | Eq => if choice (Zfloor x) then Zceil x else Zfloor x | Lt => Zfloor x | Gt => Zceil x end) < / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
H:(x - IZR (Zfloor x) < / 2)%R

(Rabs (x - IZR (Zfloor x)) < / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
H:(x - IZR (Zfloor x))%R = (/ 2)%R
(Rabs (x - IZR (if choice (Zfloor x) then Zceil x else Zfloor x)) < / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
H:(/ 2 < x - IZR (Zfloor x))%R
(Rabs (x - IZR (Zceil x)) < / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
H:(x - IZR (Zfloor x) < / 2)%R

(x - IZR (Zfloor x) < / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
H:(x - IZR (Zfloor x) < / 2)%R
(0 <= x - IZR (Zfloor x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
H:(x - IZR (Zfloor x))%R = (/ 2)%R
(Rabs (x - IZR (if choice (Zfloor x) then Zceil x else Zfloor x)) < / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
H:(/ 2 < x - IZR (Zfloor x))%R
(Rabs (x - IZR (Zceil x)) < / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
H:(x - IZR (Zfloor x) < / 2)%R

(0 <= x - IZR (Zfloor x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
H:(x - IZR (Zfloor x))%R = (/ 2)%R
(Rabs (x - IZR (if choice (Zfloor x) then Zceil x else Zfloor x)) < / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
H:(/ 2 < x - IZR (Zfloor x))%R
(Rabs (x - IZR (Zceil x)) < / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
H:(x - IZR (Zfloor x) < / 2)%R

(IZR (Zfloor x) <= x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
H:(x - IZR (Zfloor x))%R = (/ 2)%R
(Rabs (x - IZR (if choice (Zfloor x) then Zceil x else Zfloor x)) < / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
H:(/ 2 < x - IZR (Zfloor x))%R
(Rabs (x - IZR (Zceil x)) < / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
H:(x - IZR (Zfloor x))%R = (/ 2)%R

(Rabs (x - IZR (if choice (Zfloor x) then Zceil x else Zfloor x)) < / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
H:(/ 2 < x - IZR (Zfloor x))%R
(Rabs (x - IZR (Zceil x)) < / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
H:(/ 2 < x - IZR (Zfloor x))%R

(Rabs (x - IZR (Zceil x)) < / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
H:(/ 2 < x - IZR (Zfloor x))%R

(- (x - IZR (Zceil x)) < / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
H:(/ 2 < x - IZR (Zfloor x))%R
(x - IZR (Zceil x) <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
H:(/ 2 < x - IZR (Zfloor x))%R

(IZR (Zceil x) - x < / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
H:(/ 2 < x - IZR (Zfloor x))%R
(x - IZR (Zceil x) <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
H:(/ 2 < x - IZR (Zfloor x))%R

(IZR (Zfloor x + 1) - x < / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
H:(/ 2 < x - IZR (Zfloor x))%R
IZR (Zfloor x) <> x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
H:(/ 2 < x - IZR (Zfloor x))%R
(x - IZR (Zceil x) <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
H:(/ 2 < x - IZR (Zfloor x))%R

(IZR (Zfloor x) + 1 - x < / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
H:(/ 2 < x - IZR (Zfloor x))%R
IZR (Zfloor x) <> x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
H:(/ 2 < x - IZR (Zfloor x))%R
(x - IZR (Zceil x) <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
H:(/ 2 < x - IZR (Zfloor x))%R

(- / 2 < - (IZR (Zfloor x) + 1 - x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
H:(/ 2 < x - IZR (Zfloor x))%R
IZR (Zfloor x) <> x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
H:(/ 2 < x - IZR (Zfloor x))%R
(x - IZR (Zceil x) <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
H:(/ 2 < x - IZR (Zfloor x))%R

(1 + - / 2 < 1 + - (IZR (Zfloor x) + 1 - x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
H:(/ 2 < x - IZR (Zfloor x))%R
IZR (Zfloor x) <> x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
H:(/ 2 < x - IZR (Zfloor x))%R
(x - IZR (Zceil x) <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
H:(/ 2 < x - IZR (Zfloor x))%R

(/ 2 < 1 + - (IZR (Zfloor x) + 1 - x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
H:(/ 2 < x - IZR (Zfloor x))%R
IZR (Zfloor x) <> x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
H:(/ 2 < x - IZR (Zfloor x))%R
(x - IZR (Zceil x) <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
H:(/ 2 < x - IZR (Zfloor x))%R

IZR (Zfloor x) <> x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
H:(/ 2 < x - IZR (Zfloor x))%R
(x - IZR (Zceil x) <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
H:(/ 2 < x - IZR (Zfloor x))%R

(IZR (Zfloor x) < x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
H:(/ 2 < x - IZR (Zfloor x))%R
(x - IZR (Zceil x) <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
H:(/ 2 < x - IZR (Zfloor x))%R

(- IZR (Zfloor x) + IZR (Zfloor x) < - IZR (Zfloor x) + x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
H:(/ 2 < x - IZR (Zfloor x))%R
(x - IZR (Zceil x) <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
H:(/ 2 < x - IZR (Zfloor x))%R

(- IZR (Zfloor x) + IZR (Zfloor x) < / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
H:(/ 2 < x - IZR (Zfloor x))%R
(/ 2 < - IZR (Zfloor x) + x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
H:(/ 2 < x - IZR (Zfloor x))%R
(x - IZR (Zceil x) <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
H:(/ 2 < x - IZR (Zfloor x))%R

(0 < / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
H:(/ 2 < x - IZR (Zfloor x))%R
(/ 2 < - IZR (Zfloor x) + x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
H:(/ 2 < x - IZR (Zfloor x))%R
(x - IZR (Zceil x) <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
H:(/ 2 < x - IZR (Zfloor x))%R

(0 < 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
H:(/ 2 < x - IZR (Zfloor x))%R
(/ 2 < - IZR (Zfloor x) + x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
H:(/ 2 < x - IZR (Zfloor x))%R
(x - IZR (Zceil x) <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
H:(/ 2 < x - IZR (Zfloor x))%R

(/ 2 < - IZR (Zfloor x) + x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
H:(/ 2 < x - IZR (Zfloor x))%R
(x - IZR (Zceil x) <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
H:(/ 2 < x - IZR (Zfloor x))%R

(x - IZR (Zceil x) <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
H:(/ 2 < x - IZR (Zfloor x))%R

(x <= IZR (Zceil x))%R
apply Zceil_ub. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool

forall x : R, (Rabs (x - IZR (Znearest x)) <= / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool

forall x : R, (Rabs (x - IZR (Znearest x)) <= / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R

(Rabs (x - IZR (Znearest x)) <= / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R

(Rabs (x - IZR (Znearest x)) <= / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
(Rabs (x - IZR (Znearest x)) <= / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R

(Rabs (/ 2) <= / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
K:(Rabs (/ 2) <= / 2)%R
(Rabs (x - IZR (Znearest x)) <= / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
(Rabs (x - IZR (Znearest x)) <= / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R

Rabs (/ 2) = (/ 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
K:(Rabs (/ 2) <= / 2)%R
(Rabs (x - IZR (Znearest x)) <= / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
(Rabs (x - IZR (Znearest x)) <= / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R

(0 <= / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
K:(Rabs (/ 2) <= / 2)%R
(Rabs (x - IZR (Znearest x)) <= / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
(Rabs (x - IZR (Znearest x)) <= / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R

(0 < / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
K:(Rabs (/ 2) <= / 2)%R
(Rabs (x - IZR (Znearest x)) <= / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
(Rabs (x - IZR (Znearest x)) <= / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R

(0 < 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
K:(Rabs (/ 2) <= / 2)%R
(Rabs (x - IZR (Znearest x)) <= / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
(Rabs (x - IZR (Znearest x)) <= / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
K:(Rabs (/ 2) <= / 2)%R

(Rabs (x - IZR (Znearest x)) <= / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
(Rabs (x - IZR (Znearest x)) <= / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
K:(Rabs (/ 2) <= / 2)%R

(Rabs (x - IZR (Zfloor x)) <= / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
K:(Rabs (/ 2) <= / 2)%R
(Rabs (x - IZR (Zceil x)) <= / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
(Rabs (x - IZR (Znearest x)) <= / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
K:(Rabs (/ 2) <= / 2)%R

(Rabs (x - IZR (Zceil x)) <= / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
(Rabs (x - IZR (Znearest x)) <= / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
K:(Rabs (/ 2) <= / 2)%R

(Rabs (x - IZR (Zfloor x + 1)) <= / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
K:(Rabs (/ 2) <= / 2)%R
IZR (Zfloor x) <> x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
(Rabs (x - IZR (Znearest x)) <= / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
K:(Rabs (/ 2) <= / 2)%R

(Rabs (x - (IZR (Zfloor x) + 1)) <= / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
K:(Rabs (/ 2) <= / 2)%R
IZR (Zfloor x) <> x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
(Rabs (x - IZR (Znearest x)) <= / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
K:(Rabs (/ 2) <= / 2)%R

(Rabs (x - IZR (Zfloor x) - 1) <= / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
K:(Rabs (/ 2) <= / 2)%R
IZR (Zfloor x) <> x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
(Rabs (x - IZR (Znearest x)) <= / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
K:(Rabs (/ 2) <= / 2)%R

(Rabs (/ 2 - 1) <= / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
K:(Rabs (/ 2) <= / 2)%R
IZR (Zfloor x) <> x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
(Rabs (x - IZR (Znearest x)) <= / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
K:(Rabs (/ 2) <= / 2)%R

(Rabs (1 - / 2) <= / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
K:(Rabs (/ 2) <= / 2)%R
IZR (Zfloor x) <> x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
(Rabs (x - IZR (Znearest x)) <= / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
K:(Rabs (/ 2) <= / 2)%R

IZR (Zfloor x) <> x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
(Rabs (x - IZR (Znearest x)) <= / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
K:(Rabs (/ 2) <= / 2)%R

(IZR (Zfloor x) < x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
(Rabs (x - IZR (Znearest x)) <= / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
K:(Rabs (/ 2) <= / 2)%R

(- IZR (Zfloor x) + IZR (Zfloor x) < - IZR (Zfloor x) + x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
(Rabs (x - IZR (Znearest x)) <= / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
K:(Rabs (/ 2) <= / 2)%R

(0 < x + - IZR (Zfloor x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
(Rabs (x - IZR (Znearest x)) <= / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
K:(Rabs (/ 2) <= / 2)%R

(0 < x - IZR (Zfloor x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
(Rabs (x - IZR (Znearest x)) <= / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
K:(Rabs (/ 2) <= / 2)%R

(0 < / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
(Rabs (x - IZR (Znearest x)) <= / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R = (/ 2)%R
K:(Rabs (/ 2) <= / 2)%R

(0 < 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R
(Rabs (x - IZR (Znearest x)) <= / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R

(Rabs (x - IZR (Znearest x)) <= / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:(x - IZR (Zfloor x))%R <> (/ 2)%R

(Rabs (x - IZR (Znearest x)) < / 2)%R
now apply Znearest_N_strict. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool

forall (x : R) (n : Z), (Rabs (x - IZR n) < / 2)%R -> Znearest x = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool

forall (x : R) (n : Z), (Rabs (x - IZR n) < / 2)%R -> Znearest x = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
n:Z
Hd:(Rabs (x - IZR n) < / 2)%R

Znearest x = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
n:Z
Hd:(Rabs (x - IZR n) < / 2)%R

(Z.abs (Znearest x - n) < 1)%Z -> Znearest x = n
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
n:Z
Hd:(Rabs (x - IZR n) < / 2)%R
(Z.abs (Znearest x - n) < 1)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
n:Z
Hd:(Rabs (x - IZR n) < / 2)%R

(Z.abs (Znearest x - n) < 1)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
n:Z
Hd:(Rabs (x - IZR n) < / 2)%R

(IZR (Z.abs (Znearest x - n)) < 1)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
n:Z
Hd:(Rabs (x - IZR n) < / 2)%R

(Rabs (IZR (Znearest x) - IZR n) < 1)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
n:Z
Hd:(Rabs (x - IZR n) < / 2)%R

(Rabs (- (x - IZR (Znearest x)) + (x - IZR n)) < 1)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
n:Z
Hd:(Rabs (x - IZR n) < / 2)%R

(Rabs (- (x - IZR (Znearest x))) + Rabs (x - IZR n) < 1)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
n:Z
Hd:(Rabs (x - IZR n) < / 2)%R

(Rabs (- (x - IZR (Znearest x))) + Rabs (x - IZR n) < 1)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
n:Z
Hd:(Rabs (x - IZR n) < / 2)%R

(Rabs (- (x - IZR (Znearest x))) + Rabs (x - IZR n) < / 2 + / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
n:Z
Hd:(Rabs (x - IZR n) < / 2)%R

(Rabs (- (x - IZR (Znearest x))) <= / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
n:Z
Hd:(Rabs (x - IZR n) < / 2)%R

(Rabs (x - IZR (Znearest x)) <= / 2)%R
apply Znearest_half. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool

forall x : R, Rnd_N_pt generic_format x (round Znearest x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool

forall x : R, Rnd_N_pt generic_format x (round Znearest x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R

Rnd_N_pt generic_format x (round Znearest x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R

Rnd_N_pt generic_format x (round Znearest x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R

Rnd_N_pt generic_format x (round Znearest x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R

Rnd_N_pt generic_format x (round Znearest x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R

Rnd_N_pt generic_format x (round Znearest x)
(* . *)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R

(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
Rnd_N_pt generic_format x (round Znearest x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R

(Rabs (round Znearest x - scaled_mantissa x * bpow (cexp x)) <= Rmin (scaled_mantissa x * bpow (cexp x) - d) (u - scaled_mantissa x * bpow (cexp x)))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
Rnd_N_pt generic_format x (round Znearest x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R

(Rabs (IZR (Fnum {| Fnum := Znearest (scaled_mantissa x); Fexp := cexp x |}) * bpow (Fexp {| Fnum := Znearest (scaled_mantissa x); Fexp := cexp x |}) - scaled_mantissa x * bpow (cexp x)) <= Rmin (scaled_mantissa x * bpow (cexp x) - IZR (Fnum {| Fnum := Zfloor (scaled_mantissa x); Fexp := cexp x |}) * bpow (Fexp {| Fnum := Zfloor (scaled_mantissa x); Fexp := cexp x |})) (IZR (Fnum {| Fnum := Zceil (scaled_mantissa x); Fexp := cexp x |}) * bpow (Fexp {| Fnum := Zceil (scaled_mantissa x); Fexp := cexp x |}) - scaled_mantissa x * bpow (cexp x)))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
Rnd_N_pt generic_format x (round Znearest x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R

(Rabs (IZR (Znearest (scaled_mantissa x)) * bpow (cexp x) - scaled_mantissa x * bpow (cexp x)) <= Rmin (scaled_mantissa x * bpow (cexp x) - IZR (Zfloor (scaled_mantissa x)) * bpow (cexp x)) (IZR (Zceil (scaled_mantissa x)) * bpow (cexp x) - scaled_mantissa x * bpow (cexp x)))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
Rnd_N_pt generic_format x (round Znearest x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R

(Rabs (IZR (Znearest mx) * bx - mx * bx) <= Rmin (mx * bx - IZR (Zfloor mx) * bx) (IZR (Zceil mx) * bx - mx * bx))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
Rnd_N_pt generic_format x (round Znearest x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R

(Rabs ((IZR (Znearest mx) - mx) * bx) <= Rmin ((mx - IZR (Zfloor mx)) * bx) ((IZR (Zceil mx) - mx) * bx))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
Rnd_N_pt generic_format x (round Znearest x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R

(Rabs (IZR (Znearest mx) - mx) * bx <= Rmin ((mx - IZR (Zfloor mx)) * bx) ((IZR (Zceil mx) - mx) * bx))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
(0 <= bx)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
Rnd_N_pt generic_format x (round Znearest x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R

(Rabs (IZR (Znearest mx) - mx) * bx <= Rmin ((mx - IZR (Zfloor mx)) * bx) ((IZR (Zceil mx) - mx) * bx))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
Rnd_N_pt generic_format x (round Znearest x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R

(Rabs (IZR (Znearest mx) - mx) * bx <= Rmin (mx - IZR (Zfloor mx)) (IZR (Zceil mx) - mx) * bx)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
(0 <= bx)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
Rnd_N_pt generic_format x (round Znearest x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R

(Rabs (IZR (Znearest mx) - mx) * bx <= Rmin (mx - IZR (Zfloor mx)) (IZR (Zceil mx) - mx) * bx)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
Rnd_N_pt generic_format x (round Znearest x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R

(0 <= bx)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
(Rabs (IZR (Znearest mx) - mx) <= Rmin (mx - IZR (Zfloor mx)) (IZR (Zceil mx) - mx))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
Rnd_N_pt generic_format x (round Znearest x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R

(Rabs (IZR (Znearest mx) - mx) <= Rmin (mx - IZR (Zfloor mx)) (IZR (Zceil mx) - mx))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
Rnd_N_pt generic_format x (round Znearest x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R

(Rabs (IZR match Rcompare (mx - IZR (Zfloor mx)) (/ 2) with | Eq => if choice (Zfloor mx) then Zceil mx else Zfloor mx | Lt => Zfloor mx | Gt => Zceil mx end - mx) <= Rmin (mx - IZR (Zfloor mx)) (IZR (Zceil mx) - mx))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
Rnd_N_pt generic_format x (round Znearest x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) = mx

(Rabs (IZR match Rcompare (mx - IZR (Zfloor mx)) (/ 2) with | Eq => if choice (Zfloor mx) then Zceil mx else Zfloor mx | Lt => Zfloor mx | Gt => Zceil mx end - mx) <= Rmin (mx - IZR (Zfloor mx)) (IZR (Zceil mx) - mx))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) <> mx
(Rabs (IZR match Rcompare (mx - IZR (Zfloor mx)) (/ 2) with | Eq => if choice (Zfloor mx) then Zceil mx else Zfloor mx | Lt => Zfloor mx | Gt => Zceil mx end - mx) <= Rmin (mx - IZR (Zfloor mx)) (IZR (Zceil mx) - mx))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
Rnd_N_pt generic_format x (round Znearest x)
(* .. *)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) = mx

(Rabs (IZR match Rcompare (mx - mx) (/ 2) with | Eq => if choice (Zfloor mx) then Zceil mx else Zfloor mx | Lt => Zfloor mx | Gt => Zceil mx end - mx) <= Rmin (mx - mx) (IZR (Zceil mx) - mx))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) <> mx
(Rabs (IZR match Rcompare (mx - IZR (Zfloor mx)) (/ 2) with | Eq => if choice (Zfloor mx) then Zceil mx else Zfloor mx | Lt => Zfloor mx | Gt => Zceil mx end - mx) <= Rmin (mx - IZR (Zfloor mx)) (IZR (Zceil mx) - mx))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
Rnd_N_pt generic_format x (round Znearest x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) = mx

(Rabs (IZR match Rcompare (mx + - mx) (/ 2) with | Eq => if choice (Zfloor mx) then Zceil mx else Zfloor mx | Lt => Zfloor mx | Gt => Zceil mx end - mx) <= Rmin (mx - mx) (IZR (Zceil mx) - mx))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) <> mx
(Rabs (IZR match Rcompare (mx - IZR (Zfloor mx)) (/ 2) with | Eq => if choice (Zfloor mx) then Zceil mx else Zfloor mx | Lt => Zfloor mx | Gt => Zceil mx end - mx) <= Rmin (mx - IZR (Zfloor mx)) (IZR (Zceil mx) - mx))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
Rnd_N_pt generic_format x (round Znearest x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) = mx

(Rabs (IZR match Rcompare 0 (/ 2) with | Eq => if choice (Zfloor mx) then Zceil mx else Zfloor mx | Lt => Zfloor mx | Gt => Zceil mx end - mx) <= Rmin (mx - mx) (IZR (Zceil mx) - mx))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) <> mx
(Rabs (IZR match Rcompare (mx - IZR (Zfloor mx)) (/ 2) with | Eq => if choice (Zfloor mx) then Zceil mx else Zfloor mx | Lt => Zfloor mx | Gt => Zceil mx end - mx) <= Rmin (mx - IZR (Zfloor mx)) (IZR (Zceil mx) - mx))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
Rnd_N_pt generic_format x (round Znearest x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) = mx

(Rabs (IZR (Zfloor mx) - mx) <= Rmin (mx - mx) (IZR (Zceil mx) - mx))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) = mx
(0 < / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) <> mx
(Rabs (IZR match Rcompare (mx - IZR (Zfloor mx)) (/ 2) with | Eq => if choice (Zfloor mx) then Zceil mx else Zfloor mx | Lt => Zfloor mx | Gt => Zceil mx end - mx) <= Rmin (mx - IZR (Zfloor mx)) (IZR (Zceil mx) - mx))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
Rnd_N_pt generic_format x (round Znearest x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) = mx

(Rabs (mx - mx) <= Rmin (mx - mx) (IZR (Zceil mx) - mx))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) = mx
(0 < / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) <> mx
(Rabs (IZR match Rcompare (mx - IZR (Zfloor mx)) (/ 2) with | Eq => if choice (Zfloor mx) then Zceil mx else Zfloor mx | Lt => Zfloor mx | Gt => Zceil mx end - mx) <= Rmin (mx - IZR (Zfloor mx)) (IZR (Zceil mx) - mx))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
Rnd_N_pt generic_format x (round Znearest x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) = mx

(Rabs (mx + - mx) <= Rmin (mx + - mx) (IZR (Zceil mx) - mx))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) = mx
(0 < / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) <> mx
(Rabs (IZR match Rcompare (mx - IZR (Zfloor mx)) (/ 2) with | Eq => if choice (Zfloor mx) then Zceil mx else Zfloor mx | Lt => Zfloor mx | Gt => Zceil mx end - mx) <= Rmin (mx - IZR (Zfloor mx)) (IZR (Zceil mx) - mx))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
Rnd_N_pt generic_format x (round Znearest x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) = mx

(Rabs 0 <= Rmin 0 (IZR (Zceil mx) - mx))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) = mx
(0 < / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) <> mx
(Rabs (IZR match Rcompare (mx - IZR (Zfloor mx)) (/ 2) with | Eq => if choice (Zfloor mx) then Zceil mx else Zfloor mx | Lt => Zfloor mx | Gt => Zceil mx end - mx) <= Rmin (mx - IZR (Zfloor mx)) (IZR (Zceil mx) - mx))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
Rnd_N_pt generic_format x (round Znearest x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) = mx

(0 <= Rmin 0 (IZR (Zceil mx) - mx))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) = mx
(0 < / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) <> mx
(Rabs (IZR match Rcompare (mx - IZR (Zfloor mx)) (/ 2) with | Eq => if choice (Zfloor mx) then Zceil mx else Zfloor mx | Lt => Zfloor mx | Gt => Zceil mx end - mx) <= Rmin (mx - IZR (Zfloor mx)) (IZR (Zceil mx) - mx))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
Rnd_N_pt generic_format x (round Znearest x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) = mx

(0 <= (if Rle_dec 0 (IZR (Zceil mx) - mx) then 0 else IZR (Zceil mx) - mx))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) = mx
(0 < / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) <> mx
(Rabs (IZR match Rcompare (mx - IZR (Zfloor mx)) (/ 2) with | Eq => if choice (Zfloor mx) then Zceil mx else Zfloor mx | Lt => Zfloor mx | Gt => Zceil mx end - mx) <= Rmin (mx - IZR (Zfloor mx)) (IZR (Zceil mx) - mx))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
Rnd_N_pt generic_format x (round Znearest x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) = mx
H:(0 <= IZR (Zceil mx) - mx)%R

(0 <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) = mx
H:~ (0 <= IZR (Zceil mx) - mx)%R
(0 <= IZR (Zceil mx) - mx)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) = mx
(0 < / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) <> mx
(Rabs (IZR match Rcompare (mx - IZR (Zfloor mx)) (/ 2) with | Eq => if choice (Zfloor mx) then Zceil mx else Zfloor mx | Lt => Zfloor mx | Gt => Zceil mx end - mx) <= Rmin (mx - IZR (Zfloor mx)) (IZR (Zceil mx) - mx))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
Rnd_N_pt generic_format x (round Znearest x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) = mx
H:~ (0 <= IZR (Zceil mx) - mx)%R

(0 <= IZR (Zceil mx) - mx)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) = mx
(0 < / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) <> mx
(Rabs (IZR match Rcompare (mx - IZR (Zfloor mx)) (/ 2) with | Eq => if choice (Zfloor mx) then Zceil mx else Zfloor mx | Lt => Zfloor mx | Gt => Zceil mx end - mx) <= Rmin (mx - IZR (Zfloor mx)) (IZR (Zceil mx) - mx))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
Rnd_N_pt generic_format x (round Znearest x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) = mx
H:~ (0 <= IZR (Zceil mx) - mx)%R

(mx <= IZR (Zceil mx))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) = mx
(0 < / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) <> mx
(Rabs (IZR match Rcompare (mx - IZR (Zfloor mx)) (/ 2) with | Eq => if choice (Zfloor mx) then Zceil mx else Zfloor mx | Lt => Zfloor mx | Gt => Zceil mx end - mx) <= Rmin (mx - IZR (Zfloor mx)) (IZR (Zceil mx) - mx))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
Rnd_N_pt generic_format x (round Znearest x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) = mx

(0 < / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) <> mx
(Rabs (IZR match Rcompare (mx - IZR (Zfloor mx)) (/ 2) with | Eq => if choice (Zfloor mx) then Zceil mx else Zfloor mx | Lt => Zfloor mx | Gt => Zceil mx end - mx) <= Rmin (mx - IZR (Zfloor mx)) (IZR (Zceil mx) - mx))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
Rnd_N_pt generic_format x (round Znearest x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) = mx

(0 < 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) <> mx
(Rabs (IZR match Rcompare (mx - IZR (Zfloor mx)) (/ 2) with | Eq => if choice (Zfloor mx) then Zceil mx else Zfloor mx | Lt => Zfloor mx | Gt => Zceil mx end - mx) <= Rmin (mx - IZR (Zfloor mx)) (IZR (Zceil mx) - mx))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
Rnd_N_pt generic_format x (round Znearest x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) <> mx

(Rabs (IZR match Rcompare (mx - IZR (Zfloor mx)) (/ 2) with | Eq => if choice (Zfloor mx) then Zceil mx else Zfloor mx | Lt => Zfloor mx | Gt => Zceil mx end - mx) <= Rmin (mx - IZR (Zfloor mx)) (IZR (Zceil mx) - mx))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
Rnd_N_pt generic_format x (round Znearest x)
(* .. *)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) <> mx

(Rabs (IZR match Rcompare (mx - IZR (Zfloor mx)) (IZR (Zceil mx) - mx) with | Eq => if choice (Zfloor mx) then Zceil mx else Zfloor mx | Lt => Zfloor mx | Gt => Zceil mx end - mx) <= Rmin (mx - IZR (Zfloor mx)) (IZR (Zceil mx) - mx))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
Rnd_N_pt generic_format x (round Znearest x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) <> mx

(Rabs (IZR match Rcompare (mx - IZR (Zfloor mx)) (IZR (Zceil mx) - mx) with | Eq => if choice (Zfloor mx) then Zceil mx else Zfloor mx | Lt => Zfloor mx | Gt => Zceil mx end - mx) <= match Rcompare (mx - IZR (Zfloor mx)) (IZR (Zceil mx) - mx) with | Gt => IZR (Zceil mx) - mx | _ => mx - IZR (Zfloor mx) end)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
Rnd_N_pt generic_format x (round Znearest x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) <> mx

(Rabs (mx - IZR (Zfloor mx)) <= mx - IZR (Zfloor mx))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) <> mx
H:(Rabs (mx - IZR (Zfloor mx)) <= mx - IZR (Zfloor mx))%R
(Rabs (IZR match Rcompare (mx - IZR (Zfloor mx)) (IZR (Zceil mx) - mx) with | Eq => if choice (Zfloor mx) then Zceil mx else Zfloor mx | Lt => Zfloor mx | Gt => Zceil mx end - mx) <= match Rcompare (mx - IZR (Zfloor mx)) (IZR (Zceil mx) - mx) with | Gt => IZR (Zceil mx) - mx | _ => mx - IZR (Zfloor mx) end)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
Rnd_N_pt generic_format x (round Znearest x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) <> mx

(mx - IZR (Zfloor mx) <= mx - IZR (Zfloor mx))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) <> mx
(0 <= mx - IZR (Zfloor mx))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) <> mx
H:(Rabs (mx - IZR (Zfloor mx)) <= mx - IZR (Zfloor mx))%R
(Rabs (IZR match Rcompare (mx - IZR (Zfloor mx)) (IZR (Zceil mx) - mx) with | Eq => if choice (Zfloor mx) then Zceil mx else Zfloor mx | Lt => Zfloor mx | Gt => Zceil mx end - mx) <= match Rcompare (mx - IZR (Zfloor mx)) (IZR (Zceil mx) - mx) with | Gt => IZR (Zceil mx) - mx | _ => mx - IZR (Zfloor mx) end)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
Rnd_N_pt generic_format x (round Znearest x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) <> mx

(0 <= mx - IZR (Zfloor mx))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) <> mx
H:(Rabs (mx - IZR (Zfloor mx)) <= mx - IZR (Zfloor mx))%R
(Rabs (IZR match Rcompare (mx - IZR (Zfloor mx)) (IZR (Zceil mx) - mx) with | Eq => if choice (Zfloor mx) then Zceil mx else Zfloor mx | Lt => Zfloor mx | Gt => Zceil mx end - mx) <= match Rcompare (mx - IZR (Zfloor mx)) (IZR (Zceil mx) - mx) with | Gt => IZR (Zceil mx) - mx | _ => mx - IZR (Zfloor mx) end)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
Rnd_N_pt generic_format x (round Znearest x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) <> mx

(IZR (Zfloor mx) <= mx)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) <> mx
H:(Rabs (mx - IZR (Zfloor mx)) <= mx - IZR (Zfloor mx))%R
(Rabs (IZR match Rcompare (mx - IZR (Zfloor mx)) (IZR (Zceil mx) - mx) with | Eq => if choice (Zfloor mx) then Zceil mx else Zfloor mx | Lt => Zfloor mx | Gt => Zceil mx end - mx) <= match Rcompare (mx - IZR (Zfloor mx)) (IZR (Zceil mx) - mx) with | Gt => IZR (Zceil mx) - mx | _ => mx - IZR (Zfloor mx) end)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
Rnd_N_pt generic_format x (round Znearest x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) <> mx
H:(Rabs (mx - IZR (Zfloor mx)) <= mx - IZR (Zfloor mx))%R

(Rabs (IZR match Rcompare (mx - IZR (Zfloor mx)) (IZR (Zceil mx) - mx) with | Eq => if choice (Zfloor mx) then Zceil mx else Zfloor mx | Lt => Zfloor mx | Gt => Zceil mx end - mx) <= match Rcompare (mx - IZR (Zfloor mx)) (IZR (Zceil mx) - mx) with | Gt => IZR (Zceil mx) - mx | _ => mx - IZR (Zfloor mx) end)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
Rnd_N_pt generic_format x (round Znearest x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) <> mx
H:(Rabs (mx - IZR (Zfloor mx)) <= mx - IZR (Zfloor mx))%R
Hm':(mx - IZR (Zfloor mx) < IZR (Zceil mx) - mx)%R

(Rabs (IZR (Zfloor mx) - mx) <= mx - IZR (Zfloor mx))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) <> mx
H:(Rabs (mx - IZR (Zfloor mx)) <= mx - IZR (Zfloor mx))%R
Hm':(mx - IZR (Zfloor mx))%R = (IZR (Zceil mx) - mx)%R
(Rabs (IZR (if choice (Zfloor mx) then Zceil mx else Zfloor mx) - mx) <= mx - IZR (Zfloor mx))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) <> mx
H:(Rabs (mx - IZR (Zfloor mx)) <= mx - IZR (Zfloor mx))%R
Hm':(IZR (Zceil mx) - mx < mx - IZR (Zfloor mx))%R
(Rabs (IZR (Zceil mx) - mx) <= IZR (Zceil mx) - mx)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
Rnd_N_pt generic_format x (round Znearest x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) <> mx
H:(Rabs (mx - IZR (Zfloor mx)) <= mx - IZR (Zfloor mx))%R
Hm':(mx - IZR (Zfloor mx))%R = (IZR (Zceil mx) - mx)%R

(Rabs (IZR (if choice (Zfloor mx) then Zceil mx else Zfloor mx) - mx) <= mx - IZR (Zfloor mx))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) <> mx
H:(Rabs (mx - IZR (Zfloor mx)) <= mx - IZR (Zfloor mx))%R
Hm':(IZR (Zceil mx) - mx < mx - IZR (Zfloor mx))%R
(Rabs (IZR (Zceil mx) - mx) <= IZR (Zceil mx) - mx)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
Rnd_N_pt generic_format x (round Znearest x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) <> mx
H:(Rabs (mx - IZR (Zfloor mx)) <= mx - IZR (Zfloor mx))%R
Hm':(mx - IZR (Zfloor mx))%R = (IZR (Zceil mx) - mx)%R

(Rabs (IZR (Zceil mx) - mx) <= mx - IZR (Zfloor mx))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) <> mx
H:(Rabs (mx - IZR (Zfloor mx)) <= mx - IZR (Zfloor mx))%R
Hm':(mx - IZR (Zfloor mx))%R = (IZR (Zceil mx) - mx)%R
(Rabs (IZR (Zfloor mx) - mx) <= mx - IZR (Zfloor mx))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) <> mx
H:(Rabs (mx - IZR (Zfloor mx)) <= mx - IZR (Zfloor mx))%R
Hm':(IZR (Zceil mx) - mx < mx - IZR (Zfloor mx))%R
(Rabs (IZR (Zceil mx) - mx) <= IZR (Zceil mx) - mx)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
Rnd_N_pt generic_format x (round Znearest x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) <> mx
H:(Rabs (mx - IZR (Zfloor mx)) <= mx - IZR (Zfloor mx))%R
Hm':(mx - IZR (Zfloor mx))%R = (IZR (Zceil mx) - mx)%R

(Rabs (mx - IZR (Zfloor mx)) <= mx - IZR (Zfloor mx))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) <> mx
H:(Rabs (mx - IZR (Zfloor mx)) <= mx - IZR (Zfloor mx))%R
Hm':(mx - IZR (Zfloor mx))%R = (IZR (Zceil mx) - mx)%R
(Rabs (IZR (Zfloor mx) - mx) <= mx - IZR (Zfloor mx))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) <> mx
H:(Rabs (mx - IZR (Zfloor mx)) <= mx - IZR (Zfloor mx))%R
Hm':(IZR (Zceil mx) - mx < mx - IZR (Zfloor mx))%R
(Rabs (IZR (Zceil mx) - mx) <= IZR (Zceil mx) - mx)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
Rnd_N_pt generic_format x (round Znearest x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) <> mx
H:(Rabs (mx - IZR (Zfloor mx)) <= mx - IZR (Zfloor mx))%R
Hm':(mx - IZR (Zfloor mx))%R = (IZR (Zceil mx) - mx)%R

(Rabs (IZR (Zfloor mx) - mx) <= mx - IZR (Zfloor mx))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) <> mx
H:(Rabs (mx - IZR (Zfloor mx)) <= mx - IZR (Zfloor mx))%R
Hm':(IZR (Zceil mx) - mx < mx - IZR (Zfloor mx))%R
(Rabs (IZR (Zceil mx) - mx) <= IZR (Zceil mx) - mx)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
Rnd_N_pt generic_format x (round Znearest x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) <> mx
H:(Rabs (mx - IZR (Zfloor mx)) <= mx - IZR (Zfloor mx))%R
Hm':(IZR (Zceil mx) - mx < mx - IZR (Zfloor mx))%R

(Rabs (IZR (Zceil mx) - mx) <= IZR (Zceil mx) - mx)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
Rnd_N_pt generic_format x (round Znearest x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) <> mx
H:(Rabs (mx - IZR (Zfloor mx)) <= mx - IZR (Zfloor mx))%R
Hm':(IZR (Zceil mx) - mx < mx - IZR (Zfloor mx))%R

(IZR (Zceil mx) - mx <= IZR (Zceil mx) - mx)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) <> mx
H:(Rabs (mx - IZR (Zfloor mx)) <= mx - IZR (Zfloor mx))%R
Hm':(IZR (Zceil mx) - mx < mx - IZR (Zfloor mx))%R
(0 <= IZR (Zceil mx) - mx)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
Rnd_N_pt generic_format x (round Znearest x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) <> mx
H:(Rabs (mx - IZR (Zfloor mx)) <= mx - IZR (Zfloor mx))%R
Hm':(IZR (Zceil mx) - mx < mx - IZR (Zfloor mx))%R

(0 <= IZR (Zceil mx) - mx)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
Rnd_N_pt generic_format x (round Znearest x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
Hm:IZR (Zfloor mx) <> mx
H:(Rabs (mx - IZR (Zfloor mx)) <= mx - IZR (Zfloor mx))%R
Hm':(IZR (Zceil mx) - mx < mx - IZR (Zfloor mx))%R

(mx <= IZR (Zceil mx))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
Rnd_N_pt generic_format x (round Znearest x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R

Rnd_N_pt generic_format x (round Znearest x)
(* . *)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R

generic_format (round Znearest x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
Rnd_DN_pt generic_format x d
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
Rnd_UP_pt generic_format x u
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
(Rabs (round Znearest x - x) <= x - d)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
(Rabs (round Znearest x - x) <= u - x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R

Valid_rnd Znearest
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
Rnd_DN_pt generic_format x d
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
Rnd_UP_pt generic_format x u
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
(Rabs (round Znearest x - x) <= x - d)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
(Rabs (round Znearest x - x) <= u - x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R

Rnd_DN_pt generic_format x d
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
Rnd_UP_pt generic_format x u
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
(Rabs (round Znearest x - x) <= x - d)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
(Rabs (round Znearest x - x) <= u - x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R

Rnd_UP_pt generic_format x u
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
(Rabs (round Znearest x - x) <= x - d)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
(Rabs (round Znearest x - x) <= u - x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R

(Rabs (round Znearest x - x) <= x - d)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
(Rabs (round Znearest x - x) <= u - x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R

(Rmin (x - d) (u - x) <= x - d)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R
(Rabs (round Znearest x - x) <= u - x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R

(Rabs (round Znearest x - x) <= u - x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
d:=round Zfloor x:R
u:=round Zceil x:R
mx:=scaled_mantissa x:R
bx:=bpow (cexp x):R
H:(Rabs (round Znearest x - x) <= Rmin (x - d) (u - x))%R

(Rmin (x - d) (u - x) <= u - x)%R
apply Rmin_r. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool

forall x : R, (x - round Zfloor x)%R = (round Zceil x - x)%R -> round Znearest x = (if choice (Zfloor (scaled_mantissa x)) then round Zceil x else round Zfloor x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool

forall x : R, (x - round Zfloor x)%R = (round Zceil x - x)%R -> round Znearest x = (if choice (Zfloor (scaled_mantissa x)) then round Zceil x else round Zfloor x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R

(x - round Zfloor x)%R = (round Zceil x - x)%R -> round Znearest x = (if choice (Zfloor (scaled_mantissa x)) then round Zceil x else round Zfloor x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R

(scaled_mantissa x * bpow (cexp x) - round Zfloor x)%R = (round Zceil x - scaled_mantissa x * bpow (cexp x))%R -> round Znearest x = (if choice (Zfloor (scaled_mantissa x)) then round Zceil x else round Zfloor x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R

(scaled_mantissa x * bpow (cexp x) - IZR (Fnum {| Fnum := Zfloor (scaled_mantissa x); Fexp := cexp x |}) * bpow (Fexp {| Fnum := Zfloor (scaled_mantissa x); Fexp := cexp x |}))%R = (IZR (Fnum {| Fnum := Zceil (scaled_mantissa x); Fexp := cexp x |}) * bpow (Fexp {| Fnum := Zceil (scaled_mantissa x); Fexp := cexp x |}) - scaled_mantissa x * bpow (cexp x))%R -> (IZR (Fnum {| Fnum := match Rcompare (scaled_mantissa x - IZR (Zfloor (scaled_mantissa x))) (/ 2) with | Eq => if choice (Zfloor (scaled_mantissa x)) then Zceil (scaled_mantissa x) else Zfloor (scaled_mantissa x) | Lt => Zfloor (scaled_mantissa x) | Gt => Zceil (scaled_mantissa x) end; Fexp := cexp x |}) * bpow (Fexp {| Fnum := match Rcompare (scaled_mantissa x - IZR (Zfloor (scaled_mantissa x))) (/ 2) with | Eq => if choice (Zfloor (scaled_mantissa x)) then Zceil (scaled_mantissa x) else Zfloor (scaled_mantissa x) | Lt => Zfloor (scaled_mantissa x) | Gt => Zceil (scaled_mantissa x) end; Fexp := cexp x |}))%R = (if choice (Zfloor (scaled_mantissa x)) then (IZR (Fnum {| Fnum := Zceil (scaled_mantissa x); Fexp := cexp x |}) * bpow (Fexp {| Fnum := Zceil (scaled_mantissa x); Fexp := cexp x |}))%R else (IZR (Fnum {| Fnum := Zfloor (scaled_mantissa x); Fexp := cexp x |}) * bpow (Fexp {| Fnum := Zfloor (scaled_mantissa x); Fexp := cexp x |}))%R)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R

(scaled_mantissa x * bpow (cexp x) - IZR (Zfloor (scaled_mantissa x)) * bpow (cexp x))%R = (IZR (Zceil (scaled_mantissa x)) * bpow (cexp x) - scaled_mantissa x * bpow (cexp x))%R -> (IZR match Rcompare (scaled_mantissa x - IZR (Zfloor (scaled_mantissa x))) (/ 2) with | Eq => if choice (Zfloor (scaled_mantissa x)) then Zceil (scaled_mantissa x) else Zfloor (scaled_mantissa x) | Lt => Zfloor (scaled_mantissa x) | Gt => Zceil (scaled_mantissa x) end * bpow (cexp x))%R = (if choice (Zfloor (scaled_mantissa x)) then (IZR (Zceil (scaled_mantissa x)) * bpow (cexp x))%R else (IZR (Zfloor (scaled_mantissa x)) * bpow (cexp x))%R)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Fx:IZR (Zfloor (scaled_mantissa x)) = scaled_mantissa x

(scaled_mantissa x * bpow (cexp x) - IZR (Zfloor (scaled_mantissa x)) * bpow (cexp x))%R = (IZR (Zceil (scaled_mantissa x)) * bpow (cexp x) - scaled_mantissa x * bpow (cexp x))%R -> (IZR match Rcompare (scaled_mantissa x - IZR (Zfloor (scaled_mantissa x))) (/ 2) with | Eq => if choice (Zfloor (scaled_mantissa x)) then Zceil (scaled_mantissa x) else Zfloor (scaled_mantissa x) | Lt => Zfloor (scaled_mantissa x) | Gt => Zceil (scaled_mantissa x) end * bpow (cexp x))%R = (if choice (Zfloor (scaled_mantissa x)) then (IZR (Zceil (scaled_mantissa x)) * bpow (cexp x))%R else (IZR (Zfloor (scaled_mantissa x)) * bpow (cexp x))%R)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Fx:IZR (Zfloor (scaled_mantissa x)) <> scaled_mantissa x
(scaled_mantissa x * bpow (cexp x) - IZR (Zfloor (scaled_mantissa x)) * bpow (cexp x))%R = (IZR (Zceil (scaled_mantissa x)) * bpow (cexp x) - scaled_mantissa x * bpow (cexp x))%R -> (IZR match Rcompare (scaled_mantissa x - IZR (Zfloor (scaled_mantissa x))) (/ 2) with | Eq => if choice (Zfloor (scaled_mantissa x)) then Zceil (scaled_mantissa x) else Zfloor (scaled_mantissa x) | Lt => Zfloor (scaled_mantissa x) | Gt => Zceil (scaled_mantissa x) end * bpow (cexp x))%R = (if choice (Zfloor (scaled_mantissa x)) then (IZR (Zceil (scaled_mantissa x)) * bpow (cexp x))%R else (IZR (Zfloor (scaled_mantissa x)) * bpow (cexp x))%R)
(* *)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Fx:IZR (Zfloor (scaled_mantissa x)) = scaled_mantissa x

(IZR match Rcompare (scaled_mantissa x - IZR (Zfloor (scaled_mantissa x))) (/ 2) with | Eq => if choice (Zfloor (scaled_mantissa x)) then Zceil (scaled_mantissa x) else Zfloor (scaled_mantissa x) | Lt => Zfloor (scaled_mantissa x) | Gt => Zceil (scaled_mantissa x) end * bpow (cexp x))%R = (if choice (Zfloor (scaled_mantissa x)) then (IZR (Zceil (scaled_mantissa x)) * bpow (cexp x))%R else (IZR (Zfloor (scaled_mantissa x)) * bpow (cexp x))%R)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Fx:IZR (Zfloor (scaled_mantissa x)) <> scaled_mantissa x
(scaled_mantissa x * bpow (cexp x) - IZR (Zfloor (scaled_mantissa x)) * bpow (cexp x))%R = (IZR (Zceil (scaled_mantissa x)) * bpow (cexp x) - scaled_mantissa x * bpow (cexp x))%R -> (IZR match Rcompare (scaled_mantissa x - IZR (Zfloor (scaled_mantissa x))) (/ 2) with | Eq => if choice (Zfloor (scaled_mantissa x)) then Zceil (scaled_mantissa x) else Zfloor (scaled_mantissa x) | Lt => Zfloor (scaled_mantissa x) | Gt => Zceil (scaled_mantissa x) end * bpow (cexp x))%R = (if choice (Zfloor (scaled_mantissa x)) then (IZR (Zceil (scaled_mantissa x)) * bpow (cexp x))%R else (IZR (Zfloor (scaled_mantissa x)) * bpow (cexp x))%R)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Fx:IZR (Zfloor (scaled_mantissa x)) = scaled_mantissa x

(IZR match Rcompare (IZR (Zfloor (scaled_mantissa x)) - IZR (Zfloor (IZR (Zfloor (scaled_mantissa x))))) (/ 2) with | Eq => if choice (Zfloor (IZR (Zfloor (scaled_mantissa x)))) then Zceil (IZR (Zfloor (scaled_mantissa x))) else Zfloor (IZR (Zfloor (scaled_mantissa x))) | Lt => Zfloor (IZR (Zfloor (scaled_mantissa x))) | Gt => Zceil (IZR (Zfloor (scaled_mantissa x))) end * bpow (cexp x))%R = (if choice (Zfloor (IZR (Zfloor (scaled_mantissa x)))) then (IZR (Zceil (IZR (Zfloor (scaled_mantissa x)))) * bpow (cexp x))%R else (IZR (Zfloor (IZR (Zfloor (scaled_mantissa x)))) * bpow (cexp x))%R)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Fx:IZR (Zfloor (scaled_mantissa x)) <> scaled_mantissa x
(scaled_mantissa x * bpow (cexp x) - IZR (Zfloor (scaled_mantissa x)) * bpow (cexp x))%R = (IZR (Zceil (scaled_mantissa x)) * bpow (cexp x) - scaled_mantissa x * bpow (cexp x))%R -> (IZR match Rcompare (scaled_mantissa x - IZR (Zfloor (scaled_mantissa x))) (/ 2) with | Eq => if choice (Zfloor (scaled_mantissa x)) then Zceil (scaled_mantissa x) else Zfloor (scaled_mantissa x) | Lt => Zfloor (scaled_mantissa x) | Gt => Zceil (scaled_mantissa x) end * bpow (cexp x))%R = (if choice (Zfloor (scaled_mantissa x)) then (IZR (Zceil (scaled_mantissa x)) * bpow (cexp x))%R else (IZR (Zfloor (scaled_mantissa x)) * bpow (cexp x))%R)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Fx:IZR (Zfloor (scaled_mantissa x)) = scaled_mantissa x

(IZR match Rcompare (IZR (Zfloor (scaled_mantissa x)) - IZR (Zfloor (scaled_mantissa x))) (/ 2) with | Eq => if choice (Zfloor (scaled_mantissa x)) then Zfloor (scaled_mantissa x) else Zfloor (scaled_mantissa x) | _ => Zfloor (scaled_mantissa x) end * bpow (cexp x))%R = (if choice (Zfloor (scaled_mantissa x)) then (IZR (Zfloor (scaled_mantissa x)) * bpow (cexp x))%R else (IZR (Zfloor (scaled_mantissa x)) * bpow (cexp x))%R)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Fx:IZR (Zfloor (scaled_mantissa x)) <> scaled_mantissa x
(scaled_mantissa x * bpow (cexp x) - IZR (Zfloor (scaled_mantissa x)) * bpow (cexp x))%R = (IZR (Zceil (scaled_mantissa x)) * bpow (cexp x) - scaled_mantissa x * bpow (cexp x))%R -> (IZR match Rcompare (scaled_mantissa x - IZR (Zfloor (scaled_mantissa x))) (/ 2) with | Eq => if choice (Zfloor (scaled_mantissa x)) then Zceil (scaled_mantissa x) else Zfloor (scaled_mantissa x) | Lt => Zfloor (scaled_mantissa x) | Gt => Zceil (scaled_mantissa x) end * bpow (cexp x))%R = (if choice (Zfloor (scaled_mantissa x)) then (IZR (Zceil (scaled_mantissa x)) * bpow (cexp x))%R else (IZR (Zfloor (scaled_mantissa x)) * bpow (cexp x))%R)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Fx:IZR (Zfloor (scaled_mantissa x)) = scaled_mantissa x
m:=Zfloor (scaled_mantissa x):Z

(IZR match Rcompare (IZR m - IZR m) (/ 2) with | Eq => if choice m then m else m | _ => m end * bpow (cexp x))%R = (if choice m then (IZR m * bpow (cexp x))%R else (IZR m * bpow (cexp x))%R)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Fx:IZR (Zfloor (scaled_mantissa x)) <> scaled_mantissa x
(scaled_mantissa x * bpow (cexp x) - IZR (Zfloor (scaled_mantissa x)) * bpow (cexp x))%R = (IZR (Zceil (scaled_mantissa x)) * bpow (cexp x) - scaled_mantissa x * bpow (cexp x))%R -> (IZR match Rcompare (scaled_mantissa x - IZR (Zfloor (scaled_mantissa x))) (/ 2) with | Eq => if choice (Zfloor (scaled_mantissa x)) then Zceil (scaled_mantissa x) else Zfloor (scaled_mantissa x) | Lt => Zfloor (scaled_mantissa x) | Gt => Zceil (scaled_mantissa x) end * bpow (cexp x))%R = (if choice (Zfloor (scaled_mantissa x)) then (IZR (Zceil (scaled_mantissa x)) * bpow (cexp x))%R else (IZR (Zfloor (scaled_mantissa x)) * bpow (cexp x))%R)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Fx:IZR (Zfloor (scaled_mantissa x)) <> scaled_mantissa x

(scaled_mantissa x * bpow (cexp x) - IZR (Zfloor (scaled_mantissa x)) * bpow (cexp x))%R = (IZR (Zceil (scaled_mantissa x)) * bpow (cexp x) - scaled_mantissa x * bpow (cexp x))%R -> (IZR match Rcompare (scaled_mantissa x - IZR (Zfloor (scaled_mantissa x))) (/ 2) with | Eq => if choice (Zfloor (scaled_mantissa x)) then Zceil (scaled_mantissa x) else Zfloor (scaled_mantissa x) | Lt => Zfloor (scaled_mantissa x) | Gt => Zceil (scaled_mantissa x) end * bpow (cexp x))%R = (if choice (Zfloor (scaled_mantissa x)) then (IZR (Zceil (scaled_mantissa x)) * bpow (cexp x))%R else (IZR (Zfloor (scaled_mantissa x)) * bpow (cexp x))%R)
(* *)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Fx:IZR (Zfloor (scaled_mantissa x)) <> scaled_mantissa x
H:(scaled_mantissa x * bpow (cexp x) - IZR (Zfloor (scaled_mantissa x)) * bpow (cexp x))%R = (IZR (Zceil (scaled_mantissa x)) * bpow (cexp x) - scaled_mantissa x * bpow (cexp x))%R

(IZR match Rcompare (scaled_mantissa x - IZR (Zfloor (scaled_mantissa x))) (/ 2) with | Eq => if choice (Zfloor (scaled_mantissa x)) then Zceil (scaled_mantissa x) else Zfloor (scaled_mantissa x) | Lt => Zfloor (scaled_mantissa x) | Gt => Zceil (scaled_mantissa x) end * bpow (cexp x))%R = (if choice (Zfloor (scaled_mantissa x)) then (IZR (Zceil (scaled_mantissa x)) * bpow (cexp x))%R else (IZR (Zfloor (scaled_mantissa x)) * bpow (cexp x))%R)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Fx:IZR (Zfloor (scaled_mantissa x)) <> scaled_mantissa x
H:(scaled_mantissa x * bpow (cexp x) - IZR (Zfloor (scaled_mantissa x)) * bpow (cexp x))%R = (IZR (Zceil (scaled_mantissa x)) * bpow (cexp x) - scaled_mantissa x * bpow (cexp x))%R

(IZR match Rcompare (scaled_mantissa x - IZR (Zfloor (scaled_mantissa x))) (IZR (Zceil (scaled_mantissa x)) - scaled_mantissa x) with | Eq => if choice (Zfloor (scaled_mantissa x)) then Zceil (scaled_mantissa x) else Zfloor (scaled_mantissa x) | Lt => Zfloor (scaled_mantissa x) | Gt => Zceil (scaled_mantissa x) end * bpow (cexp x))%R = (if choice (Zfloor (scaled_mantissa x)) then (IZR (Zceil (scaled_mantissa x)) * bpow (cexp x))%R else (IZR (Zfloor (scaled_mantissa x)) * bpow (cexp x))%R)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Fx:IZR (Zfloor (scaled_mantissa x)) <> scaled_mantissa x
H:(scaled_mantissa x * bpow (cexp x) - IZR (Zfloor (scaled_mantissa x)) * bpow (cexp x))%R = (IZR (Zceil (scaled_mantissa x)) * bpow (cexp x) - scaled_mantissa x * bpow (cexp x))%R

(IZR (if choice (Zfloor (scaled_mantissa x)) then Zceil (scaled_mantissa x) else Zfloor (scaled_mantissa x)) * bpow (cexp x))%R = (if choice (Zfloor (scaled_mantissa x)) then (IZR (Zceil (scaled_mantissa x)) * bpow (cexp x))%R else (IZR (Zfloor (scaled_mantissa x)) * bpow (cexp x))%R)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Fx:IZR (Zfloor (scaled_mantissa x)) <> scaled_mantissa x
H:(scaled_mantissa x * bpow (cexp x) - IZR (Zfloor (scaled_mantissa x)) * bpow (cexp x))%R = (IZR (Zceil (scaled_mantissa x)) * bpow (cexp x) - scaled_mantissa x * bpow (cexp x))%R
(scaled_mantissa x - IZR (Zfloor (scaled_mantissa x)))%R = (IZR (Zceil (scaled_mantissa x)) - scaled_mantissa x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Fx:IZR (Zfloor (scaled_mantissa x)) <> scaled_mantissa x
H:(scaled_mantissa x * bpow (cexp x) - IZR (Zfloor (scaled_mantissa x)) * bpow (cexp x))%R = (IZR (Zceil (scaled_mantissa x)) * bpow (cexp x) - scaled_mantissa x * bpow (cexp x))%R

(scaled_mantissa x - IZR (Zfloor (scaled_mantissa x)))%R = (IZR (Zceil (scaled_mantissa x)) - scaled_mantissa x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Fx:IZR (Zfloor (scaled_mantissa x)) <> scaled_mantissa x
H:(scaled_mantissa x * bpow (cexp x) - IZR (Zfloor (scaled_mantissa x)) * bpow (cexp x))%R = (IZR (Zceil (scaled_mantissa x)) * bpow (cexp x) - scaled_mantissa x * bpow (cexp x))%R

((scaled_mantissa x - IZR (Zfloor (scaled_mantissa x))) * bpow (cexp x))%R = ((IZR (Zceil (scaled_mantissa x)) - scaled_mantissa x) * bpow (cexp x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Fx:IZR (Zfloor (scaled_mantissa x)) <> scaled_mantissa x
H:(scaled_mantissa x * bpow (cexp x) - IZR (Zfloor (scaled_mantissa x)) * bpow (cexp x))%R = (IZR (Zceil (scaled_mantissa x)) * bpow (cexp x) - scaled_mantissa x * bpow (cexp x))%R
bpow (cexp x) <> 0%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Fx:IZR (Zfloor (scaled_mantissa x)) <> scaled_mantissa x
H:(scaled_mantissa x * bpow (cexp x) - IZR (Zfloor (scaled_mantissa x)) * bpow (cexp x))%R = (IZR (Zceil (scaled_mantissa x)) * bpow (cexp x) - scaled_mantissa x * bpow (cexp x))%R

bpow (cexp x) <> 0%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Fx:IZR (Zfloor (scaled_mantissa x)) <> scaled_mantissa x
H:(scaled_mantissa x * bpow (cexp x) - IZR (Zfloor (scaled_mantissa x)) * bpow (cexp x))%R = (IZR (Zceil (scaled_mantissa x)) * bpow (cexp x) - scaled_mantissa x * bpow (cexp x))%R

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

forall (x : R) (ex : Z), (bpow (ex - 1) <= x < bpow ex)%R -> (ex < fexp ex)%Z -> round Znearest x = 0%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool

forall (x : R) (ex : Z), (bpow (ex - 1) <= x < bpow ex)%R -> (ex < fexp ex)%Z -> round Znearest x = 0%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Hf:(ex < fexp ex)%Z

round Znearest x = 0%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Hf:(ex < fexp ex)%Z

(IZR (Znearest (x * bpow (- fexp (mag beta x)))) * bpow (fexp (mag beta x)))%R = 0%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Hf:(ex < fexp ex)%Z

(IZR (Znearest (x * bpow (- fexp (mag beta x)))) * bpow (fexp (mag beta x)) * bpow (- fexp (mag beta x)))%R = (0 * bpow (- fexp (mag beta x)))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Hf:(ex < fexp ex)%Z

(IZR (Znearest (x * bpow (- fexp (mag beta x)))) * bpow (fexp (mag beta x) + - fexp (mag beta x)))%R = 0%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Hf:(ex < fexp ex)%Z

IZR (Znearest (x * bpow (- fexp (mag beta x)))) = 0%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Hf:(ex < fexp ex)%Z

Znearest (x * bpow (- fexp (mag beta x))) = 0%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Hf:(ex < fexp ex)%Z

(Rabs (x * bpow (- fexp (mag beta x)) - 0) < / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Hf:(ex < fexp ex)%Z

(Rabs (x * bpow (- fexp (mag beta x))) < / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Hf:(ex < fexp ex)%Z

(x >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Hf:(ex < fexp ex)%Z
H:(x >= 0)%R
(Rabs (x * bpow (- fexp (mag beta x))) < / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Hf:(ex < fexp ex)%Z

(x >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Hf:(ex < fexp ex)%Z

(0 <= bpow (ex - 1))%R
now apply bpow_ge_0.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Hf:(ex < fexp ex)%Z
H:(x >= 0)%R

(Rabs (x * bpow (- fexp (mag beta x))) < / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Hf:(ex < fexp ex)%Z
H:(x >= 0)%R

(x * bpow (- fexp (mag beta x)) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Hf:(ex < fexp ex)%Z
H:(x >= 0)%R
H':(x * bpow (- fexp (mag beta x)) >= 0)%R
(Rabs (x * bpow (- fexp (mag beta x))) < / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Hf:(ex < fexp ex)%Z
H:(x >= 0)%R

(x * bpow (- fexp (mag beta x)) >= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Hf:(ex < fexp ex)%Z
H:(x >= 0)%R

(0 <= x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Hf:(ex < fexp ex)%Z
H:(x >= 0)%R
(0 <= bpow (- fexp (mag beta x)))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Hf:(ex < fexp ex)%Z
H:(x >= 0)%R

(0 <= x)%R
now apply Rge_le.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Hf:(ex < fexp ex)%Z
H:(x >= 0)%R

(0 <= bpow (- fexp (mag beta x)))%R
now apply bpow_ge_0.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Hf:(ex < fexp ex)%Z
H:(x >= 0)%R
H':(x * bpow (- fexp (mag beta x)) >= 0)%R

(Rabs (x * bpow (- fexp (mag beta x))) < / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Hf:(ex < fexp ex)%Z
H:(x >= 0)%R
H':(x * bpow (- fexp (mag beta x)) >= 0)%R

(x * bpow (- fexp (mag beta x)) < / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Hf:(ex < fexp ex)%Z
H:(x >= 0)%R
H':(x * bpow (- fexp (mag beta x)) >= 0)%R

(x * bpow (- fexp (mag beta x)) * bpow (fexp (mag beta x)) < / 2 * bpow (fexp (mag beta x)))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Hf:(ex < fexp ex)%Z
H:(x >= 0)%R
H':(x * bpow (- fexp (mag beta x)) >= 0)%R

(x * bpow (- fexp (mag beta x) + fexp (mag beta x)) < / 2 * bpow (fexp (mag beta x)))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Hf:(ex < fexp ex)%Z
H:(x >= 0)%R
H':(x * bpow (- fexp (mag beta x)) >= 0)%R

(x < / 2 * bpow (fexp (mag beta x)))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Hf:(ex < fexp ex)%Z
H:(x >= 0)%R
H':(x * bpow (- fexp (mag beta x)) >= 0)%R

(bpow ex <= / 2 * bpow (fexp (mag beta x)))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Hf:(ex < fexp ex)%Z
H:(x >= 0)%R
H':(x * bpow (- fexp (mag beta x)) >= 0)%R

(bpow ex <= bpow (fexp (mag beta x) - 1))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Hf:(ex < fexp ex)%Z
H:(x >= 0)%R
H':(x * bpow (- fexp (mag beta x)) >= 0)%R
(bpow (fexp (mag beta x) - 1) <= / 2 * bpow (fexp (mag beta x)))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Hf:(ex < fexp ex)%Z
H:(x >= 0)%R
H':(x * bpow (- fexp (mag beta x)) >= 0)%R

(bpow ex <= bpow (fexp (mag beta x) - 1))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Hf:(ex < fexp ex)%Z
H:(x >= 0)%R
H':(x * bpow (- fexp (mag beta x)) >= 0)%R

(ex <= fexp (mag beta x) - 1)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Hf:(ex < fexp ex)%Z
H:(x >= 0)%R
H':(x * bpow (- fexp (mag beta x)) >= 0)%R

(bpow (ex - 1) <= Rabs x < bpow ex)%R
now rewrite Rabs_right.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Hf:(ex < fexp ex)%Z
H:(x >= 0)%R
H':(x * bpow (- fexp (mag beta x)) >= 0)%R

(bpow (fexp (mag beta x) - 1) <= / 2 * bpow (fexp (mag beta x)))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Hf:(ex < fexp ex)%Z
H:(x >= 0)%R
H':(x * bpow (- fexp (mag beta x)) >= 0)%R

(bpow (fexp (mag beta x)) * bpow (- (1)) <= / 2 * bpow (fexp (mag beta x)))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Hf:(ex < fexp ex)%Z
H:(x >= 0)%R
H':(x * bpow (- fexp (mag beta x)) >= 0)%R

(bpow (- (1)) * bpow (fexp (mag beta x)) <= / 2 * bpow (fexp (mag beta x)))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Hf:(ex < fexp ex)%Z
H:(x >= 0)%R
H':(x * bpow (- fexp (mag beta x)) >= 0)%R

(bpow (- (1)) <= / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Hf:(ex < fexp ex)%Z
H:(x >= 0)%R
H':(x * bpow (- fexp (mag beta x)) >= 0)%R

(/ IZR (beta * 1) <= / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Hf:(ex < fexp ex)%Z
H:(x >= 0)%R
H':(x * bpow (- fexp (mag beta x)) >= 0)%R

(/ IZR beta <= / 2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Hf:(ex < fexp ex)%Z
H:(x >= 0)%R
H':(x * bpow (- fexp (mag beta x)) >= 0)%R

(2 <= IZR beta)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
ex:Z
Hex:(bpow (ex - 1) <= x < bpow ex)%R
Hf:(ex < fexp ex)%Z
H:(x >= 0)%R
H':(x * bpow (- fexp (mag beta x)) >= 0)%R

(2 <= beta)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
ex, beta_val:Z
beta_prop:(2 <=? beta_val)%Z = true
Hex:(Raux.bpow {| radix_val := beta_val; radix_prop := beta_prop |} (ex - 1) <= x < Raux.bpow {| radix_val := beta_val; radix_prop := beta_prop |} ex)%R
Hf:(ex < fexp ex)%Z
H:(x >= 0)%R
H':(x * Raux.bpow {| radix_val := beta_val; radix_prop := beta_prop |} (- fexp (mag {| radix_val := beta_val; radix_prop := beta_prop |} x)) >= 0)%R

(2 <= {| radix_val := beta_val; radix_prop := beta_prop |})%Z
now apply Zle_bool_imp_le. Qed. End Znearest. Section rndNA. Instance valid_rnd_NA : Valid_rnd (Znearest (Zle_bool 0)) := valid_rnd_N _.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, Rnd_NA_pt generic_format x (round (Znearest (Z.leb 0)) x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, Rnd_NA_pt generic_format x (round (Znearest (Z.leb 0)) x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

Rnd_NA_pt generic_format x (round (Znearest (Z.leb 0)) x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

Rnd_N_pt generic_format x (round (Znearest (Z.leb 0)) x) -> Rnd_NA_pt generic_format x (round (Znearest (Z.leb 0)) x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R

Rnd_N_pt generic_format x f -> Rnd_NA_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f

Rnd_NA_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R

Rnd_NA_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_NA_pt generic_format x f
(* *)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R

generic_format 0
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Rnd_N_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
(Rabs x <= Rabs f)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_NA_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R

Rnd_N_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
(Rabs x <= Rabs f)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_NA_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R

(Rabs x <= Rabs f)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_NA_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(0 <= x)%R

(Rabs x <= Rabs f)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R
(Rabs x <= Rabs f)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_NA_pt generic_format x f
(* . *)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(0 <= x)%R

(x <= Rabs f)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R
(Rabs x <= Rabs f)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_NA_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(0 <= x)%R

(x <= f)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(0 <= x)%R
(0 <= f)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R
(Rabs x <= Rabs f)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_NA_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(0 <= x)%R

(x <= round (Znearest (Z.leb 0)) x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(0 <= x)%R
(0 <= f)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R
(Rabs x <= Rabs f)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_NA_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(0 <= x)%R

(x <= (if (0 <=? Zfloor (scaled_mantissa x))%Z then round Zceil x else round Zfloor x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(0 <= x)%R
(0 <= f)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R
(Rabs x <= Rabs f)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_NA_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(0 <= x)%R

(x <= round Zceil x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(0 <= x)%R
(0 <= Zfloor (scaled_mantissa x))%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(0 <= x)%R
(0 <= f)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R
(Rabs x <= Rabs f)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_NA_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(0 <= x)%R

(0 <= Zfloor (scaled_mantissa x))%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(0 <= x)%R
(0 <= f)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R
(Rabs x <= Rabs f)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_NA_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(0 <= x)%R

(0 <= scaled_mantissa x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(0 <= x)%R
(0 <= f)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R
(Rabs x <= Rabs f)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_NA_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(0 <= x)%R

(0 <= bpow (- cexp x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(0 <= x)%R
(0 <= f)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R
(Rabs x <= Rabs f)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_NA_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(0 <= x)%R

(0 <= f)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R
(Rabs x <= Rabs f)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_NA_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(0 <= x)%R

generic_format 0
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R
(Rabs x <= Rabs f)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_NA_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R

(Rabs x <= Rabs f)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_NA_pt generic_format x f
(* . *)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R

(- x <= Rabs f)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_NA_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R

(- x <= - f)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R
(f <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_NA_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R

(f <= x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R
(f <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_NA_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R

(round (Znearest (Z.leb 0)) x <= x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R
(f <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_NA_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R

((if (0 <=? Zfloor (scaled_mantissa x))%Z then round Zceil x else round Zfloor x) <= x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R
(f <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_NA_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R

(round Zfloor x <= x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R
(Zfloor (scaled_mantissa x) < 0)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R
(f <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_NA_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R

(Zfloor (scaled_mantissa x) < 0)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R
(f <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_NA_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R

(IZR (Zfloor (scaled_mantissa x)) < 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R
(f <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_NA_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R

(IZR (Zfloor (scaled_mantissa x)) <= scaled_mantissa x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R
(scaled_mantissa x < 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R
(f <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_NA_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R

(scaled_mantissa x < 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R
(f <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_NA_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R

(scaled_mantissa x < 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R
(f <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_NA_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R

(scaled_mantissa x < 0 * bpow (- cexp x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R
(f <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_NA_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R

(0 < bpow (- cexp x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R
(f <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_NA_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R

(f <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_NA_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R

generic_format 0
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R
(x <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_NA_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R

(x <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_NA_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R

Rnd_NA_pt generic_format x f
(* *)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R

Rnd_N_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
forall f2 : R, Rnd_N_pt generic_format x f2 -> (Rabs f2 <= Rabs f)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R

forall f2 : R, Rnd_N_pt generic_format x f2 -> (Rabs f2 <= Rabs f)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
g:R
Rxg:Rnd_N_pt generic_format x g

(Rabs g <= Rabs f)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
g:R
Rxg:Rnd_N_pt generic_format x g

(Rabs g <= Rabs g)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
g:R
Rxg:Rnd_N_pt generic_format x g
Rnd_DN_pt generic_format x (round Zfloor x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
g:R
Rxg:Rnd_N_pt generic_format x g
Rnd_UP_pt generic_format x (round Zceil x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
g:R
Rxg:Rnd_N_pt generic_format x g

Rnd_DN_pt generic_format x (round Zfloor x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
g:R
Rxg:Rnd_N_pt generic_format x g
Rnd_UP_pt generic_format x (round Zceil x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round (Znearest (Z.leb 0)) x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
g:R
Rxg:Rnd_N_pt generic_format x g

Rnd_UP_pt generic_format x (round Zceil x)
apply round_UP_pt. Qed. End rndNA. Notation Znearest0 := (Znearest (fun x => (Zlt_bool x 0))). Section rndN0. Instance valid_rnd_N0 : Valid_rnd Znearest0 := valid_rnd_N _.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, Rnd_N0_pt generic_format x (round Znearest0 x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, Rnd_N0_pt generic_format x (round Znearest0 x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

Rnd_N0_pt generic_format x (round Znearest0 x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

Rnd_N_pt generic_format x (round Znearest0 x) -> Rnd_N0_pt generic_format x (round Znearest0 x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R

Rnd_N_pt generic_format x f -> Rnd_N0_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f

Rnd_N0_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R

Rnd_N0_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_N0_pt generic_format x f
(* *)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R

generic_format 0
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Rnd_N_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
(Rabs f <= Rabs x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_N0_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R

Rnd_N_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
(Rabs f <= Rabs x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_N0_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R

(Rabs f <= Rabs x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_N0_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(0 <= x)%R

(Rabs f <= Rabs x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R
(Rabs f <= Rabs x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_N0_pt generic_format x f
(* . *)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(0 <= x)%R

(Rabs f <= x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R
(Rabs f <= Rabs x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_N0_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(0 <= x)%R

(f <= x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(0 <= x)%R
(0 <= f)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R
(Rabs f <= Rabs x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_N0_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(0 <= x)%R

(round Znearest0 x <= x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(0 <= x)%R
(0 <= f)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R
(Rabs f <= Rabs x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_N0_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(0 <= x)%R

((if (Zfloor (scaled_mantissa x) <? 0)%Z then round Zceil x else round Zfloor x) <= x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(0 <= x)%R
(0 <= f)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R
(Rabs f <= Rabs x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_N0_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(0 <= x)%R

(round Zfloor x <= x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(0 <= x)%R
(0 <= Zfloor (scaled_mantissa x))%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(0 <= x)%R
(0 <= f)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R
(Rabs f <= Rabs x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_N0_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(0 <= x)%R

(0 <= Zfloor (scaled_mantissa x))%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(0 <= x)%R
(0 <= f)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R
(Rabs f <= Rabs x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_N0_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(0 <= x)%R

(0 <= scaled_mantissa x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(0 <= x)%R
(0 <= f)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R
(Rabs f <= Rabs x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_N0_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(0 <= x)%R

(0 <= bpow (- cexp x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(0 <= x)%R
(0 <= f)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R
(Rabs f <= Rabs x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_N0_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(0 <= x)%R

(0 <= f)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R
(Rabs f <= Rabs x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_N0_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(0 <= x)%R

generic_format 0
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R
(Rabs f <= Rabs x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_N0_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R

(Rabs f <= Rabs x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_N0_pt generic_format x f
(* . *)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R

(Rabs f <= - x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_N0_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R

(- f <= - x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R
(f <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_N0_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R

(x <= f)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R
(f <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_N0_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R

(x <= round Znearest0 x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R
(f <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_N0_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R

(x <= (if (Zfloor (scaled_mantissa x) <? 0)%Z then round Zceil x else round Zfloor x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R
(f <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_N0_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R

(x <= round Zceil x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R
(Zfloor (scaled_mantissa x) < 0)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R
(f <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_N0_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R

(Zfloor (scaled_mantissa x) < 0)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R
(f <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_N0_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R

(IZR (Zfloor (scaled_mantissa x)) < 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R
(f <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_N0_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R

(IZR (Zfloor (scaled_mantissa x)) <= scaled_mantissa x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R
(scaled_mantissa x < 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R
(f <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_N0_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R

(scaled_mantissa x < 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R
(f <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_N0_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R

(scaled_mantissa x < 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R
(f <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_N0_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R

(scaled_mantissa x < 0 * bpow (- cexp x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R
(f <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_N0_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R

(0 < bpow (- cexp x))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R
(f <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_N0_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R

(f <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_N0_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R

generic_format 0
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R
(x <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_N0_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R = (round Zceil x - x)%R
Hx:(x < 0)%R

(x <= 0)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
Rnd_N0_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R

Rnd_N0_pt generic_format x f
(* *)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R

Rnd_N_pt generic_format x f
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
forall f2 : R, Rnd_N_pt generic_format x f2 -> (Rabs f <= Rabs f2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R

forall f2 : R, Rnd_N_pt generic_format x f2 -> (Rabs f <= Rabs f2)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
g:R
Rxg:Rnd_N_pt generic_format x g

(Rabs f <= Rabs g)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
g:R
Rxg:Rnd_N_pt generic_format x g

(Rabs g <= Rabs g)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
g:R
Rxg:Rnd_N_pt generic_format x g
Rnd_DN_pt generic_format x (round Zfloor x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
g:R
Rxg:Rnd_N_pt generic_format x g
Rnd_UP_pt generic_format x (round Zceil x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
g:R
Rxg:Rnd_N_pt generic_format x g

Rnd_DN_pt generic_format x (round Zfloor x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
g:R
Rxg:Rnd_N_pt generic_format x g
Rnd_UP_pt generic_format x (round Zceil x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
f:=round Znearest0 x:R
Rxf:Rnd_N_pt generic_format x f
Hm:(x - round Zfloor x)%R <> (round Zceil x - x)%R
g:R
Rxg:Rnd_N_pt generic_format x g

Rnd_UP_pt generic_format x (round Zceil x)
apply round_UP_pt; easy. Qed. End rndN0. Section rndN_opp.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall (choice : Z -> bool) (x : R), Znearest choice (- x) = (- Znearest (fun t : Z => negb (choice (- (t + 1)))) x)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall (choice : Z -> bool) (x : R), Znearest choice (- x) = (- Znearest (fun t : Z => negb (choice (- (t + 1)))) x)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R

Znearest choice (- x) = (- Znearest (fun t : Z => negb (choice (- (t + 1)))) x)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:IZR (Zfloor x) = x

Znearest choice (- x) = (- Znearest (fun t : Z => negb (choice (- (t + 1)))) x)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:IZR (Zfloor x) <> x
Znearest choice (- x) = (- Znearest (fun t : Z => negb (choice (- (t + 1)))) x)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:IZR (Zfloor x) = x

Znearest choice (- IZR (Zfloor x)) = (- Znearest (fun t : Z => negb (choice (- (t + 1)))) (IZR (Zfloor x)))%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:IZR (Zfloor x) <> x
Znearest choice (- x) = (- Znearest (fun t : Z => negb (choice (- (t + 1)))) x)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:IZR (Zfloor x) = x

Znearest choice (IZR (- Zfloor x)) = (- Znearest (fun t : Z => negb (choice (- (t + 1)))) (IZR (Zfloor x)))%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:IZR (Zfloor x) <> x
Znearest choice (- x) = (- Znearest (fun t : Z => negb (choice (- (t + 1)))) x)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:IZR (Zfloor x) <> x

Znearest choice (- x) = (- Znearest (fun t : Z => negb (choice (- (t + 1)))) x)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:IZR (Zfloor x) <> x

match Rcompare (- x - IZR (Zfloor (- x))) (/ 2) with | Eq => if choice (Zfloor (- x)) then Zceil (- x) else Zfloor (- x) | Lt => Zfloor (- x) | Gt => Zceil (- x) end = (- match Rcompare (x - IZR (Zfloor x)) (/ 2) with | Eq => if negb (choice (- (Zfloor x + 1))) then Zceil x else Zfloor x | Lt => Zfloor x | Gt => Zceil x end)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:IZR (Zfloor x) <> x

match Rcompare (IZR (Zceil x) - x) (/ 2) with | Eq => if choice (Zfloor (- x)) then Zceil (- x) else Zfloor (- x) | Lt => Zfloor (- x) | Gt => Zceil (- x) end = (- match Rcompare (x - IZR (Zfloor x)) (/ 2) with | Eq => if negb (choice (- (Zfloor x + 1))) then Zceil x else Zfloor x | Lt => Zfloor x | Gt => Zceil x end)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:IZR (Zfloor x) <> x
(IZR (Zceil x) - x)%R = (- x - IZR (Zfloor (- x)))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:IZR (Zfloor x) <> x

match Rcompare (IZR (Zceil x) - x) (x - IZR (Zfloor x)) with | Eq => if choice (Zfloor (- x)) then Zceil (- x) else Zfloor (- x) | Lt => Zfloor (- x) | Gt => Zceil (- x) end = (- match Rcompare (x - IZR (Zfloor x)) (/ 2) with | Eq => if negb (choice (- (Zfloor x + 1))) then Zceil x else Zfloor x | Lt => Zfloor x | Gt => Zceil x end)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:IZR (Zfloor x) <> x
(IZR (Zceil x) - x)%R = (- x - IZR (Zfloor (- x)))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:IZR (Zfloor x) <> x

match Rcompare (IZR (Zceil x) - x) (x - IZR (Zfloor x)) with | Eq => if choice (Zfloor (- x)) then Zceil (- x) else Zfloor (- x) | Lt => Zfloor (- x) | Gt => Zceil (- x) end = (- match Rcompare (x - IZR (Zfloor x)) (IZR (Zceil x) - x) with | Eq => if negb (choice (- (Zfloor x + 1))) then Zceil x else Zfloor x | Lt => Zfloor x | Gt => Zceil x end)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:IZR (Zfloor x) <> x
(IZR (Zceil x) - x)%R = (- x - IZR (Zfloor (- x)))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:IZR (Zfloor x) <> x

match CompOpp (Rcompare (x - IZR (Zfloor x)) (IZR (Zceil x) - x)) with | Eq => if choice (Zfloor (- x)) then Zceil (- x) else Zfloor (- x) | Lt => Zfloor (- x) | Gt => Zceil (- x) end = (- match Rcompare (x - IZR (Zfloor x)) (IZR (Zceil x) - x) with | Eq => if negb (choice (- (Zfloor x + 1))) then Zceil x else Zfloor x | Lt => Zfloor x | Gt => Zceil x end)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:IZR (Zfloor x) <> x
(IZR (Zceil x) - x)%R = (- x - IZR (Zfloor (- x)))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:IZR (Zfloor x) <> x

match CompOpp (Rcompare (x - IZR (Zfloor x)) (IZR (Zceil x) - x)) with | Eq => if choice (Zfloor (- x)) then Zceil (- x) else Zfloor (- x) | Lt => Zfloor (- x) | Gt => Zceil (- x) end = (- match Rcompare (x - IZR (Zfloor x)) (IZR (Zceil x) - x) with | Eq => if negb (choice (- Zceil x)) then Zceil x else Zfloor x | Lt => Zfloor x | Gt => Zceil x end)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:IZR (Zfloor x) <> x
(IZR (Zceil x) - x)%R = (- x - IZR (Zfloor (- x)))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:IZR (Zfloor x) <> x

match CompOpp (Rcompare (x - IZR (Zfloor x)) (IZR (- Zfloor (- x)) - x)) with | Eq => if choice (Zfloor (- x)) then (- Zfloor (- - x))%Z else Zfloor (- x) | Lt => Zfloor (- x) | Gt => (- Zfloor (- - x))%Z end = (- match Rcompare (x - IZR (Zfloor x)) (IZR (- Zfloor (- x)) - x) with | Eq => if negb (choice (- - Zfloor (- x))) then - Zfloor (- x) else Zfloor x | Lt => Zfloor x | Gt => - Zfloor (- x) end)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:IZR (Zfloor x) <> x
(IZR (Zceil x) - x)%R = (- x - IZR (Zfloor (- x)))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:IZR (Zfloor x) <> x

match CompOpp (Rcompare (x - IZR (Zfloor x)) (IZR (- Zfloor (- x)) - x)) with | Eq => if choice (Zfloor (- x)) then (- Zfloor x)%Z else Zfloor (- x) | Lt => Zfloor (- x) | Gt => (- Zfloor x)%Z end = (- match Rcompare (x - IZR (Zfloor x)) (IZR (- Zfloor (- x)) - x) with | Eq => if negb (choice (- - Zfloor (- x))) then - Zfloor (- x) else Zfloor x | Lt => Zfloor x | Gt => - Zfloor (- x) end)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:IZR (Zfloor x) <> x
(IZR (Zceil x) - x)%R = (- x - IZR (Zfloor (- x)))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:IZR (Zfloor x) <> x

(if choice (Zfloor (- x)) then (- Zfloor x)%Z else Zfloor (- x)) = (- (if negb (choice (- - Zfloor (- x))) then - Zfloor (- x) else Zfloor x))%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:IZR (Zfloor x) <> x
Zfloor (- x) = (- - Zfloor (- x))%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:IZR (Zfloor x) <> x
(IZR (Zceil x) - x)%R = (- x - IZR (Zfloor (- x)))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:IZR (Zfloor x) <> x

(if choice (Zfloor (- x)) then (- Zfloor x)%Z else Zfloor (- x)) = (- (if negb (choice (Zfloor (- x))) then - Zfloor (- x) else Zfloor x))%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:IZR (Zfloor x) <> x
Zfloor (- x) = (- - Zfloor (- x))%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:IZR (Zfloor x) <> x
(IZR (Zceil x) - x)%R = (- x - IZR (Zfloor (- x)))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:IZR (Zfloor x) <> x

Zfloor (- x) = (- - Zfloor (- x))%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:IZR (Zfloor x) <> x
Zfloor (- x) = (- - Zfloor (- x))%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:IZR (Zfloor x) <> x
(IZR (Zceil x) - x)%R = (- x - IZR (Zfloor (- x)))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:IZR (Zfloor x) <> x

Zfloor (- x) = (- - Zfloor (- x))%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:IZR (Zfloor x) <> x
(IZR (Zceil x) - x)%R = (- x - IZR (Zfloor (- x)))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:IZR (Zfloor x) <> x

(IZR (Zceil x) - x)%R = (- x - IZR (Zfloor (- x)))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:IZR (Zfloor x) <> x

(IZR (- Zfloor (- x)) - x)%R = (- x - IZR (Zfloor (- x)))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
Hx:IZR (Zfloor x) <> x

(- IZR (Zfloor (- x)) - x)%R = (- x - IZR (Zfloor (- x)))%R
apply Rplus_comm. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall (choice : Z -> bool) (x : R), round (Znearest choice) (- x) = (- round (Znearest (fun t : Z => negb (choice (- (t + 1))%Z))) x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall (choice : Z -> bool) (x : R), round (Znearest choice) (- x) = (- round (Znearest (fun t : Z => negb (choice (- (t + 1))%Z))) x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R

round (Znearest choice) (- x) = (- round (Znearest (fun t : Z => negb (choice (- (t + 1))%Z))) x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R

(IZR (Fnum {| Fnum := Znearest choice (scaled_mantissa (- x)); Fexp := cexp (- x) |}) * bpow (Fexp {| Fnum := Znearest choice (scaled_mantissa (- x)); Fexp := cexp (- x) |}))%R = (- (IZR (Fnum {| Fnum := Znearest (fun t : Z => negb (choice (- (t + 1))%Z)) (scaled_mantissa x); Fexp := cexp x |}) * bpow (Fexp {| Fnum := Znearest (fun t : Z => negb (choice (- (t + 1))%Z)) (scaled_mantissa x); Fexp := cexp x |})))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R

(IZR (Znearest choice (scaled_mantissa (- x))) * bpow (cexp (- x)))%R = (- (IZR (Znearest (fun t : Z => negb (choice (- (t + 1))%Z)) (scaled_mantissa x)) * bpow (cexp x)))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R

(IZR (Znearest choice (scaled_mantissa (- x))) * bpow (cexp x))%R = (- (IZR (Znearest (fun t : Z => negb (choice (- (t + 1))%Z)) (scaled_mantissa x)) * bpow (cexp x)))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R

(IZR (Znearest choice (- scaled_mantissa x)) * bpow (cexp x))%R = (- (IZR (Znearest (fun t : Z => negb (choice (- (t + 1))%Z)) (scaled_mantissa x)) * bpow (cexp x)))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R

(IZR (- Znearest (fun t : Z => negb (choice (- (t + 1))%Z)) (scaled_mantissa x)) * bpow (cexp x))%R = (- (IZR (Znearest (fun t : Z => negb (choice (- (t + 1))%Z)) (scaled_mantissa x)) * bpow (cexp x)))%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R

(- IZR (Znearest (fun t : Z => negb (choice (- (t + 1))%Z)) (scaled_mantissa x)) * bpow (cexp x))%R = (- (IZR (Znearest (fun t : Z => negb (choice (- (t + 1))%Z)) (scaled_mantissa x)) * bpow (cexp x)))%R
now rewrite Ropp_mult_distr_l_reverse. Qed.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, round Znearest0 (- x) = (- round Znearest0 x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall x : R, round Znearest0 (- x) = (- round Znearest0 x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

round Znearest0 (- x) = (- round Znearest0 x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

(- round (Znearest (fun t : Z => negb (- (t + 1) <? 0)%Z)) x)%R = (- round Znearest0 x)%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

round (Znearest (fun t : Z => negb (- (t + 1) <? 0)%Z)) x = round Znearest0 x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

forall x0 : R, Znearest (fun t : Z => negb (- (t + 1) <? 0)%Z) x0 = Znearest0 x0
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

Znearest (fun t : Z => negb (- (t + 1) <? 0)%Z) x = Znearest0 x
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R

match Rcompare (x - IZR (Zfloor x)) (/ 2) with | Eq => if negb (- (Zfloor x + 1) <? 0)%Z then Zceil x else Zfloor x | Lt => Zfloor x | Gt => Zceil x end = match Rcompare (x - IZR (Zfloor x)) (/ 2) with | Eq => if (Zfloor x <? 0)%Z then Zceil x else Zfloor x | Lt => Zfloor x | Gt => Zceil x end
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
C:Rcompare (x - IZR (Zfloor x)) (/ 2) = Eq

(if negb (- (Zfloor x + 1) <? 0)%Z then Zceil x else Zfloor x) = (if (Zfloor x <? 0)%Z then Zceil x else Zfloor x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
C:(x - IZR (Zfloor x))%R = (/ 2)%R

(if negb (- (Zfloor x + 1) <? 0)%Z then Zceil x else Zfloor x) = (if (Zfloor x <? 0)%Z then Zceil x else Zfloor x)
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
C:(x - IZR (Zfloor x))%R = (/ 2)%R

negb (- (Zfloor x + 1) <? 0)%Z = (Zfloor x <? 0)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
C:(x - IZR (Zfloor x))%R = (/ 2)%R

(0 <=? - (Zfloor x + 1))%Z = (Zfloor x <? 0)%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
C:(x - IZR (Zfloor x))%R = (/ 2)%R
C':(Zfloor x <? 0)%Z = true

(0 <=? - (Zfloor x + 1))%Z = true
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
C:(x - IZR (Zfloor x))%R = (/ 2)%R
C':(Zfloor x <? 0)%Z = false
(0 <=? - (Zfloor x + 1))%Z = false
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
C:(x - IZR (Zfloor x))%R = (/ 2)%R
C':(Zfloor x < 0)%Z

(0 <=? - (Zfloor x + 1))%Z = true
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
C:(x - IZR (Zfloor x))%R = (/ 2)%R
C':(Zfloor x <? 0)%Z = false
(0 <=? - (Zfloor x + 1))%Z = false
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
C:(x - IZR (Zfloor x))%R = (/ 2)%R
C':(Zfloor x < 0)%Z

(0 <= - (Zfloor x + 1))%Z
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
C:(x - IZR (Zfloor x))%R = (/ 2)%R
C':(Zfloor x <? 0)%Z = false
(0 <=? - (Zfloor x + 1))%Z = false
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
C:(x - IZR (Zfloor x))%R = (/ 2)%R
C':(Zfloor x <? 0)%Z = false

(0 <=? - (Zfloor x + 1))%Z = false
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
C:(x - IZR (Zfloor x))%R = (/ 2)%R
C':(0 <= Zfloor x)%Z

(0 <=? - (Zfloor x + 1))%Z = false
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
x:R
C:(x - IZR (Zfloor x))%R = (/ 2)%R
C':(0 <= Zfloor x)%Z

(- (Zfloor x + 1) < 0)%Z
omega. Qed. End rndN_opp.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall (choice : Z -> bool) (x : R) (ex : Z), (bpow (ex - 1) <= Rabs x < bpow ex)%R -> (ex < fexp ex)%Z -> round (Znearest choice) x = 0%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp

forall (choice : Z -> bool) (x : R) (ex : Z), (bpow (ex - 1) <= Rabs x < bpow ex)%R -> (ex < fexp ex)%Z -> round (Znearest choice) x = 0%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
ex:Z
Hx:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hex:(ex < fexp ex)%Z

round (Znearest choice) x = 0%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
ex:Z
Hx:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hex:(ex < fexp ex)%Z
Px:(0 <= x)%R

round (Znearest choice) x = 0%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
ex:Z
Hx:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hex:(ex < fexp ex)%Z
Nx:(x < 0)%R
round (Znearest choice) x = 0%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
ex:Z
Hx:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hex:(ex < fexp ex)%Z
Px:(0 <= x)%R

round (Znearest choice) x = 0%R
now revert Hex; apply round_N_small_pos; revert Hx; rewrite Rabs_pos_eq.
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
ex:Z
Hx:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hex:(ex < fexp ex)%Z
Nx:(x < 0)%R

round (Znearest choice) x = 0%R
beta:radix
fexp:Z -> Z
valid_exp_:Valid_exp
choice:Z -> bool
x:R
ex:Z
Hx:(bpow (ex - 1) <= Rabs x < bpow ex)%R
Hex:(ex < fexp ex)%Z
Nx:(x < 0)%R

round (Znearest (fun t : Z => negb (choice (- (t + 1))%Z))) (- x) = 0%R
now revert Hex; apply round_N_small_pos; revert Hx; rewrite Rabs_left. Qed. End Format.
Inclusion of a format into an extended format
Section Inclusion.

Variables fexp1 fexp2 : Z -> Z.

Context { valid_exp1 : Valid_exp fexp1 }.
Context { valid_exp2 : Valid_exp fexp2 }.

beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2

forall x : R, (x <> 0%R -> (fexp2 (mag beta x) <= fexp1 (mag beta x))%Z) -> generic_format fexp1 x -> generic_format fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2

forall x : R, (x <> 0%R -> (fexp2 (mag beta x) <= fexp1 (mag beta x))%Z) -> generic_format fexp1 x -> generic_format fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
x:R
He:x <> 0%R -> (fexp2 (mag beta x) <= fexp1 (mag beta x))%Z
Fx:generic_format fexp1 x

generic_format fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
x:R
He:x <> 0%R -> (fexp2 (mag beta x) <= fexp1 (mag beta x))%Z
Fx:generic_format fexp1 x

generic_format fexp2 (F2R {| Fnum := Ztrunc (scaled_mantissa fexp1 x); Fexp := cexp fexp1 x |})
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
x:R
He:x <> 0%R -> (fexp2 (mag beta x) <= fexp1 (mag beta x))%Z
Fx:generic_format fexp1 x

Ztrunc (scaled_mantissa fexp1 x) <> 0%Z -> (cexp fexp2 (F2R {| Fnum := Ztrunc (scaled_mantissa fexp1 x); Fexp := cexp fexp1 x |}) <= cexp fexp1 x)%Z
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
x:R
He:x <> 0%R -> (fexp2 (mag beta x) <= fexp1 (mag beta x))%Z
Fx:generic_format fexp1 x
Zx:Ztrunc (scaled_mantissa fexp1 x) <> 0%Z

(cexp fexp2 (F2R {| Fnum := Ztrunc (scaled_mantissa fexp1 x); Fexp := cexp fexp1 x |}) <= cexp fexp1 x)%Z
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
x:R
He:x <> 0%R -> (fexp2 (mag beta x) <= fexp1 (mag beta x))%Z
Fx:generic_format fexp1 x
Zx:Ztrunc (scaled_mantissa fexp1 x) <> 0%Z

(cexp fexp2 x <= cexp fexp1 x)%Z
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
x:R
He:x <> 0%R -> (fexp2 (mag beta x) <= fexp1 (mag beta x))%Z
Fx:generic_format fexp1 x
Zx:Ztrunc (scaled_mantissa fexp1 x) <> 0%Z

x <> 0%R
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
x:R
He:x <> 0%R -> (fexp2 (mag beta x) <= fexp1 (mag beta x))%Z
Fx:generic_format fexp1 x
Zx:x = 0%R

Ztrunc (scaled_mantissa fexp1 x) = 0%Z
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
x:R
He:x <> 0%R -> (fexp2 (mag beta x) <= fexp1 (mag beta x))%Z
Fx:generic_format fexp1 x
Zx:x = 0%R

Ztrunc 0 = 0%Z
apply Ztrunc_IZR. Qed.
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2

forall e1 e2 : Z, (forall e : Z, (e1 < e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z) -> forall x : R, (bpow e1 <= Rabs x < bpow e2)%R -> generic_format fexp1 x -> generic_format fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2

forall e1 e2 : Z, (forall e : Z, (e1 < e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z) -> forall x : R, (bpow e1 <= Rabs x < bpow e2)%R -> generic_format fexp1 x -> generic_format fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e1, e2:Z
He:forall e : Z, (e1 < e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx1:(bpow e1 <= Rabs x)%R
Hx2:(Rabs x < bpow e2)%R

generic_format fexp1 x -> generic_format fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e1, e2:Z
He:forall e : Z, (e1 < e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx1:(bpow e1 <= Rabs x)%R
Hx2:(Rabs x < bpow e2)%R

x <> 0%R -> (fexp2 (mag beta x) <= fexp1 (mag beta x))%Z
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e1, e2:Z
He:forall e : Z, (e1 < e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx1:(bpow e1 <= Rabs x)%R
Hx2:(Rabs x < bpow e2)%R
Zx:x <> 0%R

(fexp2 (mag beta x) <= fexp1 (mag beta x))%Z
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e1, e2:Z
He:forall e : Z, (e1 < e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx1:(bpow e1 <= Rabs x)%R
Hx2:(Rabs x < bpow e2)%R
Zx:x <> 0%R

(e1 < mag beta x <= e2)%Z
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e1, e2:Z
He:forall e : Z, (e1 < e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx1:(bpow e1 <= Rabs x)%R
Hx2:(Rabs x < bpow e2)%R
Zx:x <> 0%R

(e1 < mag beta x)%Z
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e1, e2:Z
He:forall e : Z, (e1 < e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx1:(bpow e1 <= Rabs x)%R
Hx2:(Rabs x < bpow e2)%R
Zx:x <> 0%R
(mag beta x <= e2)%Z
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e1, e2:Z
He:forall e : Z, (e1 < e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx1:(bpow e1 <= Rabs x)%R
Hx2:(Rabs x < bpow e2)%R
Zx:x <> 0%R

(mag beta x <= e2)%Z
now apply mag_le_bpow. Qed.
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2

forall e : Z, (fexp2 e <= fexp1 e)%Z -> forall x : R, (bpow (e - 1) <= Rabs x <= bpow e)%R -> generic_format fexp1 x -> generic_format fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2

forall e : Z, (fexp2 e <= fexp1 e)%Z -> forall x : R, (bpow (e - 1) <= Rabs x <= bpow e)%R -> generic_format fexp1 x -> generic_format fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e:Z
He:(fexp2 e <= fexp1 e)%Z
x:R
Hx1:(bpow (e - 1) <= Rabs x)%R
Hx2:(Rabs x < bpow e)%R

generic_format fexp1 x -> generic_format fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e:Z
He:(fexp2 e <= fexp1 e)%Z
x:R
Hx1:(bpow (e - 1) <= Rabs x)%R
Hx2:Rabs x = bpow e
generic_format fexp1 x -> generic_format fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e:Z
He:(fexp2 e <= fexp1 e)%Z
x:R
Hx1:(bpow (e - 1) <= Rabs x)%R
Hx2:(Rabs x < bpow e)%R

x <> 0%R -> (fexp2 (mag beta x) <= fexp1 (mag beta x))%Z
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e:Z
He:(fexp2 e <= fexp1 e)%Z
x:R
Hx1:(bpow (e - 1) <= Rabs x)%R
Hx2:Rabs x = bpow e
generic_format fexp1 x -> generic_format fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e:Z
He:(fexp2 e <= fexp1 e)%Z
x:R
Hx1:(bpow (e - 1) <= Rabs x)%R
Hx2:Rabs x = bpow e

generic_format fexp1 x -> generic_format fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e:Z
He:(fexp2 e <= fexp1 e)%Z
x:R
Hx1:(bpow (e - 1) <= Rabs x)%R
Hx2:Rabs x = bpow e
Fx:generic_format fexp1 x

generic_format fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e:Z
He:(fexp2 e <= fexp1 e)%Z
x:R
Hx1:(bpow (e - 1) <= Rabs x)%R
Hx2:Rabs x = bpow e
Fx:generic_format fexp1 x

generic_format fexp2 (Rabs x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e:Z
He:(fexp2 e <= fexp1 e)%Z
x:R
Hx1:(bpow (e - 1) <= Rabs x)%R
Hx2:Rabs x = bpow e
Fx:generic_format fexp1 x

generic_format fexp2 (bpow e)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e:Z
He:(fexp2 e <= fexp1 e)%Z
x:R
Hx1:(bpow (e - 1) <= Rabs x)%R
Hx2:Rabs x = bpow e
Fx:generic_format fexp1 x

(fexp2 e <= e)%Z
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e:Z
He:(fexp2 e <= fexp1 e)%Z
x:R
Hx1:(bpow (e - 1) <= Rabs x)%R
Hx2:Rabs x = bpow e
Fx:generic_format fexp1 x

(fexp1 e <= e)%Z
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e:Z
He:(fexp2 e <= fexp1 e)%Z
x:R
Hx1:(bpow (e - 1) <= Rabs x)%R
Hx2:Rabs x = bpow e
Fx:generic_format fexp1 x

generic_format fexp1 (bpow e)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e:Z
He:(fexp2 e <= fexp1 e)%Z
x:R
Hx1:(bpow (e - 1) <= Rabs x)%R
Hx2:Rabs x = bpow e
Fx:generic_format fexp1 x

generic_format fexp1 (Rabs x)
now apply generic_format_abs. Qed.
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2

forall e1 e2 : Z, (e1 < e2)%Z -> (forall e : Z, (e1 < e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z) -> forall x : R, (bpow e1 <= Rabs x <= bpow e2)%R -> generic_format fexp1 x -> generic_format fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2

forall e1 e2 : Z, (e1 < e2)%Z -> (forall e : Z, (e1 < e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z) -> forall x : R, (bpow e1 <= Rabs x <= bpow e2)%R -> generic_format fexp1 x -> generic_format fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e1, e2:Z
He':(e1 < e2)%Z
He:forall e : Z, (e1 < e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx1:(bpow e1 <= Rabs x)%R
Hx2:(Rabs x < bpow e2)%R

generic_format fexp1 x -> generic_format fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e1, e2:Z
He':(e1 < e2)%Z
He:forall e : Z, (e1 < e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx1:(bpow e1 <= Rabs x)%R
Hx2:Rabs x = bpow e2
generic_format fexp1 x -> generic_format fexp2 x
(* *)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e1, e2:Z
He':(e1 < e2)%Z
He:forall e : Z, (e1 < e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx1:(bpow e1 <= Rabs x)%R
Hx2:(Rabs x < bpow e2)%R

x <> 0%R -> (fexp2 (mag beta x) <= fexp1 (mag beta x))%Z
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e1, e2:Z
He':(e1 < e2)%Z
He:forall e : Z, (e1 < e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx1:(bpow e1 <= Rabs x)%R
Hx2:Rabs x = bpow e2
generic_format fexp1 x -> generic_format fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e1, e2:Z
He':(e1 < e2)%Z
He:forall e : Z, (e1 < e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx1:(bpow e1 <= Rabs x)%R
Hx2:(Rabs x < bpow e2)%R
Zx:x <> 0%R

(fexp2 (mag beta x) <= fexp1 (mag beta x))%Z
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e1, e2:Z
He':(e1 < e2)%Z
He:forall e : Z, (e1 < e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx1:(bpow e1 <= Rabs x)%R
Hx2:Rabs x = bpow e2
generic_format fexp1 x -> generic_format fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e1, e2:Z
He':(e1 < e2)%Z
He:forall e : Z, (e1 < e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx1:(bpow e1 <= Rabs x)%R
Hx2:(Rabs x < bpow e2)%R
Zx:x <> 0%R

(e1 < mag beta x <= e2)%Z
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e1, e2:Z
He':(e1 < e2)%Z
He:forall e : Z, (e1 < e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx1:(bpow e1 <= Rabs x)%R
Hx2:Rabs x = bpow e2
generic_format fexp1 x -> generic_format fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e1, e2:Z
He':(e1 < e2)%Z
He:forall e : Z, (e1 < e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx1:(bpow e1 <= Rabs x)%R
Hx2:(Rabs x < bpow e2)%R
Zx:x <> 0%R

(e1 < mag beta x)%Z
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e1, e2:Z
He':(e1 < e2)%Z
He:forall e : Z, (e1 < e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx1:(bpow e1 <= Rabs x)%R
Hx2:(Rabs x < bpow e2)%R
Zx:x <> 0%R
(mag beta x <= e2)%Z
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e1, e2:Z
He':(e1 < e2)%Z
He:forall e : Z, (e1 < e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx1:(bpow e1 <= Rabs x)%R
Hx2:Rabs x = bpow e2
generic_format fexp1 x -> generic_format fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e1, e2:Z
He':(e1 < e2)%Z
He:forall e : Z, (e1 < e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx1:(bpow e1 <= Rabs x)%R
Hx2:(Rabs x < bpow e2)%R
Zx:x <> 0%R

(mag beta x <= e2)%Z
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e1, e2:Z
He':(e1 < e2)%Z
He:forall e : Z, (e1 < e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx1:(bpow e1 <= Rabs x)%R
Hx2:Rabs x = bpow e2
generic_format fexp1 x -> generic_format fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e1, e2:Z
He':(e1 < e2)%Z
He:forall e : Z, (e1 < e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx1:(bpow e1 <= Rabs x)%R
Hx2:Rabs x = bpow e2

generic_format fexp1 x -> generic_format fexp2 x
(* *)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e1, e2:Z
He':(e1 < e2)%Z
He:forall e : Z, (e1 < e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx1:(bpow e1 <= Rabs x)%R
Hx2:Rabs x = bpow e2

(fexp2 e2 <= fexp1 e2)%Z
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e1, e2:Z
He':(e1 < e2)%Z
He:forall e : Z, (e1 < e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx1:(bpow e1 <= Rabs x)%R
Hx2:Rabs x = bpow e2
(bpow (e2 - 1) <= Rabs x <= bpow e2)%R
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e1, e2:Z
He':(e1 < e2)%Z
He:forall e : Z, (e1 < e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx1:(bpow e1 <= Rabs x)%R
Hx2:Rabs x = bpow e2

(e1 < e2 <= e2)%Z
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e1, e2:Z
He':(e1 < e2)%Z
He:forall e : Z, (e1 < e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx1:(bpow e1 <= Rabs x)%R
Hx2:Rabs x = bpow e2
(bpow (e2 - 1) <= Rabs x <= bpow e2)%R
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e1, e2:Z
He':(e1 < e2)%Z
He:forall e : Z, (e1 < e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx1:(bpow e1 <= Rabs x)%R
Hx2:Rabs x = bpow e2

(e1 < e2)%Z
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e1, e2:Z
He':(e1 < e2)%Z
He:forall e : Z, (e1 < e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx1:(bpow e1 <= Rabs x)%R
Hx2:Rabs x = bpow e2
(e2 <= e2)%Z
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e1, e2:Z
He':(e1 < e2)%Z
He:forall e : Z, (e1 < e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx1:(bpow e1 <= Rabs x)%R
Hx2:Rabs x = bpow e2
(bpow (e2 - 1) <= Rabs x <= bpow e2)%R
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e1, e2:Z
He':(e1 < e2)%Z
He:forall e : Z, (e1 < e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx1:(bpow e1 <= Rabs x)%R
Hx2:Rabs x = bpow e2

(e2 <= e2)%Z
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e1, e2:Z
He':(e1 < e2)%Z
He:forall e : Z, (e1 < e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx1:(bpow e1 <= Rabs x)%R
Hx2:Rabs x = bpow e2
(bpow (e2 - 1) <= Rabs x <= bpow e2)%R
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e1, e2:Z
He':(e1 < e2)%Z
He:forall e : Z, (e1 < e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx1:(bpow e1 <= Rabs x)%R
Hx2:Rabs x = bpow e2

(bpow (e2 - 1) <= Rabs x <= bpow e2)%R
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e1, e2:Z
He':(e1 < e2)%Z
He:forall e : Z, (e1 < e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx1:(bpow e1 <= Rabs x)%R
Hx2:Rabs x = bpow e2

(bpow (e2 - 1) <= bpow e2 <= bpow e2)%R
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e1, e2:Z
He':(e1 < e2)%Z
He:forall e : Z, (e1 < e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx1:(bpow e1 <= Rabs x)%R
Hx2:Rabs x = bpow e2

(bpow (e2 - 1) <= bpow e2)%R
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e1, e2:Z
He':(e1 < e2)%Z
He:forall e : Z, (e1 < e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx1:(bpow e1 <= Rabs x)%R
Hx2:Rabs x = bpow e2
(bpow e2 <= bpow e2)%R
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e1, e2:Z
He':(e1 < e2)%Z
He:forall e : Z, (e1 < e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx1:(bpow e1 <= Rabs x)%R
Hx2:Rabs x = bpow e2

(e2 - 1 <= e2)%Z
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e1, e2:Z
He':(e1 < e2)%Z
He:forall e : Z, (e1 < e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx1:(bpow e1 <= Rabs x)%R
Hx2:Rabs x = bpow e2
(bpow e2 <= bpow e2)%R
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e1, e2:Z
He':(e1 < e2)%Z
He:forall e : Z, (e1 < e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx1:(bpow e1 <= Rabs x)%R
Hx2:Rabs x = bpow e2

(bpow e2 <= bpow e2)%R
apply Rle_refl. Qed.
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2

forall e2 : Z, (forall e : Z, (e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z) -> forall x : R, (Rabs x <= bpow e2)%R -> generic_format fexp1 x -> generic_format fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2

forall e2 : Z, (forall e : Z, (e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z) -> forall x : R, (Rabs x <= bpow e2)%R -> generic_format fexp1 x -> generic_format fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e2:Z
He:forall e : Z, (e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx:(Rabs x < bpow e2)%R

generic_format fexp1 x -> generic_format fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e2:Z
He:forall e : Z, (e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx:Rabs x = bpow e2
generic_format fexp1 x -> generic_format fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e2:Z
He:forall e : Z, (e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx:(Rabs x < bpow e2)%R

x <> 0%R -> (fexp2 (mag beta x) <= fexp1 (mag beta x))%Z
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e2:Z
He:forall e : Z, (e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx:Rabs x = bpow e2
generic_format fexp1 x -> generic_format fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e2:Z
He:forall e : Z, (e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx:(Rabs x < bpow e2)%R
Zx:x <> 0%R

(fexp2 (mag beta x) <= fexp1 (mag beta x))%Z
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e2:Z
He:forall e : Z, (e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx:Rabs x = bpow e2
generic_format fexp1 x -> generic_format fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e2:Z
He:forall e : Z, (e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx:(Rabs x < bpow e2)%R
Zx:x <> 0%R

(mag beta x <= e2)%Z
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e2:Z
He:forall e : Z, (e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx:Rabs x = bpow e2
generic_format fexp1 x -> generic_format fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e2:Z
He:forall e : Z, (e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx:Rabs x = bpow e2

generic_format fexp1 x -> generic_format fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e2:Z
He:forall e : Z, (e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx:Rabs x = bpow e2

(fexp2 e2 <= fexp1 e2)%Z
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e2:Z
He:forall e : Z, (e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx:Rabs x = bpow e2
(bpow (e2 - 1) <= Rabs x <= bpow e2)%R
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e2:Z
He:forall e : Z, (e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx:Rabs x = bpow e2

(e2 <= e2)%Z
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e2:Z
He:forall e : Z, (e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx:Rabs x = bpow e2
(bpow (e2 - 1) <= Rabs x <= bpow e2)%R
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e2:Z
He:forall e : Z, (e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx:Rabs x = bpow e2

(bpow (e2 - 1) <= Rabs x <= bpow e2)%R
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e2:Z
He:forall e : Z, (e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx:Rabs x = bpow e2

(bpow (e2 - 1) <= bpow e2 <= bpow e2)%R
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e2:Z
He:forall e : Z, (e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx:Rabs x = bpow e2

(bpow (e2 - 1) <= bpow e2)%R
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e2:Z
He:forall e : Z, (e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx:Rabs x = bpow e2
(bpow e2 <= bpow e2)%R
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e2:Z
He:forall e : Z, (e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx:Rabs x = bpow e2

(e2 - 1 <= e2)%Z
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e2:Z
He:forall e : Z, (e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx:Rabs x = bpow e2
(bpow e2 <= bpow e2)%R
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e2:Z
He:forall e : Z, (e <= e2)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx:Rabs x = bpow e2

(bpow e2 <= bpow e2)%R
apply Rle_refl. Qed.
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2

forall e1 : Z, (forall e : Z, (e1 < e)%Z -> (fexp2 e <= fexp1 e)%Z) -> forall x : R, (bpow e1 <= Rabs x)%R -> generic_format fexp1 x -> generic_format fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2

forall e1 : Z, (forall e : Z, (e1 < e)%Z -> (fexp2 e <= fexp1 e)%Z) -> forall x : R, (bpow e1 <= Rabs x)%R -> generic_format fexp1 x -> generic_format fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e1:Z
He:forall e : Z, (e1 < e)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx:(bpow e1 <= Rabs x)%R

generic_format fexp1 x -> generic_format fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e1:Z
He:forall e : Z, (e1 < e)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx:(bpow e1 <= Rabs x)%R

x <> 0%R -> (fexp2 (mag beta x) <= fexp1 (mag beta x))%Z
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e1:Z
He:forall e : Z, (e1 < e)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx:(bpow e1 <= Rabs x)%R
Zx:x <> 0%R

(fexp2 (mag beta x) <= fexp1 (mag beta x))%Z
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
e1:Z
He:forall e : Z, (e1 < e)%Z -> (fexp2 e <= fexp1 e)%Z
x:R
Hx:(bpow e1 <= Rabs x)%R
Zx:x <> 0%R

(e1 < mag beta x)%Z
now apply mag_gt_bpow. Qed. Variable rnd : R -> Z. Context { valid_rnd : Valid_rnd rnd }.
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd

forall x : R, generic_format fexp1 x -> generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd

forall x : R, generic_format fexp1 x -> generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Gx:generic_format fexp1 x

generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Gx:generic_format fexp1 x

generic_format fexp1 (Rabs (round fexp2 rnd x))
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Gx:generic_format fexp1 (Rabs x)

generic_format fexp1 (Rabs (round fexp2 rnd x))
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2

forall rnd : R -> Z, Valid_rnd rnd -> forall x : R, generic_format fexp1 (Rabs x) -> generic_format fexp1 (Rabs (round fexp2 rnd x))
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2

forall rnd : R -> Z, Valid_rnd rnd -> forall x : R, (0 <= x)%R -> generic_format fexp1 x -> generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x

generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:0%R = x
Gx:generic_format fexp1 x
generic_format fexp1 (round fexp2 rnd x)
(* x > 0 *)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x

(cexp fexp1 x < mag beta x)%Z -> generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:0%R = x
Gx:generic_format fexp1 x
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x

(fexp1 (mag beta x) < mag beta x)%Z -> generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:0%R = x
Gx:generic_format fexp1 x
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R

(fexp1 ex < ex)%Z -> generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:0%R = x
Gx:generic_format fexp1 x
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R

(fexp1 ex < ex)%Z -> generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:0%R = x
Gx:generic_format fexp1 x
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
He':(fexp1 ex < ex)%Z

generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:0%R = x
Gx:generic_format fexp1 x
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
He':(fexp1 ex < ex)%Z

generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:0%R = x
Gx:generic_format fexp1 x
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
He':(fexp1 ex < ex)%Z
He:(ex <= fexp2 ex)%Z

generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
He':(fexp1 ex < ex)%Z
He:(fexp2 ex < ex)%Z
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:0%R = x
Gx:generic_format fexp1 x
generic_format fexp1 (round fexp2 rnd x)
(* - x near 0 for fexp2 *)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
He':(fexp1 ex < ex)%Z
He:(ex <= fexp2 ex)%Z
Hr:round fexp2 rnd x = 0%R

generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
He':(fexp1 ex < ex)%Z
He:(ex <= fexp2 ex)%Z
Hr:round fexp2 rnd x = bpow (fexp2 ex)
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
He':(fexp1 ex < ex)%Z
He:(fexp2 ex < ex)%Z
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:0%R = x
Gx:generic_format fexp1 x
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
He':(fexp1 ex < ex)%Z
He:(ex <= fexp2 ex)%Z
Hr:round fexp2 rnd x = 0%R

generic_format fexp1 0
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
He':(fexp1 ex < ex)%Z
He:(ex <= fexp2 ex)%Z
Hr:round fexp2 rnd x = bpow (fexp2 ex)
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
He':(fexp1 ex < ex)%Z
He:(fexp2 ex < ex)%Z
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:0%R = x
Gx:generic_format fexp1 x
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
He':(fexp1 ex < ex)%Z
He:(ex <= fexp2 ex)%Z
Hr:round fexp2 rnd x = bpow (fexp2 ex)

generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
He':(fexp1 ex < ex)%Z
He:(fexp2 ex < ex)%Z
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:0%R = x
Gx:generic_format fexp1 x
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
He':(fexp1 ex < ex)%Z
He:(ex <= fexp2 ex)%Z
Hr:round fexp2 rnd x = bpow (fexp2 ex)

generic_format fexp1 (bpow (fexp2 ex))
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
He':(fexp1 ex < ex)%Z
He:(fexp2 ex < ex)%Z
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:0%R = x
Gx:generic_format fexp1 x
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
He':(fexp1 ex < ex)%Z
He:(ex <= fexp2 ex)%Z
Hr:round fexp2 rnd x = bpow (fexp2 ex)

(fexp1 (fexp2 ex) <= fexp2 ex)%Z
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
He':(fexp1 ex < ex)%Z
He:(fexp2 ex < ex)%Z
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:0%R = x
Gx:generic_format fexp1 x
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
He':(fexp1 ex < ex)%Z
He:(ex <= fexp2 ex)%Z
Hr:round fexp2 rnd x = bpow (fexp2 ex)

(fexp1 (fexp2 ex) < fexp2 ex)%Z
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
He':(fexp1 ex < ex)%Z
He:(fexp2 ex < ex)%Z
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:0%R = x
Gx:generic_format fexp1 x
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
He':(fexp1 ex < ex)%Z
He:(fexp2 ex < ex)%Z

generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:0%R = x
Gx:generic_format fexp1 x
generic_format fexp1 (round fexp2 rnd x)
(* - x large for fexp2 *)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
He':(fexp1 ex < ex)%Z
He:(fexp2 ex < ex)%Z
He'':(cexp fexp2 x <= cexp fexp1 x)%Z

generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
He':(fexp1 ex < ex)%Z
He:(fexp2 ex < ex)%Z
He'':(cexp fexp1 x < cexp fexp2 x)%Z
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:0%R = x
Gx:generic_format fexp1 x
generic_format fexp1 (round fexp2 rnd x)
(* - - round fexp2 x is representable for fexp1 *)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
He':(fexp1 ex < ex)%Z
He:(fexp2 ex < ex)%Z
He'':(cexp fexp2 x <= cexp fexp1 x)%Z

generic_format fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
He':(fexp1 ex < ex)%Z
He:(fexp2 ex < ex)%Z
He'':(cexp fexp1 x < cexp fexp2 x)%Z
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:0%R = x
Gx:generic_format fexp1 x
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
He':(fexp1 ex < ex)%Z
He:(fexp2 ex < ex)%Z
He'':(cexp fexp2 x <= cexp fexp1 x)%Z

generic_format fexp2 (F2R {| Fnum := Ztrunc (scaled_mantissa fexp1 x); Fexp := cexp fexp1 x |})
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
He':(fexp1 ex < ex)%Z
He:(fexp2 ex < ex)%Z
He'':(cexp fexp1 x < cexp fexp2 x)%Z
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:0%R = x
Gx:generic_format fexp1 x
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
He':(fexp1 ex < ex)%Z
He:(fexp2 ex < ex)%Z
He'':(cexp fexp2 x <= cexp fexp1 x)%Z

Ztrunc (scaled_mantissa fexp1 x) <> 0%Z -> (cexp fexp2 (F2R {| Fnum := Ztrunc (scaled_mantissa fexp1 x); Fexp := cexp fexp1 x |}) <= cexp fexp1 x)%Z
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
He':(fexp1 ex < ex)%Z
He:(fexp2 ex < ex)%Z
He'':(cexp fexp1 x < cexp fexp2 x)%Z
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:0%R = x
Gx:generic_format fexp1 x
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
He':(fexp1 ex < ex)%Z
He:(fexp2 ex < ex)%Z
He'':(cexp fexp2 x <= cexp fexp1 x)%Z

Ztrunc (scaled_mantissa fexp1 x) <> 0%Z -> (cexp fexp2 (round fexp1 Ztrunc x) <= cexp fexp1 x)%Z
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
He':(fexp1 ex < ex)%Z
He:(fexp2 ex < ex)%Z
He'':(cexp fexp1 x < cexp fexp2 x)%Z
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:0%R = x
Gx:generic_format fexp1 x
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
He':(fexp1 ex < ex)%Z
He:(fexp2 ex < ex)%Z
He'':(cexp fexp2 x <= cexp fexp1 x)%Z
Zx:Ztrunc (scaled_mantissa fexp1 x) <> 0%Z

(cexp fexp2 (round fexp1 Ztrunc x) <= cexp fexp1 x)%Z
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
He':(fexp1 ex < ex)%Z
He:(fexp2 ex < ex)%Z
He'':(cexp fexp1 x < cexp fexp2 x)%Z
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:0%R = x
Gx:generic_format fexp1 x
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
He':(fexp1 ex < ex)%Z
He:(fexp2 ex < ex)%Z
He'':(cexp fexp2 x <= cexp fexp1 x)%Z
Zx:Ztrunc (scaled_mantissa fexp1 x) <> 0%Z

(fexp2 (mag beta (round fexp1 Ztrunc x)) <= cexp fexp1 x)%Z
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
He':(fexp1 ex < ex)%Z
He:(fexp2 ex < ex)%Z
He'':(cexp fexp1 x < cexp fexp2 x)%Z
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:0%R = x
Gx:generic_format fexp1 x
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
He':(fexp1 ex < ex)%Z
He:(fexp2 ex < ex)%Z
He'':(cexp fexp2 x <= cexp fexp1 x)%Z
Zx:Ztrunc (scaled_mantissa fexp1 x) <> 0%Z

round fexp1 Ztrunc x <> 0%R
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
He':(fexp1 ex < ex)%Z
He:(fexp2 ex < ex)%Z
He'':(cexp fexp1 x < cexp fexp2 x)%Z
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:0%R = x
Gx:generic_format fexp1 x
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
He':(fexp1 ex < ex)%Z
He:(fexp2 ex < ex)%Z
He'':(cexp fexp2 x <= cexp fexp1 x)%Z
Zx:round fexp1 Ztrunc x = 0%R

Ztrunc (scaled_mantissa fexp1 x) = 0%Z
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
He':(fexp1 ex < ex)%Z
He:(fexp2 ex < ex)%Z
He'':(cexp fexp1 x < cexp fexp2 x)%Z
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:0%R = x
Gx:generic_format fexp1 x
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
He':(fexp1 ex < ex)%Z
He:(fexp2 ex < ex)%Z
He'':(cexp fexp1 x < cexp fexp2 x)%Z

generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:0%R = x
Gx:generic_format fexp1 x
generic_format fexp1 (round fexp2 rnd x)
(* - - round fexp2 x has too many digits for fexp1 *)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
He':(fexp1 ex < ex)%Z
He:(fexp2 ex < ex)%Z
He'':(cexp fexp1 x < cexp fexp2 x)%Z
Hr1:(bpow (ex - 1) <= round fexp2 rnd x)%R
Hr2:(round fexp2 rnd x < bpow ex)%R

generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
He':(fexp1 ex < ex)%Z
He:(fexp2 ex < ex)%Z
He'':(cexp fexp1 x < cexp fexp2 x)%Z
Hr1:(bpow (ex - 1) <= round fexp2 rnd x)%R
Hr2:round fexp2 rnd x = bpow ex
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:0%R = x
Gx:generic_format fexp1 x
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
He':(fexp1 ex < ex)%Z
He:(fexp2 ex < ex)%Z
He'':(cexp fexp1 x < cexp fexp2 x)%Z
Hr1:(bpow (ex - 1) <= round fexp2 rnd x)%R
Hr2:(round fexp2 rnd x < bpow ex)%R

rnd (scaled_mantissa fexp2 x) <> 0%Z -> (cexp fexp1 (F2R {| Fnum := rnd (scaled_mantissa fexp2 x); Fexp := cexp fexp2 x |}) <= cexp fexp2 x)%Z
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
He':(fexp1 ex < ex)%Z
He:(fexp2 ex < ex)%Z
He'':(cexp fexp1 x < cexp fexp2 x)%Z
Hr1:(bpow (ex - 1) <= round fexp2 rnd x)%R
Hr2:round fexp2 rnd x = bpow ex
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:0%R = x
Gx:generic_format fexp1 x
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
He':(fexp1 ex < ex)%Z
He:(fexp2 ex < ex)%Z
He'':(cexp fexp1 x < cexp fexp2 x)%Z
Hr1:(bpow (ex - 1) <= round fexp2 rnd x)%R
Hr2:(round fexp2 rnd x < bpow ex)%R
Zx:rnd (scaled_mantissa fexp2 x) <> 0%Z

(cexp fexp1 (F2R {| Fnum := rnd (scaled_mantissa fexp2 x); Fexp := cexp fexp2 x |}) <= cexp fexp2 x)%Z
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
He':(fexp1 ex < ex)%Z
He:(fexp2 ex < ex)%Z
He'':(cexp fexp1 x < cexp fexp2 x)%Z
Hr1:(bpow (ex - 1) <= round fexp2 rnd x)%R
Hr2:round fexp2 rnd x = bpow ex
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:0%R = x
Gx:generic_format fexp1 x
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
He':(fexp1 ex < ex)%Z
He:(fexp2 ex < ex)%Z
He'':(cexp fexp1 x < cexp fexp2 x)%Z
Hr1:(bpow (ex - 1) <= round fexp2 rnd x)%R
Hr2:(round fexp2 rnd x < bpow ex)%R
Zx:rnd (scaled_mantissa fexp2 x) <> 0%Z

(cexp fexp1 (round fexp2 rnd x) <= cexp fexp2 x)%Z
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
He':(fexp1 ex < ex)%Z
He:(fexp2 ex < ex)%Z
He'':(cexp fexp1 x < cexp fexp2 x)%Z
Hr1:(bpow (ex - 1) <= round fexp2 rnd x)%R
Hr2:round fexp2 rnd x = bpow ex
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:0%R = x
Gx:generic_format fexp1 x
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
He':(fexp1 ex < ex)%Z
He:(fexp2 ex < ex)%Z
He'':(cexp fexp1 x < cexp fexp2 x)%Z
Hr1:(bpow (ex - 1) <= round fexp2 rnd x)%R
Hr2:(round fexp2 rnd x < bpow ex)%R
Zx:rnd (scaled_mantissa fexp2 x) <> 0%Z

(fexp1 (mag beta (round fexp2 rnd x)) <= cexp fexp2 x)%Z
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
He':(fexp1 ex < ex)%Z
He:(fexp2 ex < ex)%Z
He'':(cexp fexp1 x < cexp fexp2 x)%Z
Hr1:(bpow (ex - 1) <= round fexp2 rnd x)%R
Hr2:round fexp2 rnd x = bpow ex
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:0%R = x
Gx:generic_format fexp1 x
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
He':(fexp1 ex < ex)%Z
He:(fexp2 ex < ex)%Z
He'':(cexp fexp1 x < cexp fexp2 x)%Z
Hr1:(bpow (ex - 1) <= round fexp2 rnd x)%R
Hr2:(round fexp2 rnd x < bpow ex)%R
Zx:rnd (scaled_mantissa fexp2 x) <> 0%Z

(fexp1 ex <= cexp fexp2 x)%Z
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
He':(fexp1 ex < ex)%Z
He:(fexp2 ex < ex)%Z
He'':(cexp fexp1 x < cexp fexp2 x)%Z
Hr1:(bpow (ex - 1) <= round fexp2 rnd x)%R
Hr2:round fexp2 rnd x = bpow ex
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:0%R = x
Gx:generic_format fexp1 x
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
He':(fexp1 ex < ex)%Z
He:(fexp2 ex < ex)%Z
He'':(cexp fexp1 x < cexp fexp2 x)%Z
Hr1:(bpow (ex - 1) <= round fexp2 rnd x)%R
Hr2:(round fexp2 rnd x < bpow ex)%R
Zx:rnd (scaled_mantissa fexp2 x) <> 0%Z

(fexp1 (mag beta x) <= cexp fexp2 x)%Z
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
He':(fexp1 ex < ex)%Z
He:(fexp2 ex < ex)%Z
He'':(cexp fexp1 x < cexp fexp2 x)%Z
Hr1:(bpow (ex - 1) <= round fexp2 rnd x)%R
Hr2:round fexp2 rnd x = bpow ex
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:0%R = x
Gx:generic_format fexp1 x
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
He':(fexp1 ex < ex)%Z
He:(fexp2 ex < ex)%Z
He'':(cexp fexp1 x < cexp fexp2 x)%Z
Hr1:(bpow (ex - 1) <= round fexp2 rnd x)%R
Hr2:round fexp2 rnd x = bpow ex

generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:0%R = x
Gx:generic_format fexp1 x
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
He':(fexp1 ex < ex)%Z
He:(fexp2 ex < ex)%Z
He'':(cexp fexp1 x < cexp fexp2 x)%Z
Hr1:(bpow (ex - 1) <= round fexp2 rnd x)%R
Hr2:round fexp2 rnd x = bpow ex

generic_format fexp1 (bpow ex)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:0%R = x
Gx:generic_format fexp1 x
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:(0 < x)%R
Gx:generic_format fexp1 x
ex:Z
Ex:(bpow (ex - 1) <= x < bpow ex)%R
He':(fexp1 ex < ex)%Z
He:(fexp2 ex < ex)%Z
He'':(cexp fexp1 x < cexp fexp2 x)%Z
Hr1:(bpow (ex - 1) <= round fexp2 rnd x)%R
Hr2:round fexp2 rnd x = bpow ex

(fexp1 ex <= ex)%Z
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:0%R = x
Gx:generic_format fexp1 x
generic_format fexp1 (round fexp2 rnd x)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:0%R = x
Gx:generic_format fexp1 x

generic_format fexp1 (round fexp2 rnd x)
(* x = 0 *)
beta:radix
fexp1, fexp2:Z -> Z
valid_exp1:Valid_exp fexp1
valid_exp2:Valid_exp fexp2
rnd:R -> Z
valid_rnd:Valid_rnd rnd
x:R
Hx:0%R = x
Gx:generic_format fexp1 x

generic_format fexp1 0
apply generic_format_0. Qed. End Inclusion. End Generic. Notation ZnearestA := (Znearest (Zle_bool 0)). Section rndNA_opp.

forall (beta : radix) (fexp : Z -> Z) (x : R), round beta fexp ZnearestA (- x) = (- round beta fexp ZnearestA x)%R

forall (beta : radix) (fexp : Z -> Z) (x : R), round beta fexp ZnearestA (- x) = (- round beta fexp ZnearestA x)%R
beta:radix
fexp:Z -> Z
x:R

round beta fexp ZnearestA (- x) = (- round beta fexp ZnearestA x)%R
beta:radix
fexp:Z -> Z
x:R

(- round beta fexp (Znearest (fun t : Z => negb (0 <=? - (t + 1))%Z)) x)%R = (- round beta fexp ZnearestA x)%R
beta:radix
fexp:Z -> Z
x:R

round beta fexp (Znearest (fun t : Z => negb (0 <=? - (t + 1))%Z)) x = round beta fexp ZnearestA x
beta:radix
fexp:Z -> Z
x:R

forall x0 : R, Znearest (fun t : Z => negb (0 <=? - (t + 1))%Z) x0 = ZnearestA x0
beta:radix
fexp:Z -> Z
x:R

Znearest (fun t : Z => negb (0 <=? - (t + 1))%Z) x = ZnearestA x
beta:radix
fexp:Z -> Z
x:R

match Rcompare (x - IZR (Zfloor x)) (/ 2) with | Eq => if negb (0 <=? - (Zfloor x + 1))%Z then Zceil x else Zfloor x | Lt => Zfloor x | Gt => Zceil x end = match Rcompare (x - IZR (Zfloor x)) (/ 2) with | Eq => if (0 <=? Zfloor x)%Z then Zceil x else Zfloor x | Lt => Zfloor x | Gt => Zceil x end
beta:radix
fexp:Z -> Z
x:R
C:Rcompare (x - IZR (Zfloor x)) (/ 2) = Eq

(if negb (0 <=? - (Zfloor x + 1))%Z then Zceil x else Zfloor x) = (if (0 <=? Zfloor x)%Z then Zceil x else Zfloor x)
beta:radix
fexp:Z -> Z
x:R
C:(x - IZR (Zfloor x))%R = (/ 2)%R

(if negb (0 <=? - (Zfloor x + 1))%Z then Zceil x else Zfloor x) = (if (0 <=? Zfloor x)%Z then Zceil x else Zfloor x)
beta:radix
fexp:Z -> Z
x:R
C:(x - IZR (Zfloor x))%R = (/ 2)%R

negb (0 <=? - (Zfloor x + 1))%Z = (0 <=? Zfloor x)%Z
beta:radix
fexp:Z -> Z
x:R
C:(x - IZR (Zfloor x))%R = (/ 2)%R

(- (Zfloor x + 1) <? 0)%Z = (0 <=? Zfloor x)%Z
beta:radix
fexp:Z -> Z
x:R
C:(x - IZR (Zfloor x))%R = (/ 2)%R
C':(0 <=? Zfloor x)%Z = true

(- (Zfloor x + 1) <? 0)%Z = true
beta:radix
fexp:Z -> Z
x:R
C:(x - IZR (Zfloor x))%R = (/ 2)%R
C':(0 <=? Zfloor x)%Z = false
(- (Zfloor x + 1) <? 0)%Z = false
beta:radix
fexp:Z -> Z
x:R
C:(x - IZR (Zfloor x))%R = (/ 2)%R
C':(0 <=? Zfloor x)%Z = true

(- (Zfloor x + 1) <? 0)%Z = true
beta:radix
fexp:Z -> Z
x:R
C:(x - IZR (Zfloor x))%R = (/ 2)%R
C':(0 <= Zfloor x)%Z

(- (Zfloor x + 1) <? 0)%Z = true
beta:radix
fexp:Z -> Z
x:R
C:(x - IZR (Zfloor x))%R = (/ 2)%R
C':(0 <= Zfloor x)%Z

(- (Zfloor x + 1) < 0)%Z
omega.
beta:radix
fexp:Z -> Z
x:R
C:(x - IZR (Zfloor x))%R = (/ 2)%R
C':(0 <=? Zfloor x)%Z = false

(- (Zfloor x + 1) <? 0)%Z = false
beta:radix
fexp:Z -> Z
x:R
C:(x - IZR (Zfloor x))%R = (/ 2)%R
C':(Zfloor x < 0)%Z

(- (Zfloor x + 1) <? 0)%Z = false
beta:radix
fexp:Z -> Z
x:R
C:(x - IZR (Zfloor x))%R = (/ 2)%R
C':(Zfloor x < 0)%Z

(0 <= - (Zfloor x + 1))%Z
omega. Qed. End rndNA_opp.
Notations for backward-compatibility with Flocq 1.4.
Notation rndDN := Zfloor (only parsing).
Notation rndUP := Zceil (only parsing).
Notation rndZR := Ztrunc (only parsing).
Notation rndNA := ZnearestA (only parsing).