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) 2014-2018 Sylvie Boldo
Copyright (C) 2014-2018 Guillaume Melquiond
Copyright (C) 2014-2018 Pierre Roux
This library is free software; you can redistribute it and/or modify it under the terms of the GNU Lesser General Public License as published by the Free Software Foundation; either version 3 of the License, or (at your option) any later version.
This library is distributed in the hope that it will be useful, but WITHOUT ANY WARRANTY; without even the implied warranty of MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the COPYING file for more details.

Conditions for innocuous double rounding.

Require Import Psatz.
Require Import Raux Defs Generic_fmt Operations Ulp FLX FLT FTZ.

Open Scope R_scope.

Section Double_round.

Variable beta : radix.
Notation bpow e := (bpow beta e).
Notation mag x := (mag beta x).

Definition round_round_eq fexp1 fexp2 choice1 choice2 x :=
  round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x)
  = round beta fexp1 (Znearest choice1) x.
A little tactic to simplify terms of the form bpow a × bpow b.
Ltac bpow_simplify :=
  (* bpow ex * bpow ey ~~> bpow (ex + ey) *)
  repeat
    match goal with
      | |- context [(Raux.bpow _ _ * Raux.bpow _ _)] =>
        rewrite <- bpow_plus
      | |- context [(?X1 * Raux.bpow _ _ * Raux.bpow _ _)] =>
        rewrite (Rmult_assoc X1); rewrite <- bpow_plus
      | |- context [(?X1 * (?X2 * Raux.bpow _ _) * Raux.bpow _ _)] =>
        rewrite <- (Rmult_assoc X1 X2); rewrite (Rmult_assoc (X1 * X2));
        rewrite <- bpow_plus
    end;
  (* ring_simplify arguments of bpow *)
  repeat
    match goal with
      | |- context [(Raux.bpow _ ?X)] =>
        progress ring_simplify X
    end;
  (* bpow 0 ~~> 1 *)
  change (Raux.bpow _ 0) with 1;
  repeat
    match goal with
      | |- context [(_ * 1)] =>
        rewrite Rmult_1_r
    end.

Definition midp (fexp : Z -> Z) (x : R) :=
  round beta fexp Zfloor x + / 2 * ulp beta fexp x.

Definition midp' (fexp : Z -> Z) (x : R) :=
  round beta fexp Zceil x - / 2 * ulp beta fexp x.

beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall (choice1 choice2 : Z -> bool) (x : R), 0 < x -> (fexp2 (mag x) <= fexp1 (mag x) - 1)%Z -> x < bpow (mag x) - / 2 * ulp beta fexp2 x -> x < midp fexp1 x - / 2 * ulp beta fexp2 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall (choice1 choice2 : Z -> bool) (x : R), 0 < x -> (fexp2 (mag x) <= fexp1 (mag x) - 1)%Z -> x < bpow (mag x) - / 2 * ulp beta fexp2 x -> x < midp fexp1 x - / 2 * ulp beta fexp2 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x

x < midp fexp1 x - / 2 * ulp beta fexp2 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x

x < midp fexp1 x - / 2 * ulp beta fexp2 x -> round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R

x < midp fexp1 x - / 2 * ulp beta fexp2 x -> round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x

x - round beta fexp1 Zfloor x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x

x - round beta fexp1 Zfloor x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
now apply (Rplus_lt_reg_r (round beta fexp1 Zfloor x)); ring_simplify.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R

round beta fexp1 (Znearest choice1) x'' = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R

Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
round beta fexp1 (Znearest choice1) x'' = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R

Rabs (x'' - x) <= / 2 * ulp beta fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
/ 2 * ulp beta fexp2 x <= / 2 * bpow (fexp2 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
round beta fexp1 (Znearest choice1) x'' = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R

/ 2 * ulp beta fexp2 x <= / 2 * bpow (fexp2 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
round beta fexp1 (Znearest choice1) x'' = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))

round beta fexp1 (Znearest choice1) x'' = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))

0 <= x - x'
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
round beta fexp1 (Znearest choice1) x'' = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))

0 <= x - x'
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))

x' <= x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))

Valid_exp fexp1
exact Vfexp1.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'

round beta fexp1 (Znearest choice1) x'' = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'

round beta fexp1 (Znearest choice1) x'' = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'

Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
round beta fexp1 (Znearest choice1) x'' = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'

Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'

Rabs (x'' - x + (x - x')) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'

Rabs (x'' - x) + Rabs (x - x') < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'

Rabs (x'' - x) + Rabs (x - x') < / 2 * bpow (fexp2 (mag x)) + / 2 * (bpow (fexp1 (mag x)) - bpow (fexp2 (mag x)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'

Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Rabs (x - x') < / 2 * (bpow (fexp1 (mag x)) - bpow (fexp2 (mag x)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'

Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
exact Hr1.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'

Rabs (x - x') < / 2 * (bpow (fexp1 (mag x)) - bpow (fexp2 (mag x)))
now rewrite Rabs_right; [|now apply Rle_ge]; apply Hx2.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))

round beta fexp1 (Znearest choice1) x'' = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Zx'':x'' = 0

round beta fexp1 (Znearest choice1) x'' = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
round beta fexp1 (Znearest choice1) x'' = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Zx'':x'' = 0

round beta fexp1 (Znearest choice1) x'' = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (0 - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Zx'':x'' = 0

round beta fexp1 (Znearest choice1) 0 = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (0 - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Zx'':x'' = 0

0 = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (0 - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Zx'':x'' = 0

0 = IZR (Znearest choice1 (x * bpow (- fexp1 (mag x)))) * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (0 - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Zx'':x'' = 0

Rabs (x * bpow (- fexp1 (mag x)) - 0) < / 2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (0 - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Zx'':x'' = 0

Rabs (x * bpow (- fexp1 (mag x)) - 0) * bpow (fexp1 (mag x)) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (0 - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Zx'':x'' = 0

Rabs (x * bpow (- fexp1 (mag x)) - 0) * Rabs (bpow (fexp1 (mag x))) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (0 - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Zx'':x'' = 0

Rabs (x * bpow (- fexp1 (mag x)) * bpow (fexp1 (mag x)) - 0 * bpow (fexp1 (mag x))) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (0 - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Zx'':x'' = 0

Rabs (x * bpow (- fexp1 (mag x)) * bpow (fexp1 (mag x)) - 0) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (0 - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Zx'':x'' = 0

Rabs (x - 0) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (0 - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Zx'':x'' = 0

Rabs (0 - x) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (0 - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Zx'':x'' = 0

/ 2 * bpow (fexp2 (mag x)) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (0 - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Zx'':x'' = 0

bpow (fexp2 (mag x)) < bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (0 - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Zx'':x'' = 0

(fexp2 (mag x) < fexp1 (mag x))%Z
omega.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0

round beta fexp1 (Znearest choice1) x'' = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0

mag x'' = mag x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
Lx'':mag x'' = mag x
round beta fexp1 (Znearest choice1) x'' = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0

mag x'' = mag x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0

(mag x'' <= mag x)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
(mag x <= mag x'')%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0

(mag x'' <= mag x)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0

Rabs x'' < bpow (mag x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0

Rabs (x'' - x + x) < bpow (mag x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0

Rabs (x'' - x) + Rabs x < bpow (mag x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0

Rabs (x'' - x) + Rabs x < / 2 * bpow (fexp2 (mag x)) + (bpow (mag x) - / 2 * bpow (fexp2 (mag x)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0

Rabs x < bpow (mag x) - / 2 * bpow (fexp2 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * bpow (cexp beta fexp2 x)
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0

Rabs x < bpow (mag x) - / 2 * bpow (fexp2 (mag x))
now rewrite Rabs_right; [|apply Rle_ge; apply Rlt_le].
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0

(mag x <= mag x'')%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':round beta fexp2 (Znearest choice2) x <> 0

(mag x <= mag (round beta fexp2 (Znearest choice2) x))%Z
now apply mag_round_ge; [|apply valid_rnd_N|].
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
Lx'':mag x'' = mag x

round beta fexp1 (Znearest choice1) x'' = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
Lx'':mag x'' = mag x

IZR (Znearest choice1 (x'' * bpow (- fexp1 (mag x'')))) * bpow (fexp1 (mag x'')) = IZR (Znearest choice1 (x * bpow (- fexp1 (mag x)))) * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
Lx'':mag x'' = mag x

IZR (Znearest choice1 (x'' * bpow (- fexp1 (mag x)))) * bpow (fexp1 (mag x)) = IZR (Znearest choice1 (x * bpow (- fexp1 (mag x)))) * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
Lx'':mag x'' = mag x

IZR (Zfloor (scaled_mantissa beta fexp1 x)) * bpow (fexp1 (mag x)) = IZR (Znearest choice1 (x * bpow (- fexp1 (mag x)))) * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
Lx'':mag x'' = mag x
Rabs (x'' * bpow (- fexp1 (mag x)) - IZR (Zfloor (scaled_mantissa beta fexp1 x))) < / 2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
Lx'':mag x'' = mag x

IZR (Zfloor (scaled_mantissa beta fexp1 x)) * bpow (fexp1 (mag x)) = IZR (Znearest choice1 (x * bpow (- fexp1 (mag x)))) * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
Lx'':mag x'' = mag x

Rabs (x * bpow (- fexp1 (mag x)) - IZR (Zfloor (scaled_mantissa beta fexp1 x))) < / 2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
Lx'':mag x'' = mag x

Rabs (x * bpow (- fexp1 (mag x)) - IZR (Zfloor (scaled_mantissa beta fexp1 x))) * bpow (fexp1 (mag x)) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
Lx'':mag x'' = mag x

Rabs (x * bpow (- fexp1 (mag x)) - IZR (Zfloor (scaled_mantissa beta fexp1 x))) * Rabs (bpow (fexp1 (mag x))) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
Lx'':mag x'' = mag x

Rabs ((x * bpow (- fexp1 (mag x)) - IZR (Zfloor (scaled_mantissa beta fexp1 x))) * bpow (fexp1 (mag x))) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
Lx'':mag x'' = mag x

Rabs (x * bpow (- fexp1 (mag x)) * bpow (fexp1 (mag x)) - IZR (Zfloor (scaled_mantissa beta fexp1 x)) * bpow (fexp1 (mag x))) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
Lx'':mag x'' = mag x

Rabs (x * bpow (- fexp1 (mag x)) * bpow (fexp1 (mag x)) - IZR (Zfloor (scaled_mantissa beta fexp1 x)) * bpow (fexp1 (mag x))) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
Lx'':mag x'' = mag x

Rabs (x - IZR (Zfloor (scaled_mantissa beta fexp1 x)) * bpow (fexp1 (mag x))) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
Lx'':mag x'' = mag x

x - IZR (Zfloor (scaled_mantissa beta fexp1 x)) * bpow (fexp1 (mag x)) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
Lx'':mag x'' = mag x

/ 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x)) <= / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
Lx'':mag x'' = mag x

bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x) <= bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
Lx'':mag x'' = mag x

0 <= bpow (fexp2 (mag x)) -> bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x) <= bpow (fexp1 (mag x))
unfold ulp, cexp; lra.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
Lx'':mag x'' = mag x

Rabs (x'' * bpow (- fexp1 (mag x)) - IZR (Zfloor (scaled_mantissa beta fexp1 x))) < / 2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
Lx'':mag x'' = mag x

Rabs (x'' * bpow (- fexp1 (mag x)) - IZR (Zfloor (scaled_mantissa beta fexp1 x))) * bpow (fexp1 (mag x)) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
Lx'':mag x'' = mag x

Rabs (x'' * bpow (- fexp1 (mag x)) - IZR (Zfloor (scaled_mantissa beta fexp1 x))) * Rabs (bpow (fexp1 (mag x))) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
Lx'':mag x'' = mag x

Rabs ((x'' * bpow (- fexp1 (mag x)) - IZR (Zfloor (scaled_mantissa beta fexp1 x))) * bpow (fexp1 (mag x))) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
Lx'':mag x'' = mag x

Rabs (x'' * bpow (- fexp1 (mag x)) * bpow (fexp1 (mag x)) - IZR (Zfloor (scaled_mantissa beta fexp1 x)) * bpow (fexp1 (mag x))) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:x < bpow (mag x) - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
x'':=round beta fexp2 (Znearest choice2) x:R
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Pxx':0 <= x - x'
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
Lx'':mag x'' = mag x

Rabs (x'' * bpow (- fexp1 (mag x)) * bpow (fexp1 (mag x)) - IZR (Zfloor (scaled_mantissa beta fexp1 x)) * bpow (fexp1 (mag x))) < / 2 * bpow (fexp1 (mag x))
now bpow_simplify. Qed.
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall (choice1 choice2 : Z -> bool) (x : R), 0 < x -> (fexp2 (mag x) <= fexp1 (mag x) - 1)%Z -> (fexp1 (mag x) <= mag x)%Z -> x < midp fexp1 x - / 2 * ulp beta fexp2 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall (choice1 choice2 : Z -> bool) (x : R), 0 < x -> (fexp2 (mag x) <= fexp1 (mag x) - 1)%Z -> (fexp1 (mag x) <= mag x)%Z -> x < midp fexp1 x - / 2 * ulp beta fexp2 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z

x < midp fexp1 x - / 2 * ulp beta fexp2 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x

round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x

x - round beta fexp1 Zfloor x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x

x - round beta fexp1 Zfloor x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
now apply (Rplus_lt_reg_r (round beta fexp1 Zfloor x)); ring_simplify.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hx2:x - round beta fexp1 Zfloor x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)

round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x

x - round beta fexp1 Zfloor x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x) -> round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x

x - round beta fexp1 Zfloor x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x) -> round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R

x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x) -> round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)

0 <= x - x'
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)

0 <= x - x'
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)

x' <= x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)

Valid_exp fexp1
exact Vfexp1.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'

x < bpow (mag x) - / 2 * bpow (fexp2 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
H:x < bpow (mag x) - / 2 * bpow (fexp2 (mag x))
x < bpow (mag x) - / 2 * ulp beta fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'

x < bpow (mag x) - / 2 * bpow (fexp2 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Zx':x' = 0

x < bpow (mag x) - / 2 * bpow (fexp2 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Nzx':x' <> 0
x < bpow (mag x) - / 2 * bpow (fexp2 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Zx':x' = 0

x < bpow (mag x) - / 2 * bpow (fexp2 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Zx':x' = 0

x < bpow (mag x) - / 2 * bpow (fexp2 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Zx':x' = 0

/ 2 * (ulp beta fexp1 x - ulp beta fexp2 x) <= bpow (mag x) - / 2 * bpow (fexp2 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Zx':x' = 0

/ 2 * ulp beta fexp1 x - / 2 * ulp beta fexp2 x <= bpow (mag x) - / 2 * bpow (fexp2 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Zx':x' = 0

/ 2 * bpow (cexp beta fexp1 x) - / 2 * bpow (cexp beta fexp2 x) <= bpow (mag x) - / 2 * bpow (fexp2 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Zx':x' = 0

/ 2 * bpow (cexp beta fexp1 x) <= bpow (mag x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Zx':x' = 0

/ 2 * bpow (cexp beta fexp1 x) * bpow (- mag x) <= bpow (mag x) * bpow (- mag x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Zx':x' = 0

/ 2 * bpow (fexp1 (mag x) - mag x) <= 1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Zx':x' = 0

2 * (/ 2 * bpow (fexp1 (mag x) - mag x)) <= 2 * 1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Zx':x' = 0

bpow (fexp1 (mag x) - mag x) <= 2 * 1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Zx':x' = 0

bpow (fexp1 (mag x) - mag x) <= 1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Zx':x' = 0

(fexp1 (mag x) - mag x <= 0)%Z
omega.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Nzx':x' <> 0

x < bpow (mag x) - / 2 * bpow (fexp2 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Nzx':x' <> 0

0 < x'
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Nzx':x' <> 0
Px':0 < x'
x < bpow (mag x) - / 2 * bpow (fexp2 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Nzx':x' <> 0

0 < x'
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Nzx':x' <> 0

0 <= x'
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Nzx':x' <> 0

0 <= round beta fexp1 Zfloor x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Nzx':x' <> 0

0 * bpow (- fexp1 (mag x)) <= round beta fexp1 Zfloor x * bpow (- fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Nzx':x' <> 0

0 <= round beta fexp1 Zfloor x * bpow (- fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Nzx':x' <> 0

0 <= IZR (Zfloor (scaled_mantissa beta fexp1 x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Nzx':x' <> 0

(0 <= Zfloor (scaled_mantissa beta fexp1 x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Nzx':x' <> 0

0 <= scaled_mantissa beta fexp1 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Nzx':x' <> 0

0 <= scaled_mantissa beta fexp1 (Rabs x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Nzx':x' <> 0

0 <= Rabs (scaled_mantissa beta fexp1 x)
apply Rabs_pos.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Nzx':x' <> 0
Px':0 < x'

x < bpow (mag x) - / 2 * bpow (fexp2 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Nzx':x' <> 0
Px':0 < x'

x' <= bpow (mag x) - ulp beta fexp1 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Nzx':x' <> 0
Px':0 < x'
Hx':x' <= bpow (mag x) - ulp beta fexp1 x
x < bpow (mag x) - / 2 * bpow (fexp2 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Nzx':x' <> 0
Px':0 < x'

x' <= bpow (mag x) - ulp beta fexp1 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Nzx':x' <> 0
Px':0 < x'

x' + ulp beta fexp1 x <= bpow (mag x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Nzx':x' <> 0
Px':0 < x'

x' + ulp beta fexp1 (round beta fexp1 Zfloor x) <= bpow (mag x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Nzx':x' <> 0
Px':0 < x'
Valid_exp fexp1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Nzx':x' <> 0
Px':0 < x'
0 <= x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Nzx':x' <> 0
Px':0 < x'

x' + ulp beta fexp1 (round beta fexp1 Zfloor x) <= bpow (mag x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Nzx':x' <> 0
Px':0 < x'

x' + ulp beta fexp1 x' <= bpow (mag x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Nzx':x' <> 0
Px':0 < x'

0 < x'
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Nzx':x' <> 0
Px':0 < x'
generic_format beta fexp1 x'
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Nzx':x' <> 0
Px':0 < x'
x' < bpow (mag x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Nzx':x' <> 0
Px':0 < x'

0 < x'
exact Px'.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Nzx':x' <> 0
Px':0 < x'

generic_format beta fexp1 x'
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Nzx':x' <> 0
Px':0 < x'

generic_format beta fexp1 (round beta fexp1 Zfloor x)
now apply generic_format_round; [|apply valid_rnd_DN].
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Nzx':x' <> 0
Px':0 < x'

x' < bpow (mag x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Nzx':x' <> 0
Px':0 < x'

x' <= x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Nzx':x' <> 0
Px':0 < x'
x < bpow (mag x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Nzx':x' <> 0
Px':0 < x'

x' <= x
now apply round_DN_pt.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Nzx':x' <> 0
Px':0 < x'

x < bpow (mag x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Nzx':x' <> 0
Px':0 < x'

Rabs x < bpow (mag x)
apply bpow_mag_gt.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Nzx':x' <> 0
Px':0 < x'

Valid_exp fexp1
exact Vfexp1.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Nzx':x' <> 0
Px':0 < x'

0 <= x
now apply Rlt_le.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Nzx':x' <> 0
Px':0 < x'
Hx':x' <= bpow (mag x) - ulp beta fexp1 x

x < bpow (mag x) - / 2 * bpow (fexp2 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Nzx':x' <> 0
Px':0 < x'
Hx':x' <= bpow (mag x) - ulp beta fexp1 x

x < bpow (mag x) - / 2 * bpow (cexp beta fexp2 x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Nzx':x' <> 0
Px':0 < x'
Hx':x' <= bpow (mag x) - ulp beta fexp1 x

/ 2 * ulp beta fexp1 x <= ulp beta fexp1 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Nzx':x' <> 0
Px':0 < x'
Hx':x' <= bpow (mag x) - ulp beta fexp1 x
H:/ 2 * ulp beta fexp1 x <= ulp beta fexp1 x
x < bpow (mag x) - / 2 * bpow (cexp beta fexp2 x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Nzx':x' <> 0
Px':0 < x'
Hx':x' <= bpow (mag x) - ulp beta fexp1 x

/ 2 * ulp beta fexp1 x <= 1 * ulp beta fexp1 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Nzx':x' <> 0
Px':0 < x'
Hx':x' <= bpow (mag x) - ulp beta fexp1 x
H:/ 2 * ulp beta fexp1 x <= ulp beta fexp1 x
x < bpow (mag x) - / 2 * bpow (cexp beta fexp2 x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Nzx':x' <> 0
Px':0 < x'
Hx':x' <= bpow (mag x) - ulp beta fexp1 x

0 <= ulp beta fexp1 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Nzx':x' <> 0
Px':0 < x'
Hx':x' <= bpow (mag x) - ulp beta fexp1 x
H:/ 2 * ulp beta fexp1 x <= ulp beta fexp1 x
x < bpow (mag x) - / 2 * bpow (cexp beta fexp2 x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Pxx':0 <= x - x'
Nzx':x' <> 0
Px':0 < x'
Hx':x' <= bpow (mag x) - ulp beta fexp1 x
H:/ 2 * ulp beta fexp1 x <= ulp beta fexp1 x

x < bpow (mag x) - / 2 * bpow (cexp beta fexp2 x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
Pxx':0 <= x - x'
Nzx':x' <> 0
Px':0 < x'
Hx':x' <= bpow (mag x) - ulp beta fexp1 x
H:/ 2 * ulp beta fexp1 x <= ulp beta fexp1 x

x < bpow (mag x) - / 2 * bpow (cexp beta fexp2 x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
Pxx':0 <= x - x'
Nzx':x' <> 0
Px':0 < x'
Hx':x' <= bpow (mag x) - bpow (cexp beta fexp1 x)
H:/ 2 * ulp beta fexp1 x <= ulp beta fexp1 x

x < bpow (mag x) - / 2 * bpow (cexp beta fexp2 x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':x < midp fexp1 x - / 2 * ulp beta fexp2 x
x':=round beta fexp1 Zfloor x:R
Hx2:x - x' < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
Pxx':0 <= x - x'
Nzx':x' <> 0
Px':0 < x'
Hx':x' <= bpow (mag x) - bpow (cexp beta fexp1 x)
H:/ 2 * bpow (cexp beta fexp1 x) <= bpow (cexp beta fexp1 x)

x < bpow (mag x) - / 2 * bpow (cexp beta fexp2 x)
lra. Qed.
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> forall (choice1 choice2 : Z -> bool) (x : R), 0 < x -> fexp2 (mag x) = fexp1 (mag x) -> x < midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> forall (choice1 choice2 : Z -> bool) (x : R), 0 < x -> fexp2 (mag x) = fexp1 (mag x) -> x < midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)

x < midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':x < midp fexp1 x

round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':x < midp fexp1 x

x - round beta fexp1 Zfloor x < / 2 * ulp beta fexp1 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':x < midp fexp1 x
Hx:x - round beta fexp1 Zfloor x < / 2 * ulp beta fexp1 x
round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':x < midp fexp1 x

x - round beta fexp1 Zfloor x < / 2 * ulp beta fexp1 x
now apply (Rplus_lt_reg_r (round beta fexp1 Zfloor x)); ring_simplify.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':x < midp fexp1 x
Hx:x - round beta fexp1 Zfloor x < / 2 * ulp beta fexp1 x

round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':x < midp fexp1 x

x - round beta fexp1 Zfloor x < / 2 * ulp beta fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':x < midp fexp1 x

x - round beta fexp1 Zfloor x < / 2 * ulp beta fexp1 x -> round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':x < midp fexp1 x
x':=round beta fexp1 Zfloor x:R

x - x' < / 2 * ulp beta fexp1 x -> round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':x < midp fexp1 x
x':=round beta fexp1 Zfloor x:R
Hx:x - x' < / 2 * ulp beta fexp1 x

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':x < midp fexp1 x
x':=round beta fexp1 Zfloor x:R
Hx:x - x' < / 2 * ulp beta fexp1 x

0 <= x - x'
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':x < midp fexp1 x
x':=round beta fexp1 Zfloor x:R
Hx:x - x' < / 2 * ulp beta fexp1 x
Pxx':0 <= x - x'
round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':x < midp fexp1 x
x':=round beta fexp1 Zfloor x:R
Hx:x - x' < / 2 * ulp beta fexp1 x

0 <= x - x'
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':x < midp fexp1 x
x':=round beta fexp1 Zfloor x:R
Hx:x - x' < / 2 * ulp beta fexp1 x

x' <= x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':x < midp fexp1 x
x':=round beta fexp1 Zfloor x:R
Hx:x - x' < / 2 * ulp beta fexp1 x

Valid_exp fexp1
exact Vfexp1.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':x < midp fexp1 x
x':=round beta fexp1 Zfloor x:R
Hx:x - x' < / 2 * ulp beta fexp1 x
Pxx':0 <= x - x'

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':x < midp fexp1 x
x':=round beta fexp1 Zfloor x:R
Hx:x - x' < / 2 * ulp beta fexp1 x
Pxx':0 <= x - x'

Rabs (x * bpow (- fexp1 (mag x)) - IZR (Zfloor (x * bpow (- fexp1 (mag x))))) < / 2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':x < midp fexp1 x
x':=round beta fexp1 Zfloor x:R
Hx:x - x' < / 2 * ulp beta fexp1 x
Pxx':0 <= x - x'
H:Rabs (x * bpow (- fexp1 (mag x)) - IZR (Zfloor (x * bpow (- fexp1 (mag x))))) < / 2
round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':x < midp fexp1 x
x':=round beta fexp1 Zfloor x:R
Hx:x - x' < / 2 * ulp beta fexp1 x
Pxx':0 <= x - x'

Rabs (x * bpow (- fexp1 (mag x)) - IZR (Zfloor (x * bpow (- fexp1 (mag x))))) < / 2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':x < midp fexp1 x
x':=round beta fexp1 Zfloor x:R
Hx:x - x' < / 2 * ulp beta fexp1 x
Pxx':0 <= x - x'

Rabs (x * bpow (- fexp1 (mag x)) - IZR (Zfloor (x * bpow (- fexp1 (mag x))))) * bpow (fexp1 (mag x)) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':x < midp fexp1 x
x':=round beta fexp1 Zfloor x:R
Hx:x - x' < / 2 * ulp beta fexp1 x
Pxx':0 <= x - x'

Rabs (x * bpow (- fexp1 (mag x)) - IZR (Zfloor (x * bpow (- fexp1 (mag x))))) * bpow (fexp1 (mag x)) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':x < midp fexp1 x
x':=round beta fexp1 Zfloor x:R
Hx:x - x' < / 2 * ulp beta fexp1 x
Pxx':0 <= x - x'

Rabs (x * bpow (- fexp1 (mag x)) - IZR (Zfloor (x * bpow (- fexp1 (mag x))))) * Rabs (bpow (fexp1 (mag x))) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':x < midp fexp1 x
x':=round beta fexp1 Zfloor x:R
Hx:x - x' < / 2 * ulp beta fexp1 x
Pxx':0 <= x - x'

Rabs ((x * bpow (- fexp1 (mag x)) - IZR (Zfloor (x * bpow (- fexp1 (mag x))))) * bpow (fexp1 (mag x))) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':x < midp fexp1 x
x':=round beta fexp1 Zfloor x:R
Hx:x - x' < / 2 * ulp beta fexp1 x
Pxx':0 <= x - x'

Rabs (x * bpow (- fexp1 (mag x)) * bpow (fexp1 (mag x)) - IZR (Zfloor (x * bpow (- fexp1 (mag x)))) * bpow (fexp1 (mag x))) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':x < midp fexp1 x
x':=round beta fexp1 Zfloor x:R
Hx:x - x' < / 2 * ulp beta fexp1 x
Pxx':0 <= x - x'

Rabs (x - IZR (Zfloor (x * bpow (- fexp1 (mag x)))) * bpow (fexp1 (mag x))) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':x < midp fexp1 x
x':=round beta fexp1 Zfloor x:R
Hx:x - x' < / 2 * ulp beta fexp1 x
Pxx':0 <= x - x'

- (/ 2 * bpow (fexp1 (mag x))) < x - IZR (Zfloor (x * bpow (- fexp1 (mag x)))) * bpow (fexp1 (mag x)) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':x < midp fexp1 x
x':=round beta fexp1 Zfloor x:R
Hx:x - x' < / 2 * ulp beta fexp1 x
Pxx':0 <= x - x'

- (/ 2 * bpow (fexp1 (mag x))) < x - x' < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':x < midp fexp1 x
x':=round beta fexp1 Zfloor x:R
Hx:x - x' < / 2 * ulp beta fexp1 x
Pxx':0 <= x - x'

- (/ 2 * bpow (fexp1 (mag x))) < x - x'
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':x < midp fexp1 x
x':=round beta fexp1 Zfloor x:R
Hx:x - x' < / 2 * ulp beta fexp1 x
Pxx':0 <= x - x'
x - x' < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':x < midp fexp1 x
x':=round beta fexp1 Zfloor x:R
Hx:x - x' < / 2 * ulp beta fexp1 x
Pxx':0 <= x - x'

- (/ 2 * bpow (fexp1 (mag x))) < x - x'
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':x < midp fexp1 x
x':=round beta fexp1 Zfloor x:R
Hx:x - x' < / 2 * ulp beta fexp1 x
Pxx':0 <= x - x'

- (/ 2 * bpow (fexp1 (mag x))) < 0
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':x < midp fexp1 x
x':=round beta fexp1 Zfloor x:R
Hx:x - x' < / 2 * ulp beta fexp1 x
Pxx':0 <= x - x'

- (/ 2 * bpow (fexp1 (mag x))) < - 0
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':x < midp fexp1 x
x':=round beta fexp1 Zfloor x:R
Hx:x - x' < / 2 * ulp beta fexp1 x
Pxx':0 <= x - x'

0 < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':x < midp fexp1 x
x':=round beta fexp1 Zfloor x:R
Hx:x - x' < / 2 * ulp beta fexp1 x
Pxx':0 <= x - x'

/ 2 * 0 < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':x < midp fexp1 x
x':=round beta fexp1 Zfloor x:R
Hx:x - x' < / 2 * ulp beta fexp1 x
Pxx':0 <= x - x'

0 < bpow (fexp1 (mag x))
apply bpow_gt_0.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':x < midp fexp1 x
x':=round beta fexp1 Zfloor x:R
Hx:x - x' < / 2 * ulp beta fexp1 x
Pxx':0 <= x - x'

x - x' < / 2 * bpow (fexp1 (mag x))
rewrite ulp_neq_0 in Hx;try apply Rgt_not_eq; assumption.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':x < midp fexp1 x
x':=round beta fexp1 Zfloor x:R
Hx:x - x' < / 2 * ulp beta fexp1 x
Pxx':0 <= x - x'
H:Rabs (x * bpow (- fexp1 (mag x)) - IZR (Zfloor (x * bpow (- fexp1 (mag x))))) < / 2

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':x < midp fexp1 x
x':=round beta fexp1 Zfloor x:R
Hx:x - x' < / 2 * ulp beta fexp1 x
Pxx':0 <= x - x'
H:Rabs (x * bpow (- fexp1 (mag x)) - IZR (Zfloor (x * bpow (- fexp1 (mag x))))) < / 2

round beta fexp1 (Znearest choice1) (F2R {| Fnum := Znearest choice2 (scaled_mantissa beta fexp2 x); Fexp := cexp beta fexp2 x |}) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':x < midp fexp1 x
x':=round beta fexp1 Zfloor x:R
Hx:x - x' < / 2 * ulp beta fexp1 x
Pxx':0 <= x - x'
H:Rabs (x * bpow (- fexp1 (mag x)) - IZR (Zfloor (x * bpow (- fexp1 (mag x))))) < / 2

round beta fexp1 (Znearest choice1) (IZR (Znearest choice2 (x * bpow (- fexp2 (mag x)))) * bpow (fexp2 (mag x))) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':x < midp fexp1 x
x':=round beta fexp1 Zfloor x:R
Hx:x - x' < / 2 * ulp beta fexp1 x
Pxx':0 <= x - x'
H:Rabs (x * bpow (- fexp1 (mag x)) - IZR (Zfloor (x * bpow (- fexp1 (mag x))))) < / 2

round beta fexp1 (Znearest choice1) (IZR (Znearest choice2 (x * bpow (- fexp1 (mag x)))) * bpow (fexp1 (mag x))) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':x < midp fexp1 x
x':=round beta fexp1 Zfloor x:R
Hx:x - x' < / 2 * ulp beta fexp1 x
Pxx':0 <= x - x'
H:Rabs (x * bpow (- fexp1 (mag x)) - IZR (Zfloor (x * bpow (- fexp1 (mag x))))) < / 2

round beta fexp1 (Znearest choice1) (IZR (Zfloor (scaled_mantissa beta fexp1 x)) * bpow (fexp1 (mag x))) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':x < midp fexp1 x
x':=round beta fexp1 Zfloor x:R
Hx:x - x' < / 2 * ulp beta fexp1 x
Pxx':0 <= x - x'
H:Rabs (x * bpow (- fexp1 (mag x)) - IZR (Zfloor (x * bpow (- fexp1 (mag x))))) < / 2

IZR (Zfloor (scaled_mantissa beta fexp1 x)) * bpow (fexp1 (mag x)) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':x < midp fexp1 x
x':=round beta fexp1 Zfloor x:R
Hx:x - x' < / 2 * ulp beta fexp1 x
Pxx':0 <= x - x'
H:Rabs (x * bpow (- fexp1 (mag x)) - IZR (Zfloor (x * bpow (- fexp1 (mag x))))) < / 2
Valid_rnd (Znearest choice1)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':x < midp fexp1 x
x':=round beta fexp1 Zfloor x:R
Hx:x - x' < / 2 * ulp beta fexp1 x
Pxx':0 <= x - x'
H:Rabs (x * bpow (- fexp1 (mag x)) - IZR (Zfloor (x * bpow (- fexp1 (mag x))))) < / 2
generic_format beta fexp1 (IZR (Zfloor (scaled_mantissa beta fexp1 x)) * bpow (fexp1 (mag x)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':x < midp fexp1 x
x':=round beta fexp1 Zfloor x:R
Hx:x - x' < / 2 * ulp beta fexp1 x
Pxx':0 <= x - x'
H:Rabs (x * bpow (- fexp1 (mag x)) - IZR (Zfloor (x * bpow (- fexp1 (mag x))))) < / 2

IZR (Zfloor (scaled_mantissa beta fexp1 x)) * bpow (fexp1 (mag x)) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':x < midp fexp1 x
x':=round beta fexp1 Zfloor x:R
Hx:x - x' < / 2 * ulp beta fexp1 x
Pxx':0 <= x - x'
H:Rabs (x * bpow (- fexp1 (mag x)) - IZR (Zfloor (x * bpow (- fexp1 (mag x))))) < / 2

IZR (Zfloor (x * bpow (- fexp1 (mag x)))) * bpow (fexp1 (mag x)) = IZR (Znearest choice1 (x * bpow (- fexp1 (mag x)))) * bpow (fexp1 (mag x))
now rewrite (Znearest_imp _ _ (Zfloor (x * bpow (- fexp1 (mag x))))).
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':x < midp fexp1 x
x':=round beta fexp1 Zfloor x:R
Hx:x - x' < / 2 * ulp beta fexp1 x
Pxx':0 <= x - x'
H:Rabs (x * bpow (- fexp1 (mag x)) - IZR (Zfloor (x * bpow (- fexp1 (mag x))))) < / 2

Valid_rnd (Znearest choice1)
now apply valid_rnd_N.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':x < midp fexp1 x
x':=round beta fexp1 Zfloor x:R
Hx:x - x' < / 2 * ulp beta fexp1 x
Pxx':0 <= x - x'
H:Rabs (x * bpow (- fexp1 (mag x)) - IZR (Zfloor (x * bpow (- fexp1 (mag x))))) < / 2

generic_format beta fexp1 (IZR (Zfloor (scaled_mantissa beta fexp1 x)) * bpow (fexp1 (mag x)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':x < midp fexp1 x
x':=round beta fexp1 Zfloor x:R
Hx:x - x' < / 2 * ulp beta fexp1 x
Pxx':0 <= x - x'
H:Rabs (x * bpow (- fexp1 (mag x)) - IZR (Zfloor (x * bpow (- fexp1 (mag x))))) < / 2

generic_format beta fexp1 (IZR (Zfloor (scaled_mantissa beta fexp1 x)) * bpow (cexp beta fexp1 x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':x < midp fexp1 x
x':=round beta fexp1 Zfloor x:R
Hx:x - x' < / 2 * ulp beta fexp1 x
Pxx':0 <= x - x'
H:Rabs (x * bpow (- fexp1 (mag x)) - IZR (Zfloor (x * bpow (- fexp1 (mag x))))) < / 2

generic_format beta fexp1 (round beta fexp1 Zfloor x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':x < midp fexp1 x
x':=round beta fexp1 Zfloor x:R
Hx:x - x' < / 2 * ulp beta fexp1 x
Pxx':0 <= x - x'
H:Rabs (x * bpow (- fexp1 (mag x)) - IZR (Zfloor (x * bpow (- fexp1 (mag x))))) < / 2

Valid_exp fexp1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':x < midp fexp1 x
x':=round beta fexp1 Zfloor x:R
Hx:x - x' < / 2 * ulp beta fexp1 x
Pxx':0 <= x - x'
H:Rabs (x * bpow (- fexp1 (mag x)) - IZR (Zfloor (x * bpow (- fexp1 (mag x))))) < / 2
Valid_rnd Zfloor
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':x < midp fexp1 x
x':=round beta fexp1 Zfloor x:R
Hx:x - x' < / 2 * ulp beta fexp1 x
Pxx':0 <= x - x'
H:Rabs (x * bpow (- fexp1 (mag x)) - IZR (Zfloor (x * bpow (- fexp1 (mag x))))) < / 2

Valid_rnd Zfloor
now apply valid_rnd_DN. Qed.
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall (choice1 choice2 : Z -> bool) (x : R), 0 < x -> (fexp2 (mag x) <= fexp1 (mag x))%Z -> (fexp1 (mag x) <= mag x)%Z -> x < midp fexp1 x -> ((fexp2 (mag x) <= fexp1 (mag x) - 1)%Z -> x < midp fexp1 x - / 2 * ulp beta fexp2 x) -> round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall (choice1 choice2 : Z -> bool) (x : R), 0 < x -> (fexp2 (mag x) <= fexp1 (mag x))%Z -> (fexp1 (mag x) <= mag x)%Z -> x < midp fexp1 x -> ((fexp2 (mag x) <= fexp1 (mag x) - 1)%Z -> x < midp fexp1 x - / 2 * ulp beta fexp2 x) -> round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx:x < midp fexp1 x
Hx':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z -> x < midp fexp1 x - / 2 * ulp beta fexp2 x

round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx:x < midp fexp1 x
Hx':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z -> x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hf2':(fexp1 (mag x) <= fexp2 (mag x))%Z

round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx:x < midp fexp1 x
Hx':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z -> x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hf2':(fexp2 (mag x) < fexp1 (mag x))%Z
round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx:x < midp fexp1 x
Hx':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z -> x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hf2':(fexp1 (mag x) <= fexp2 (mag x))%Z

round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx:x < midp fexp1 x
Hx':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z -> x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hf2':(fexp1 (mag x) <= fexp2 (mag x))%Z
Hf2'':fexp2 (mag x) = fexp1 (mag x)

round_round_eq fexp1 fexp2 choice1 choice2 x
now apply round_round_lt_mid_same_place.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx:x < midp fexp1 x
Hx':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z -> x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hf2':(fexp2 (mag x) < fexp1 (mag x))%Z

round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx:x < midp fexp1 x
Hx':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z -> x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hf2':(fexp2 (mag x) < fexp1 (mag x))%Z
Hf2'':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z

round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx:x < midp fexp1 x
Hx':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z -> x < midp fexp1 x - / 2 * ulp beta fexp2 x
Hf2':(fexp2 (mag x) < fexp1 (mag x))%Z
Hf2'':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx'':x < midp fexp1 x - / 2 * ulp beta fexp2 x

round_round_eq fexp1 fexp2 choice1 choice2 x
now apply round_round_lt_mid_further_place. Qed.
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall (choice1 choice2 : Z -> bool) (x : R), 0 < x -> (fexp2 (mag x) <= fexp1 (mag x) - 1)%Z -> round beta fexp2 (Znearest choice2) x < bpow (mag x) -> midp' fexp1 x + / 2 * ulp beta fexp2 x < x -> round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall (choice1 choice2 : Z -> bool) (x : R), 0 < x -> (fexp2 (mag x) <= fexp1 (mag x) - 1)%Z -> round beta fexp2 (Znearest choice2) x < bpow (mag x) -> midp' fexp1 x + / 2 * ulp beta fexp2 x < x -> round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z

round beta fexp2 (Znearest choice2) x < bpow (mag x) -> midp' fexp1 x + / 2 * ulp beta fexp2 x < x -> round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:round beta fexp2 (Znearest choice2) x < bpow (mag x)
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x

round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:round beta fexp2 (Znearest choice2) x < bpow (mag x)
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x

round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:round beta fexp2 (Znearest choice2) x < bpow (mag x)
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:round beta fexp2 (Znearest choice2) x < bpow (mag x)
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x

round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:round beta fexp2 (Znearest choice2) x < bpow (mag x)
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x

round beta fexp1 Zceil x - / 2 * ulp beta fexp1 x + / 2 * ulp beta fexp2 x < x
now unfold midp' in Hx2'.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx1:round beta fexp2 (Znearest choice2) x < bpow (mag x)
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)

round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x

round beta fexp2 (Znearest choice2) x < bpow (mag x) -> round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x) -> round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x

round beta fexp2 (Znearest choice2) x < bpow (mag x) -> round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x) -> round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R

round beta fexp2 (Znearest choice2) x < bpow (mag x) -> x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x) -> round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R

x'' < bpow (mag x) -> x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x) -> round beta fexp1 (Znearest choice1) x'' = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)

round beta fexp1 (Znearest choice1) x'' = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)

Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
round beta fexp1 (Znearest choice1) x'' = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)

Rabs (x'' - x) <= / 2 * ulp beta fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
/ 2 * ulp beta fexp2 x <= / 2 * bpow (fexp2 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
round beta fexp1 (Znearest choice1) x'' = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)

/ 2 * ulp beta fexp2 x <= / 2 * bpow (fexp2 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
round beta fexp1 (Znearest choice1) x'' = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))

round beta fexp1 (Znearest choice1) x'' = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))

0 <= x' - x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
round beta fexp1 (Znearest choice1) x'' = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))

0 <= x' - x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))

x <= x'
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))

Valid_exp fexp1
exact Vfexp1.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x

round beta fexp1 (Znearest choice1) x'' = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x

Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
round beta fexp1 (Znearest choice1) x'' = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x

Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x

Rabs (x'' - x + (x - x')) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x

Rabs (x'' - x) + Rabs (x - x') < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x

Rabs (x'' - x) + Rabs (x - x') < / 2 * bpow (fexp2 (mag x)) + / 2 * (bpow (fexp1 (mag x)) - bpow (fexp2 (mag x)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x

Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Rabs (x - x') < / 2 * (bpow (fexp1 (mag x)) - bpow (fexp2 (mag x)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x

Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
exact Hr1.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x

Rabs (x - x') < / 2 * (bpow (fexp1 (mag x)) - bpow (fexp2 (mag x)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x

Rabs (x' - x) < / 2 * (bpow (fexp1 (mag x)) - bpow (fexp2 (mag x)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x))
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x

Rabs (x' - x) < / 2 * (bpow (fexp1 (mag x)) - bpow (fexp2 (mag x)))
now rewrite Rabs_right; [|now apply Rle_ge]; apply Hx2.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))

round beta fexp1 (Znearest choice1) x'' = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Zx'':x'' = 0

round beta fexp1 (Znearest choice1) x'' = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
round beta fexp1 (Znearest choice1) x'' = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Zx'':x'' = 0

round beta fexp1 (Znearest choice1) x'' = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (0 - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Zx'':x'' = 0

round beta fexp1 (Znearest choice1) 0 = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (0 - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Zx'':x'' = 0

0 = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (0 - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Zx'':x'' = 0

0 = IZR (Znearest choice1 (x * bpow (- fexp1 (mag x)))) * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (0 - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Zx'':x'' = 0

Rabs (x * bpow (- fexp1 (mag x)) - 0) < / 2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (0 - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Zx'':x'' = 0

Rabs (x * bpow (- fexp1 (mag x)) - 0) * bpow (fexp1 (mag x)) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (0 - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Zx'':x'' = 0

Rabs (x * bpow (- fexp1 (mag x)) - 0) * Rabs (bpow (fexp1 (mag x))) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (0 - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Zx'':x'' = 0

Rabs (x * bpow (- fexp1 (mag x)) * bpow (fexp1 (mag x)) - 0 * bpow (fexp1 (mag x))) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (0 - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Zx'':x'' = 0

Rabs (x * bpow (- fexp1 (mag x)) * bpow (fexp1 (mag x)) - 0) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (0 - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Zx'':x'' = 0

Rabs (x - 0) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (0 - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Zx'':x'' = 0

Rabs (0 - x) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (0 - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Zx'':x'' = 0

/ 2 * bpow (fexp2 (mag x)) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (0 - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Zx'':x'' = 0

bpow (fexp2 (mag x)) < bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (0 - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Zx'':x'' = 0

(fexp2 (mag x) < fexp1 (mag x))%Z
omega.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0

round beta fexp1 (Znearest choice1) x'' = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0

mag x'' = mag x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
Lx'':mag x'' = mag x
round beta fexp1 (Znearest choice1) x'' = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0

mag x'' = mag x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0

(mag x'' <= mag x)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
(mag x <= mag x'')%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0

(mag x'' <= mag x)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0

Rabs x'' < bpow (mag x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0

0 <= x''
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0

Valid_exp fexp2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
Valid_rnd (Znearest choice2)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
generic_format beta fexp2 0
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
0 <= x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0

Valid_exp fexp2
exact Vfexp2.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0

Valid_rnd (Znearest choice2)
now apply valid_rnd_N.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0

generic_format beta fexp2 0
apply generic_format_0.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0

0 <= x
now apply Rlt_le.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0

(mag x <= mag x'')%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':round beta fexp2 (Znearest choice2) x <> 0

(mag x <= mag (round beta fexp2 (Znearest choice2) x))%Z
now apply mag_round_ge; [|apply valid_rnd_N|].
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
Lx'':mag x'' = mag x

round beta fexp1 (Znearest choice1) x'' = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
Lx'':mag x'' = mag x

IZR (Znearest choice1 (x'' * bpow (- fexp1 (mag x'')))) * bpow (fexp1 (mag x'')) = IZR (Znearest choice1 (x * bpow (- fexp1 (mag x)))) * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
Lx'':mag x'' = mag x

IZR (Znearest choice1 (x'' * bpow (- fexp1 (mag x)))) * bpow (fexp1 (mag x)) = IZR (Znearest choice1 (x * bpow (- fexp1 (mag x)))) * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
Lx'':mag x'' = mag x

IZR (Zceil (scaled_mantissa beta fexp1 x)) * bpow (fexp1 (mag x)) = IZR (Znearest choice1 (x * bpow (- fexp1 (mag x)))) * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
Lx'':mag x'' = mag x
Rabs (x'' * bpow (- fexp1 (mag x)) - IZR (Zceil (scaled_mantissa beta fexp1 x))) < / 2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
Lx'':mag x'' = mag x

IZR (Zceil (scaled_mantissa beta fexp1 x)) * bpow (fexp1 (mag x)) = IZR (Znearest choice1 (x * bpow (- fexp1 (mag x)))) * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
Lx'':mag x'' = mag x

Rabs (x * bpow (- fexp1 (mag x)) - IZR (Zceil (scaled_mantissa beta fexp1 x))) < / 2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
Lx'':mag x'' = mag x

Rabs (x * bpow (- fexp1 (mag x)) - IZR (Zceil (scaled_mantissa beta fexp1 x))) * bpow (fexp1 (mag x)) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
Lx'':mag x'' = mag x

Rabs (x * bpow (- fexp1 (mag x)) - IZR (Zceil (scaled_mantissa beta fexp1 x))) * Rabs (bpow (fexp1 (mag x))) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
Lx'':mag x'' = mag x

Rabs ((x * bpow (- fexp1 (mag x)) - IZR (Zceil (scaled_mantissa beta fexp1 x))) * bpow (fexp1 (mag x))) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
Lx'':mag x'' = mag x

Rabs (x * bpow (- fexp1 (mag x)) * bpow (fexp1 (mag x)) - IZR (Zceil (scaled_mantissa beta fexp1 x)) * bpow (fexp1 (mag x))) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
Lx'':mag x'' = mag x

Rabs (x * bpow (- fexp1 (mag x)) * bpow (fexp1 (mag x)) - IZR (Zceil (scaled_mantissa beta fexp1 x)) * bpow (fexp1 (mag x))) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
Lx'':mag x'' = mag x

Rabs (x - IZR (Zceil (scaled_mantissa beta fexp1 x)) * bpow (fexp1 (mag x))) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
Lx'':mag x'' = mag x

Rabs (IZR (Zceil (scaled_mantissa beta fexp1 x)) * bpow (fexp1 (mag x)) - x) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
Lx'':mag x'' = mag x

IZR (Zceil (scaled_mantissa beta fexp1 x)) * bpow (fexp1 (mag x)) - x < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
Lx'':mag x'' = mag x

/ 2 * (ulp beta fexp1 x - ulp beta fexp2 x) <= / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
Lx'':mag x'' = mag x

ulp beta fexp1 x - ulp beta fexp2 x <= bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
Lx'':mag x'' = mag x

0 <= bpow (fexp2 (mag x)) -> ulp beta fexp1 x - ulp beta fexp2 x <= bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
Lx'':mag x'' = mag x

0 <= bpow (fexp2 (mag x)) -> bpow (cexp beta fexp1 x) - bpow (cexp beta fexp2 x) <= bpow (fexp1 (mag x))
unfold cexp; lra.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
Lx'':mag x'' = mag x

Rabs (x'' * bpow (- fexp1 (mag x)) - IZR (Zceil (scaled_mantissa beta fexp1 x))) < / 2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
Lx'':mag x'' = mag x

Rabs (x'' * bpow (- fexp1 (mag x)) - IZR (Zceil (scaled_mantissa beta fexp1 x))) * bpow (fexp1 (mag x)) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
Lx'':mag x'' = mag x

Rabs (x'' * bpow (- fexp1 (mag x)) - IZR (Zceil (scaled_mantissa beta fexp1 x))) * Rabs (bpow (fexp1 (mag x))) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
Lx'':mag x'' = mag x

Rabs ((x'' * bpow (- fexp1 (mag x)) - IZR (Zceil (scaled_mantissa beta fexp1 x))) * bpow (fexp1 (mag x))) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
Lx'':mag x'' = mag x

Rabs (x'' * bpow (- fexp1 (mag x)) * bpow (fexp1 (mag x)) - IZR (Zceil (scaled_mantissa beta fexp1 x)) * bpow (fexp1 (mag x))) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zceil x:R
x'':=round beta fexp2 (Znearest choice2) x:R
Hx1:x'' < bpow (mag x)
Hx2:x' - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
Hr1:Rabs (x'' - x) <= / 2 * bpow (fexp2 (mag x))
Px'x:0 <= x' - x
Hr2:Rabs (x'' - x') < / 2 * bpow (fexp1 (mag x))
Nzx'':x'' <> 0
Lx'':mag x'' = mag x

Rabs (x'' * bpow (- fexp1 (mag x)) * bpow (fexp1 (mag x)) - IZR (Zceil (scaled_mantissa beta fexp1 x)) * bpow (fexp1 (mag x))) < / 2 * bpow (fexp1 (mag x))
now bpow_simplify. Qed.
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall (choice1 choice2 : Z -> bool) (x : R), 0 < x -> (fexp2 (mag x) <= fexp1 (mag x) - 1)%Z -> (fexp1 (mag x) <= mag x)%Z -> midp' fexp1 x + / 2 * ulp beta fexp2 x < x -> round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall (choice1 choice2 : Z -> bool) (x : R), 0 < x -> (fexp2 (mag x) <= fexp1 (mag x) - 1)%Z -> (fexp1 (mag x) <= mag x)%Z -> midp' fexp1 x + / 2 * ulp beta fexp2 x < x -> round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x

round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x

round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x

round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x

round beta fexp1 Zceil x - / 2 * ulp beta fexp1 x + / 2 * ulp beta fexp2 x < x
now unfold midp' in Hx2'.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)

round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x

round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x) -> round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x

round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x) -> round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R

round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x) -> round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R

round beta fexp1 (Znearest choice1) x'' = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''

round beta fexp1 (Znearest choice1) x'' = round beta fexp1 (Znearest choice1) x
(* bpow (mag x) <= x'' *)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''

x'' = bpow (mag x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
Hx''pow:x'' = bpow (mag x)
round beta fexp1 (Znearest choice1) x'' = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''

x'' = bpow (mag x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''

x'' < bpow (mag x) + / 2 * ulp beta fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
H'x'':x'' < bpow (mag x) + / 2 * ulp beta fexp2 x
x'' = bpow (mag x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''

x'' < bpow (mag x) + / 2 * ulp beta fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''

x'' <= x + / 2 * ulp beta fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
x + / 2 * ulp beta fexp2 x < bpow (mag x) + / 2 * ulp beta fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''

x'' <= x + / 2 * ulp beta fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''

x'' - x <= / 2 * ulp beta fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''

Rabs (x'' - x) <= / 2 * ulp beta fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''

Valid_exp fexp2
exact Vfexp2.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''

x + / 2 * ulp beta fexp2 x < bpow (mag x) + / 2 * ulp beta fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''

x < bpow (mag x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''

Rabs x < bpow (mag x)
apply bpow_mag_gt.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
H'x'':x'' < bpow (mag x) + / 2 * ulp beta fexp2 x

x'' = bpow (mag x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
H'x'':x'' < bpow (mag x) + / 2 * ulp beta fexp2 x

x'' <= bpow (mag x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
H'x'':x'' < bpow (mag x) + / 2 * ulp beta fexp2 x

IZR (Znearest choice2 (x * bpow (- fexp2 (mag x)))) * bpow (fexp2 (mag x)) <= bpow (mag x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
H'x'':x'' < bpow (mag x) + / 2 * ulp beta fexp2 x

IZR (Znearest choice2 (x * bpow (- fexp2 (mag x)))) * bpow (fexp2 (mag x)) * bpow (- fexp2 (mag x)) <= bpow (mag x) * bpow (- fexp2 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
H'x'':x'' < bpow (mag x) + / 2 * ulp beta fexp2 x

IZR (Znearest choice2 (x * bpow (- fexp2 (mag x)))) <= bpow (mag x - fexp2 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
H'x'':x'' < bpow (mag x) + / 2 * ulp beta fexp2 x

IZR (Znearest choice2 (x * bpow (- fexp2 (mag x)))) <= IZR (beta ^ (mag x - fexp2 (mag x)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
H'x'':x'' < bpow (mag x) + / 2 * ulp beta fexp2 x

(Znearest choice2 (x * bpow (- fexp2 (mag x))) <= beta ^ (mag x - fexp2 (mag x)))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
H'x'':x'' < bpow (mag x) + / 2 * ulp beta fexp2 x

(Znearest choice2 (x * bpow (- fexp2 (mag x))) < beta ^ (mag x - fexp2 (mag x)) + 1)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
H'x'':x'' < bpow (mag x) + / 2 * ulp beta fexp2 x

IZR (Znearest choice2 (x * bpow (- fexp2 (mag x)))) < IZR (beta ^ (mag x - fexp2 (mag x)) + 1)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
H'x'':x'' < bpow (mag x) + / 2 * ulp beta fexp2 x

IZR (Znearest choice2 (x * bpow (- fexp2 (mag x)))) < bpow (mag x - fexp2 (mag x)) + 1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
H'x'':x'' < bpow (mag x) + / 2 * ulp beta fexp2 x

IZR (Znearest choice2 (x * bpow (- fexp2 (mag x)))) * bpow (fexp2 (mag x)) < (bpow (mag x - fexp2 (mag x)) + 1) * bpow (fexp2 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
H'x'':x'' < bpow (mag x) + / 2 * ulp beta fexp2 x

IZR (Znearest choice2 (x * bpow (- fexp2 (mag x)))) * bpow (fexp2 (mag x)) < bpow (mag x - fexp2 (mag x)) * bpow (fexp2 (mag x)) + bpow (fexp2 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
H'x'':x'' < bpow (mag x) + / 2 * ulp beta fexp2 x

IZR (Znearest choice2 (x * bpow (- fexp2 (mag x)))) * bpow (fexp2 (mag x)) < bpow (mag x) + bpow (fexp2 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
H'x'':x'' < bpow (mag x) + / 2 * ulp beta fexp2 x

bpow (mag x) + / 2 * ulp beta fexp2 x <= bpow (mag x) + bpow (fexp2 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
H'x'':x'' < bpow (mag x) + / 2 * ulp beta fexp2 x

/ 2 * ulp beta fexp2 x <= bpow (fexp2 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
H'x'':x'' < bpow (mag x) + / 2 * ulp beta fexp2 x

/ 2 * ulp beta fexp2 x <= 1 * bpow (fexp2 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
H'x'':x'' < bpow (mag x) + / 2 * ulp beta fexp2 x

/ 2 * bpow (cexp beta fexp2 x) <= 1 * bpow (fexp2 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
H'x'':x'' < bpow (mag x) + / 2 * ulp beta fexp2 x

/ 2 <= 1
lra.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
Hx''pow:x'' = bpow (mag x)

round beta fexp1 (Znearest choice1) x'' = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
Hx''pow:x'' = bpow (mag x)

Rabs (x - x'') < / 2 * ulp beta fexp1 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
Hx''pow:x'' = bpow (mag x)
Hr:Rabs (x - x'') < / 2 * ulp beta fexp1 x
round beta fexp1 (Znearest choice1) x'' = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
Hx''pow:x'' = bpow (mag x)

Rabs (x - x'') < / 2 * ulp beta fexp1 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
Hx''pow:x'' = bpow (mag x)

Rabs (x - x'') <= / 2 * ulp beta fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
Hx''pow:x'' = bpow (mag x)
/ 2 * ulp beta fexp2 x < / 2 * ulp beta fexp1 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
Hx''pow:x'' = bpow (mag x)

Rabs (x - x'') <= / 2 * ulp beta fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
Hx''pow:x'' = bpow (mag x)

Rabs (x'' - x) <= / 2 * ulp beta fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
Hx''pow:x'' = bpow (mag x)

Valid_exp fexp2
exact Vfexp2.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
Hx''pow:x'' = bpow (mag x)

/ 2 * ulp beta fexp2 x < / 2 * ulp beta fexp1 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
Hx''pow:x'' = bpow (mag x)

ulp beta fexp2 x < ulp beta fexp1 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
Hx''pow:x'' = bpow (mag x)

bpow (cexp beta fexp2 x) < bpow (cexp beta fexp1 x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
Hx''pow:x'' = bpow (mag x)

(fexp2 (mag x) < fexp1 (mag x))%Z
omega.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
Hx''pow:x'' = bpow (mag x)
Hr:Rabs (x - x'') < / 2 * ulp beta fexp1 x

round beta fexp1 (Znearest choice1) x'' = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
Hx''pow:x'' = bpow (mag x)
Hr:Rabs (x - x'') < / 2 * ulp beta fexp1 x

IZR (Znearest choice1 (x'' * bpow (- fexp1 (mag x'')))) * bpow (fexp1 (mag x'')) = IZR (Znearest choice1 (x * bpow (- fexp1 (mag x)))) * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
Hx''pow:x'' = bpow (mag x)
Hr:Rabs (x - x'') < / 2 * ulp beta fexp1 x

(0 <= mag x - fexp1 (mag x''))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
Hx''pow:x'' = bpow (mag x)
Hr:Rabs (x - x'') < / 2 * ulp beta fexp1 x
Hf:(0 <= mag x - fexp1 (mag x''))%Z
IZR (Znearest choice1 (x'' * bpow (- fexp1 (mag x'')))) * bpow (fexp1 (mag x'')) = IZR (Znearest choice1 (x * bpow (- fexp1 (mag x)))) * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
Hx''pow:x'' = bpow (mag x)
Hr:Rabs (x - x'') < / 2 * ulp beta fexp1 x

(0 <= mag x - fexp1 (mag x''))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
Hx''pow:x'' = bpow (mag x)
Hr:Rabs (x - x'') < / 2 * ulp beta fexp1 x

(0 <= mag x - fexp1 (mag (bpow (mag x))))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
Hx''pow:x'' = bpow (mag x)
Hr:Rabs (x - x'') < / 2 * ulp beta fexp1 x

(0 <= mag x - fexp1 (mag x + 1))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
Hx''pow:x'' = bpow (mag x)
Hr:Rabs (x - x'') < / 2 * ulp beta fexp1 x

(fexp1 (mag x + 1) <= mag x)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
Hx''pow:x'' = bpow (mag x)
Hr:Rabs (x - x'') < / 2 * ulp beta fexp1 x
Hle:(mag x <= fexp1 (mag x))%Z

(fexp1 (mag x + 1) <= mag x)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
Hx''pow:x'' = bpow (mag x)
Hr:Rabs (x - x'') < / 2 * ulp beta fexp1 x
Hle:(mag x <= fexp1 (mag x))%Z
H:mag x = fexp1 (mag x)

(fexp1 (mag x + 1) <= mag x)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
Hx''pow:x'' = bpow (mag x)
Hr:Rabs (x - x'') < / 2 * ulp beta fexp1 x
Hle:(mag x <= fexp1 (mag x))%Z
H:mag x = fexp1 (mag x)

(fexp1 (fexp1 (mag x) + 1) <= fexp1 (mag x))%Z
now apply Vfexp1.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
Hx''pow:x'' = bpow (mag x)
Hr:Rabs (x - x'') < / 2 * ulp beta fexp1 x
Hf:(0 <= mag x - fexp1 (mag x''))%Z

IZR (Znearest choice1 (x'' * bpow (- fexp1 (mag x'')))) * bpow (fexp1 (mag x'')) = IZR (Znearest choice1 (x * bpow (- fexp1 (mag x)))) * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
Hx''pow:x'' = bpow (mag x)
Hr:Rabs (x - x'') < / 2 * ulp beta fexp1 x
Hf:(0 <= mag x - fexp1 (mag x''))%Z

IZR (beta ^ (mag x - fexp1 (mag x''))) * bpow (fexp1 (mag x'')) = IZR (Znearest choice1 (x * bpow (- fexp1 (mag x)))) * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
Hx''pow:x'' = bpow (mag x)
Hr:Rabs (x - x'') < / 2 * ulp beta fexp1 x
Hf:(0 <= mag x - fexp1 (mag x''))%Z
Rabs (x'' * bpow (- fexp1 (mag x'')) - IZR (beta ^ (mag x - fexp1 (mag x'')))) < / 2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
Hx''pow:x'' = bpow (mag x)
Hr:Rabs (x - x'') < / 2 * ulp beta fexp1 x
Hf:(0 <= mag x - fexp1 (mag x''))%Z

IZR (beta ^ (mag x - fexp1 (mag x''))) * bpow (fexp1 (mag x'')) = IZR (Znearest choice1 (x * bpow (- fexp1 (mag x)))) * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
Hx''pow:x'' = bpow (mag x)
Hr:Rabs (x - x'') < / 2 * ulp beta fexp1 x
Hf:(0 <= mag x - fexp1 (mag x''))%Z

IZR (beta ^ (mag x - fexp1 (mag x''))) * bpow (fexp1 (mag x'')) = IZR (beta ^ (mag x - fexp1 (mag x))) * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
Hx''pow:x'' = bpow (mag x)
Hr:Rabs (x - x'') < / 2 * ulp beta fexp1 x
Hf:(0 <= mag x - fexp1 (mag x''))%Z
Rabs (x * bpow (- fexp1 (mag x)) - IZR (beta ^ (mag x - fexp1 (mag x)))) < / 2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
Hx''pow:x'' = bpow (mag x)
Hr:Rabs (x - x'') < / 2 * ulp beta fexp1 x
Hf:(0 <= mag x - fexp1 (mag x''))%Z

IZR (beta ^ (mag x - fexp1 (mag x''))) * bpow (fexp1 (mag x'')) = IZR (beta ^ (mag x - fexp1 (mag x))) * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
Hx''pow:x'' = bpow (mag x)
Hr:Rabs (x - x'') < / 2 * ulp beta fexp1 x
Hf:(0 <= mag x - fexp1 (mag x''))%Z

bpow (mag x - fexp1 (mag x'')) * bpow (fexp1 (mag x'')) = IZR (beta ^ (mag x - fexp1 (mag x))) * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
Hx''pow:x'' = bpow (mag x)
Hr:Rabs (x - x'') < / 2 * ulp beta fexp1 x
Hf:(0 <= mag x - fexp1 (mag x''))%Z

bpow (mag x - fexp1 (mag x'')) * bpow (fexp1 (mag x'')) = bpow (mag x - fexp1 (mag x)) * bpow (fexp1 (mag x))
now bpow_simplify.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
Hx''pow:x'' = bpow (mag x)
Hr:Rabs (x - x'') < / 2 * ulp beta fexp1 x
Hf:(0 <= mag x - fexp1 (mag x''))%Z

Rabs (x * bpow (- fexp1 (mag x)) - IZR (beta ^ (mag x - fexp1 (mag x)))) < / 2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
Hx''pow:x'' = bpow (mag x)
Hr:Rabs (x - x'') < / 2 * ulp beta fexp1 x
Hf:(0 <= mag x - fexp1 (mag x''))%Z

Rabs (x * bpow (- fexp1 (mag x)) - bpow (mag x - fexp1 (mag x))) < / 2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
Hx''pow:x'' = bpow (mag x)
Hr:Rabs (x - x'') < / 2 * ulp beta fexp1 x
Hf:(0 <= mag x - fexp1 (mag x''))%Z

Rabs (x * bpow (- fexp1 (mag x)) - bpow (mag x - fexp1 (mag x))) * bpow (fexp1 (mag x)) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
Hx''pow:x'' = bpow (mag x)
Hr:Rabs (x - x'') < / 2 * ulp beta fexp1 x
Hf:(0 <= mag x - fexp1 (mag x''))%Z

Rabs (x * bpow (- fexp1 (mag x)) - bpow (mag x - fexp1 (mag x))) * Rabs (bpow (fexp1 (mag x))) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
Hx''pow:x'' = bpow (mag x)
Hr:Rabs (x - x'') < / 2 * ulp beta fexp1 x
Hf:(0 <= mag x - fexp1 (mag x''))%Z

Rabs ((x * bpow (- fexp1 (mag x)) - bpow (mag x - fexp1 (mag x))) * bpow (fexp1 (mag x))) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
Hx''pow:x'' = bpow (mag x)
Hr:Rabs (x - x'') < / 2 * ulp beta fexp1 x
Hf:(0 <= mag x - fexp1 (mag x''))%Z

Rabs (x * bpow (- fexp1 (mag x)) * bpow (fexp1 (mag x)) - bpow (mag x - fexp1 (mag x)) * bpow (fexp1 (mag x))) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
Hx''pow:x'' = bpow (mag x)
Hr:Rabs (x - x'') < / 2 * ulp beta fexp1 x
Hf:(0 <= mag x - fexp1 (mag x''))%Z

Rabs (x - bpow (mag x)) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
Hx''pow:x'' = bpow (mag x)
Hr:Rabs (x - x'') < / 2 * bpow (cexp beta fexp1 x)
Hf:(0 <= mag x - fexp1 (mag x''))%Z

Rabs (x - bpow (mag x)) < / 2 * bpow (fexp1 (mag x))
rewrite <- Hx''pow; exact Hr.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
Hx''pow:x'' = bpow (mag x)
Hr:Rabs (x - x'') < / 2 * ulp beta fexp1 x
Hf:(0 <= mag x - fexp1 (mag x''))%Z

Rabs (x'' * bpow (- fexp1 (mag x'')) - IZR (beta ^ (mag x - fexp1 (mag x'')))) < / 2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
Hx''pow:x'' = bpow (mag x)
Hr:Rabs (x - x'') < / 2 * ulp beta fexp1 x
Hf:(0 <= mag x - fexp1 (mag x''))%Z

Rabs (x'' * bpow (- fexp1 (mag x'')) - bpow (mag x - fexp1 (mag x''))) < / 2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
Hx''pow:x'' = bpow (mag x)
Hr:Rabs (x - x'') < / 2 * ulp beta fexp1 x
Hf:(0 <= mag x - fexp1 (mag x''))%Z

Rabs (x'' * bpow (- fexp1 (mag x'')) - bpow (mag x - fexp1 (mag x''))) * bpow (fexp1 (mag x'')) < / 2 * bpow (fexp1 (mag x''))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
Hx''pow:x'' = bpow (mag x)
Hr:Rabs (x - x'') < / 2 * ulp beta fexp1 x
Hf:(0 <= mag x - fexp1 (mag x''))%Z

Rabs (x'' * bpow (- fexp1 (mag x'')) - bpow (mag x - fexp1 (mag x''))) * Rabs (bpow (fexp1 (mag x''))) < / 2 * bpow (fexp1 (mag x''))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
Hx''pow:x'' = bpow (mag x)
Hr:Rabs (x - x'') < / 2 * ulp beta fexp1 x
Hf:(0 <= mag x - fexp1 (mag x''))%Z

Rabs ((x'' * bpow (- fexp1 (mag x'')) - bpow (mag x - fexp1 (mag x''))) * bpow (fexp1 (mag x''))) < / 2 * bpow (fexp1 (mag x''))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
Hx''pow:x'' = bpow (mag x)
Hr:Rabs (x - x'') < / 2 * ulp beta fexp1 x
Hf:(0 <= mag x - fexp1 (mag x''))%Z

Rabs (x'' * bpow (- fexp1 (mag x'')) * bpow (fexp1 (mag x'')) - bpow (mag x - fexp1 (mag x'')) * bpow (fexp1 (mag x''))) < / 2 * bpow (fexp1 (mag x''))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
Hx''pow:x'' = bpow (mag x)
Hr:Rabs (x - x'') < / 2 * ulp beta fexp1 x
Hf:(0 <= mag x - fexp1 (mag x''))%Z

Rabs (x'' - bpow (mag x)) < / 2 * bpow (fexp1 (mag x''))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
Hx''pow:x'' = bpow (mag x)
Hr:Rabs (x - x'') < / 2 * ulp beta fexp1 x
Hf:(0 <= mag x - fexp1 (mag x''))%Z

Rabs 0 < / 2 * bpow (fexp1 (mag x''))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
Hx''pow:x'' = bpow (mag x)
Hr:Rabs (x - x'') < / 2 * ulp beta fexp1 x
Hf:(0 <= mag x - fexp1 (mag x''))%Z

0 < / 2 * bpow (fexp1 (mag x''))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx2':midp' fexp1 x + / 2 * ulp beta fexp2 x < x
x':=round beta fexp1 Zfloor x:R
Hx2:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
x'':=round beta fexp2 (Znearest choice2) x:R
Hx'':bpow (mag x) <= x''
Hx''pow:x'' = bpow (mag x)
Hr:Rabs (x - x'') < / 2 * ulp beta fexp1 x
Hf:(0 <= mag x - fexp1 (mag x''))%Z

/ 2 * 0 < / 2 * bpow (fexp1 (mag x''))
apply Rmult_lt_compat_l; [lra|apply bpow_gt_0]. Qed.
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> forall (choice1 choice2 : Z -> bool) (x : R), 0 < x -> fexp2 (mag x) = fexp1 (mag x) -> midp' fexp1 x < x -> round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> forall (choice1 choice2 : Z -> bool) (x : R), 0 < x -> fexp2 (mag x) = fexp1 (mag x) -> midp' fexp1 x < x -> round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':midp' fexp1 x < x

round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':midp' fexp1 x < x

round beta fexp1 Zceil x - x < / 2 * ulp beta fexp1 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':midp' fexp1 x < x
Hx:round beta fexp1 Zceil x - x < / 2 * ulp beta fexp1 x
round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':midp' fexp1 x < x

round beta fexp1 Zceil x - x < / 2 * ulp beta fexp1 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':midp' fexp1 x < x

round beta fexp1 Zceil x - / 2 * ulp beta fexp1 x < x
now unfold midp' in Hx'.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':midp' fexp1 x < x
Hx:round beta fexp1 Zceil x - x < / 2 * ulp beta fexp1 x

round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':midp' fexp1 x < x
Hx:round beta fexp1 Zceil x - x < / 2 * ulp beta fexp1 x

Rabs (IZR (Zceil (x * bpow (- fexp1 (mag x)))) - x * bpow (- fexp1 (mag x))) < / 2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':midp' fexp1 x < x
Hx:round beta fexp1 Zceil x - x < / 2 * ulp beta fexp1 x
H:Rabs (IZR (Zceil (x * bpow (- fexp1 (mag x)))) - x * bpow (- fexp1 (mag x))) < / 2
round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':midp' fexp1 x < x
Hx:round beta fexp1 Zceil x - x < / 2 * ulp beta fexp1 x

Rabs (IZR (Zceil (x * bpow (- fexp1 (mag x)))) - x * bpow (- fexp1 (mag x))) < / 2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':midp' fexp1 x < x
Hx:round beta fexp1 Zceil x - x < / 2 * ulp beta fexp1 x

Rabs (IZR (Zceil (x * bpow (- fexp1 (mag x)))) - x * bpow (- fexp1 (mag x))) * bpow (fexp1 (mag x)) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':midp' fexp1 x < x
Hx:round beta fexp1 Zceil x - x < / 2 * ulp beta fexp1 x

Rabs (IZR (Zceil (x * bpow (- fexp1 (mag x)))) - x * bpow (- fexp1 (mag x))) * bpow (fexp1 (mag x)) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':midp' fexp1 x < x
Hx:round beta fexp1 Zceil x - x < / 2 * ulp beta fexp1 x

Rabs (IZR (Zceil (x * bpow (- fexp1 (mag x)))) - x * bpow (- fexp1 (mag x))) * Rabs (bpow (fexp1 (mag x))) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':midp' fexp1 x < x
Hx:round beta fexp1 Zceil x - x < / 2 * ulp beta fexp1 x

Rabs ((IZR (Zceil (x * bpow (- fexp1 (mag x)))) - x * bpow (- fexp1 (mag x))) * bpow (fexp1 (mag x))) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':midp' fexp1 x < x
Hx:round beta fexp1 Zceil x - x < / 2 * ulp beta fexp1 x

Rabs (IZR (Zceil (x * bpow (- fexp1 (mag x)))) * bpow (fexp1 (mag x)) - x * bpow (- fexp1 (mag x)) * bpow (fexp1 (mag x))) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':midp' fexp1 x < x
Hx:round beta fexp1 Zceil x - x < / 2 * ulp beta fexp1 x

Rabs (IZR (Zceil (x * bpow (- fexp1 (mag x)))) * bpow (fexp1 (mag x)) - x) < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':midp' fexp1 x < x
Hx:round beta fexp1 Zceil x - x < / 2 * ulp beta fexp1 x

- (/ 2 * bpow (fexp1 (mag x))) < IZR (Zceil (x * bpow (- fexp1 (mag x)))) * bpow (fexp1 (mag x)) - x < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':midp' fexp1 x < x
Hx:round beta fexp1 Zceil x - x < / 2 * ulp beta fexp1 x

- (/ 2 * bpow (fexp1 (mag x))) < IZR (Zceil (x * bpow (- fexp1 (mag x)))) * bpow (fexp1 (mag x)) - x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':midp' fexp1 x < x
Hx:round beta fexp1 Zceil x - x < / 2 * ulp beta fexp1 x
IZR (Zceil (x * bpow (- fexp1 (mag x)))) * bpow (fexp1 (mag x)) - x < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':midp' fexp1 x < x
Hx:round beta fexp1 Zceil x - x < / 2 * ulp beta fexp1 x

- (/ 2 * bpow (fexp1 (mag x))) < IZR (Zceil (x * bpow (- fexp1 (mag x)))) * bpow (fexp1 (mag x)) - x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':midp' fexp1 x < x
Hx:round beta fexp1 Zceil x - x < / 2 * ulp beta fexp1 x

- (/ 2 * bpow (fexp1 (mag x))) < 0
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':midp' fexp1 x < x
Hx:round beta fexp1 Zceil x - x < / 2 * ulp beta fexp1 x
0 <= IZR (Zceil (x * bpow (- fexp1 (mag x)))) * bpow (fexp1 (mag x)) - x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':midp' fexp1 x < x
Hx:round beta fexp1 Zceil x - x < / 2 * ulp beta fexp1 x

- (/ 2 * bpow (fexp1 (mag x))) < 0
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':midp' fexp1 x < x
Hx:round beta fexp1 Zceil x - x < / 2 * ulp beta fexp1 x

0 < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':midp' fexp1 x < x
Hx:round beta fexp1 Zceil x - x < / 2 * ulp beta fexp1 x

/ 2 * 0 < / 2 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':midp' fexp1 x < x
Hx:round beta fexp1 Zceil x - x < / 2 * ulp beta fexp1 x

0 < bpow (fexp1 (mag x))
apply bpow_gt_0.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':midp' fexp1 x < x
Hx:round beta fexp1 Zceil x - x < / 2 * ulp beta fexp1 x

0 <= IZR (Zceil (x * bpow (- fexp1 (mag x)))) * bpow (fexp1 (mag x)) - x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':midp' fexp1 x < x
Hx:round beta fexp1 Zceil x - x < / 2 * ulp beta fexp1 x

x <= IZR (Zceil (x * bpow (- fexp1 (mag x)))) * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':midp' fexp1 x < x
Hx:round beta fexp1 Zceil x - x < / 2 * ulp beta fexp1 x

Valid_exp fexp1
exact Vfexp1.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':midp' fexp1 x < x
Hx:round beta fexp1 Zceil x - x < / 2 * ulp beta fexp1 x

IZR (Zceil (x * bpow (- fexp1 (mag x)))) * bpow (fexp1 (mag x)) - x < / 2 * bpow (fexp1 (mag x))
rewrite ulp_neq_0 in Hx;[exact Hx|now apply Rgt_not_eq].
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':midp' fexp1 x < x
Hx:round beta fexp1 Zceil x - x < / 2 * ulp beta fexp1 x
H:Rabs (IZR (Zceil (x * bpow (- fexp1 (mag x)))) - x * bpow (- fexp1 (mag x))) < / 2

round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':midp' fexp1 x < x
Hx:round beta fexp1 Zceil x - x < / 2 * ulp beta fexp1 x
H:Rabs (IZR (Zceil (x * bpow (- fexp1 (mag x)))) - x * bpow (- fexp1 (mag x))) < / 2

round beta fexp1 (Znearest choice1) (F2R {| Fnum := Znearest choice2 (scaled_mantissa beta fexp2 x); Fexp := cexp beta fexp2 x |}) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':midp' fexp1 x < x
Hx:round beta fexp1 Zceil x - x < / 2 * ulp beta fexp1 x
H:Rabs (IZR (Zceil (x * bpow (- fexp1 (mag x)))) - x * bpow (- fexp1 (mag x))) < / 2

round beta fexp1 (Znearest choice1) (IZR (Znearest choice2 (x * bpow (- fexp2 (mag x)))) * bpow (fexp2 (mag x))) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':midp' fexp1 x < x
Hx:round beta fexp1 Zceil x - x < / 2 * ulp beta fexp1 x
H:Rabs (IZR (Zceil (x * bpow (- fexp1 (mag x)))) - x * bpow (- fexp1 (mag x))) < / 2

round beta fexp1 (Znearest choice1) (IZR (Znearest choice2 (x * bpow (- fexp1 (mag x)))) * bpow (fexp1 (mag x))) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':midp' fexp1 x < x
Hx:round beta fexp1 Zceil x - x < / 2 * ulp beta fexp1 x
H:Rabs (IZR (Zceil (x * bpow (- fexp1 (mag x)))) - x * bpow (- fexp1 (mag x))) < / 2

round beta fexp1 (Znearest choice1) (IZR (Zceil (scaled_mantissa beta fexp1 x)) * bpow (fexp1 (mag x))) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':midp' fexp1 x < x
Hx:round beta fexp1 Zceil x - x < / 2 * ulp beta fexp1 x
H:Rabs (IZR (Zceil (x * bpow (- fexp1 (mag x)))) - x * bpow (- fexp1 (mag x))) < / 2
Rabs (x * bpow (- fexp1 (mag x)) - IZR (Zceil (scaled_mantissa beta fexp1 x))) < / 2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':midp' fexp1 x < x
Hx:round beta fexp1 Zceil x - x < / 2 * ulp beta fexp1 x
H:Rabs (IZR (Zceil (x * bpow (- fexp1 (mag x)))) - x * bpow (- fexp1 (mag x))) < / 2

round beta fexp1 (Znearest choice1) (IZR (Zceil (scaled_mantissa beta fexp1 x)) * bpow (fexp1 (mag x))) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':midp' fexp1 x < x
Hx:round beta fexp1 Zceil x - x < / 2 * ulp beta fexp1 x
H:Rabs (IZR (Zceil (x * bpow (- fexp1 (mag x)))) - x * bpow (- fexp1 (mag x))) < / 2

IZR (Zceil (scaled_mantissa beta fexp1 x)) * bpow (fexp1 (mag x)) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':midp' fexp1 x < x
Hx:round beta fexp1 Zceil x - x < / 2 * ulp beta fexp1 x
H:Rabs (IZR (Zceil (x * bpow (- fexp1 (mag x)))) - x * bpow (- fexp1 (mag x))) < / 2
Valid_rnd (Znearest choice1)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':midp' fexp1 x < x
Hx:round beta fexp1 Zceil x - x < / 2 * ulp beta fexp1 x
H:Rabs (IZR (Zceil (x * bpow (- fexp1 (mag x)))) - x * bpow (- fexp1 (mag x))) < / 2
generic_format beta fexp1 (IZR (Zceil (scaled_mantissa beta fexp1 x)) * bpow (fexp1 (mag x)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':midp' fexp1 x < x
Hx:round beta fexp1 Zceil x - x < / 2 * ulp beta fexp1 x
H:Rabs (IZR (Zceil (x * bpow (- fexp1 (mag x)))) - x * bpow (- fexp1 (mag x))) < / 2

IZR (Zceil (scaled_mantissa beta fexp1 x)) * bpow (fexp1 (mag x)) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':midp' fexp1 x < x
Hx:round beta fexp1 Zceil x - x < / 2 * ulp beta fexp1 x
H:Rabs (IZR (Zceil (x * bpow (- fexp1 (mag x)))) - x * bpow (- fexp1 (mag x))) < / 2

IZR (Zceil (x * bpow (- fexp1 (mag x)))) * bpow (fexp1 (mag x)) = IZR (Znearest choice1 (x * bpow (- fexp1 (mag x)))) * bpow (fexp1 (mag x))
now rewrite (Znearest_imp _ _ (Zceil (x * bpow (- fexp1 (mag x))))); [|rewrite Rabs_minus_sym].
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':midp' fexp1 x < x
Hx:round beta fexp1 Zceil x - x < / 2 * ulp beta fexp1 x
H:Rabs (IZR (Zceil (x * bpow (- fexp1 (mag x)))) - x * bpow (- fexp1 (mag x))) < / 2

Valid_rnd (Znearest choice1)
now apply valid_rnd_N.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':midp' fexp1 x < x
Hx:round beta fexp1 Zceil x - x < / 2 * ulp beta fexp1 x
H:Rabs (IZR (Zceil (x * bpow (- fexp1 (mag x)))) - x * bpow (- fexp1 (mag x))) < / 2

generic_format beta fexp1 (IZR (Zceil (scaled_mantissa beta fexp1 x)) * bpow (fexp1 (mag x)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':midp' fexp1 x < x
Hx:round beta fexp1 Zceil x - x < / 2 * ulp beta fexp1 x
H:Rabs (IZR (Zceil (x * bpow (- fexp1 (mag x)))) - x * bpow (- fexp1 (mag x))) < / 2

generic_format beta fexp1 (IZR (Zceil (scaled_mantissa beta fexp1 x)) * bpow (cexp beta fexp1 x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':midp' fexp1 x < x
Hx:round beta fexp1 Zceil x - x < / 2 * ulp beta fexp1 x
H:Rabs (IZR (Zceil (x * bpow (- fexp1 (mag x)))) - x * bpow (- fexp1 (mag x))) < / 2

generic_format beta fexp1 (round beta fexp1 Zceil x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':midp' fexp1 x < x
Hx:round beta fexp1 Zceil x - x < / 2 * ulp beta fexp1 x
H:Rabs (IZR (Zceil (x * bpow (- fexp1 (mag x)))) - x * bpow (- fexp1 (mag x))) < / 2

Valid_exp fexp1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':midp' fexp1 x < x
Hx:round beta fexp1 Zceil x - x < / 2 * ulp beta fexp1 x
H:Rabs (IZR (Zceil (x * bpow (- fexp1 (mag x)))) - x * bpow (- fexp1 (mag x))) < / 2
Valid_rnd Zceil
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':midp' fexp1 x < x
Hx:round beta fexp1 Zceil x - x < / 2 * ulp beta fexp1 x
H:Rabs (IZR (Zceil (x * bpow (- fexp1 (mag x)))) - x * bpow (- fexp1 (mag x))) < / 2

Valid_rnd Zceil
now apply valid_rnd_UP.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:fexp2 (mag x) = fexp1 (mag x)
Hx':midp' fexp1 x < x
Hx:round beta fexp1 Zceil x - x < / 2 * ulp beta fexp1 x
H:Rabs (IZR (Zceil (x * bpow (- fexp1 (mag x)))) - x * bpow (- fexp1 (mag x))) < / 2

Rabs (x * bpow (- fexp1 (mag x)) - IZR (Zceil (scaled_mantissa beta fexp1 x))) < / 2
now rewrite Rabs_minus_sym. Qed.
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall (choice1 choice2 : Z -> bool) (x : R), 0 < x -> (fexp2 (mag x) <= fexp1 (mag x))%Z -> (fexp1 (mag x) <= mag x)%Z -> midp' fexp1 x < x -> ((fexp2 (mag x) <= fexp1 (mag x) - 1)%Z -> midp' fexp1 x + / 2 * ulp beta fexp2 x < x) -> round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall (choice1 choice2 : Z -> bool) (x : R), 0 < x -> (fexp2 (mag x) <= fexp1 (mag x))%Z -> (fexp1 (mag x) <= mag x)%Z -> midp' fexp1 x < x -> ((fexp2 (mag x) <= fexp1 (mag x) - 1)%Z -> midp' fexp1 x + / 2 * ulp beta fexp2 x < x) -> round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx:midp' fexp1 x < x
Hx':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z -> midp' fexp1 x + / 2 * ulp beta fexp2 x < x

round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx:midp' fexp1 x < x
Hx':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z -> midp' fexp1 x + / 2 * ulp beta fexp2 x < x
Hf2':(fexp1 (mag x) <= fexp2 (mag x))%Z

round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx:midp' fexp1 x < x
Hx':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z -> midp' fexp1 x + / 2 * ulp beta fexp2 x < x
Hf2':(fexp2 (mag x) < fexp1 (mag x))%Z
round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx:midp' fexp1 x < x
Hx':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z -> midp' fexp1 x + / 2 * ulp beta fexp2 x < x
Hf2':(fexp1 (mag x) <= fexp2 (mag x))%Z

round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx:midp' fexp1 x < x
Hx':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z -> midp' fexp1 x + / 2 * ulp beta fexp2 x < x
Hf2':(fexp1 (mag x) <= fexp2 (mag x))%Z
Hf2'':fexp2 (mag x) = fexp1 (mag x)

round_round_eq fexp1 fexp2 choice1 choice2 x
now apply round_round_gt_mid_same_place.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx:midp' fexp1 x < x
Hx':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z -> midp' fexp1 x + / 2 * ulp beta fexp2 x < x
Hf2':(fexp2 (mag x) < fexp1 (mag x))%Z

round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx:midp' fexp1 x < x
Hx':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z -> midp' fexp1 x + / 2 * ulp beta fexp2 x < x
Hf2':(fexp2 (mag x) < fexp1 (mag x))%Z
Hf2'':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z

round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
Hx:midp' fexp1 x < x
Hx':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z -> midp' fexp1 x + / 2 * ulp beta fexp2 x < x
Hf2':(fexp2 (mag x) < fexp1 (mag x))%Z
Hf2'':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hx'':midp' fexp1 x + / 2 * ulp beta fexp2 x < x

round_round_eq fexp1 fexp2 choice1 choice2 x
now apply round_round_gt_mid_further_place. Qed. Section Double_round_mult.
beta:radix

forall x y : R, x <> 0 -> y <> 0 -> mag (x * y) = (mag x + mag y - 1)%Z \/ mag (x * y) = (mag x + mag y)%Z
beta:radix

forall x y : R, x <> 0 -> y <> 0 -> mag (x * y) = (mag x + mag y - 1)%Z \/ mag (x * y) = (mag x + mag y)%Z
beta:radix
x, y:R
Zx:x <> 0
Zy:y <> 0

mag (x * y) = (mag x + mag y - 1)%Z \/ mag (x * y) = (mag x + mag y)%Z
beta:radix
x, y:R
Zx:x <> 0
Zy:y <> 0
H:(mag x + mag y - 1 <= mag (x * y))%Z
H0:(mag (x * y) <= mag x + mag y)%Z

mag (x * y) = (mag x + mag y - 1)%Z \/ mag (x * y) = (mag x + mag y)%Z
omega. Qed. Definition round_round_mult_hyp fexp1 fexp2 := (forall ex ey, (fexp2 (ex + ey) <= fexp1 ex + fexp1 ey)%Z) /\ (forall ex ey, (fexp2 (ex + ey - 1) <= fexp1 ex + fexp1 ey)%Z).
beta:radix

forall fexp1 fexp2 : Z -> Z, round_round_mult_hyp fexp1 fexp2 -> forall x y : R, generic_format beta fexp1 x -> generic_format beta fexp1 y -> generic_format beta fexp2 (x * y)
beta:radix

forall fexp1 fexp2 : Z -> Z, round_round_mult_hyp fexp1 fexp2 -> forall x y : R, generic_format beta fexp1 x -> generic_format beta fexp1 y -> generic_format beta fexp2 (x * y)
beta:radix
fexp1, fexp2:Z -> Z
Hfexp:round_round_mult_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y

generic_format beta fexp2 (x * y)
beta:radix
fexp1, fexp2:Z -> Z
Hfexp:round_round_mult_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0

generic_format beta fexp2 (x * y)
beta:radix
fexp1, fexp2:Z -> Z
Hfexp:round_round_mult_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x <> 0
generic_format beta fexp2 (x * y)
beta:radix
fexp1, fexp2:Z -> Z
Hfexp:round_round_mult_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0

generic_format beta fexp2 (x * y)
beta:radix
fexp1, fexp2:Z -> Z
Hfexp:round_round_mult_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0

generic_format beta fexp2 (0 * y)
beta:radix
fexp1, fexp2:Z -> Z
Hfexp:round_round_mult_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0

generic_format beta fexp2 0
now apply generic_format_0.
beta:radix
fexp1, fexp2:Z -> Z
Hfexp:round_round_mult_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x <> 0

generic_format beta fexp2 (x * y)
beta:radix
fexp1, fexp2:Z -> Z
Hfexp:round_round_mult_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x <> 0
Zy:y = 0

generic_format beta fexp2 (x * y)
beta:radix
fexp1, fexp2:Z -> Z
Hfexp:round_round_mult_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x <> 0
Zy:y <> 0
generic_format beta fexp2 (x * y)
beta:radix
fexp1, fexp2:Z -> Z
Hfexp:round_round_mult_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x <> 0
Zy:y = 0

generic_format beta fexp2 (x * y)
beta:radix
fexp1, fexp2:Z -> Z
Hfexp:round_round_mult_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x <> 0
Zy:y = 0

generic_format beta fexp2 (x * 0)
beta:radix
fexp1, fexp2:Z -> Z
Hfexp:round_round_mult_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x <> 0
Zy:y = 0

generic_format beta fexp2 0
now apply generic_format_0.
beta:radix
fexp1, fexp2:Z -> Z
Hfexp:round_round_mult_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x <> 0
Zy:y <> 0

generic_format beta fexp2 (x * y)
beta:radix
fexp1, fexp2:Z -> Z
Hfexp:round_round_mult_hyp fexp1 fexp2
x, y:R
Zx:x <> 0
Zy:y <> 0

generic_format beta fexp1 x -> generic_format beta fexp1 y -> generic_format beta fexp2 (x * y)
beta:radix
fexp1, fexp2:Z -> Z
Hfexp:round_round_mult_hyp fexp1 fexp2
x, y:R
Zx:x <> 0
Zy:y <> 0

x = F2R {| Fnum := Ztrunc (scaled_mantissa beta fexp1 x); Fexp := cexp beta fexp1 x |} -> y = F2R {| Fnum := Ztrunc (scaled_mantissa beta fexp1 y); Fexp := cexp beta fexp1 y |} -> x * y = F2R {| Fnum := Ztrunc (scaled_mantissa beta fexp2 (x * y)); Fexp := cexp beta fexp2 (x * y) |}
beta:radix
fexp1, fexp2:Z -> Z
Hfexp:round_round_mult_hyp fexp1 fexp2
x, y:R
Zx:x <> 0
Zy:y <> 0

x = F2R {| Fnum := Ztrunc (scaled_mantissa beta fexp1 x); Fexp := fexp1 (mag x) |} -> y = F2R {| Fnum := Ztrunc (scaled_mantissa beta fexp1 y); Fexp := fexp1 (mag y) |} -> x * y = F2R {| Fnum := Ztrunc (scaled_mantissa beta fexp2 (x * y)); Fexp := fexp2 (mag (x * y)) |}
beta:radix
fexp1, fexp2:Z -> Z
Hfexp:round_round_mult_hyp fexp1 fexp2
x, y:R
Zx:x <> 0
Zy:y <> 0
mx:=Ztrunc (scaled_mantissa beta fexp1 x):Z

x = F2R {| Fnum := mx; Fexp := fexp1 (mag x) |} -> y = F2R {| Fnum := Ztrunc (scaled_mantissa beta fexp1 y); Fexp := fexp1 (mag y) |} -> x * y = F2R {| Fnum := Ztrunc (scaled_mantissa beta fexp2 (x * y)); Fexp := fexp2 (mag (x * y)) |}
beta:radix
fexp1, fexp2:Z -> Z
Hfexp:round_round_mult_hyp fexp1 fexp2
x, y:R
Zx:x <> 0
Zy:y <> 0
mx:=Ztrunc (scaled_mantissa beta fexp1 x):Z
my:=Ztrunc (scaled_mantissa beta fexp1 y):Z

x = F2R {| Fnum := mx; Fexp := fexp1 (mag x) |} -> y = F2R {| Fnum := my; Fexp := fexp1 (mag y) |} -> x * y = F2R {| Fnum := Ztrunc (scaled_mantissa beta fexp2 (x * y)); Fexp := fexp2 (mag (x * y)) |}
beta:radix
fexp1, fexp2:Z -> Z
Hfexp:round_round_mult_hyp fexp1 fexp2
x, y:R
Zx:x <> 0
Zy:y <> 0
mx:=Ztrunc (scaled_mantissa beta fexp1 x):Z
my:=Ztrunc (scaled_mantissa beta fexp1 y):Z

x = IZR mx * bpow (fexp1 (mag x)) -> y = IZR my * bpow (fexp1 (mag y)) -> x * y = IZR (Ztrunc (scaled_mantissa beta fexp2 (x * y))) * bpow (fexp2 (mag (x * y)))
beta:radix
fexp1, fexp2:Z -> Z
Hfexp:round_round_mult_hyp fexp1 fexp2
x, y:R
Zx:x <> 0
Zy:y <> 0
mx:=Ztrunc (scaled_mantissa beta fexp1 x):Z
my:=Ztrunc (scaled_mantissa beta fexp1 y):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))

x * y = IZR (Ztrunc (scaled_mantissa beta fexp2 (x * y))) * bpow (fexp2 (mag (x * y)))
beta:radix
fexp1, fexp2:Z -> Z
Hfexp:round_round_mult_hyp fexp1 fexp2
x, y:R
Zx:x <> 0
Zy:y <> 0
mx:=Ztrunc (scaled_mantissa beta fexp1 x):Z
my:=Ztrunc (scaled_mantissa beta fexp1 y):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
fxy:={| Fnum := mx * my; Fexp := fexp1 (mag x) + fexp1 (mag y) |}:float beta

x * y = IZR (Ztrunc (scaled_mantissa beta fexp2 (x * y))) * bpow (fexp2 (mag (x * y)))
beta:radix
fexp1, fexp2:Z -> Z
Hfexp:round_round_mult_hyp fexp1 fexp2
x, y:R
Zx:x <> 0
Zy:y <> 0
mx:=Ztrunc (scaled_mantissa beta fexp1 x):Z
my:=Ztrunc (scaled_mantissa beta fexp1 y):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
fxy:={| Fnum := mx * my; Fexp := fexp1 (mag x) + fexp1 (mag y) |}:float beta

x * y = F2R fxy
beta:radix
fexp1, fexp2:Z -> Z
Hfexp:round_round_mult_hyp fexp1 fexp2
x, y:R
Zx:x <> 0
Zy:y <> 0
mx:=Ztrunc (scaled_mantissa beta fexp1 x):Z
my:=Ztrunc (scaled_mantissa beta fexp1 y):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
fxy:={| Fnum := mx * my; Fexp := fexp1 (mag x) + fexp1 (mag y) |}:float beta
Hxy:x * y = F2R fxy
x * y = IZR (Ztrunc (scaled_mantissa beta fexp2 (x * y))) * bpow (fexp2 (mag (x * y)))
beta:radix
fexp1, fexp2:Z -> Z
Hfexp:round_round_mult_hyp fexp1 fexp2
x, y:R
Zx:x <> 0
Zy:y <> 0
mx:=Ztrunc (scaled_mantissa beta fexp1 x):Z
my:=Ztrunc (scaled_mantissa beta fexp1 y):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
fxy:={| Fnum := mx * my; Fexp := fexp1 (mag x) + fexp1 (mag y) |}:float beta

x * y = F2R fxy
beta:radix
fexp1, fexp2:Z -> Z
Hfexp:round_round_mult_hyp fexp1 fexp2
x, y:R
Zx:x <> 0
Zy:y <> 0
mx:=Ztrunc (scaled_mantissa beta fexp1 x):Z
my:=Ztrunc (scaled_mantissa beta fexp1 y):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
fxy:={| Fnum := mx * my; Fexp := fexp1 (mag x) + fexp1 (mag y) |}:float beta

x * y = IZR (mx * my) * bpow (fexp1 (mag x) + fexp1 (mag y))
beta:radix
fexp1, fexp2:Z -> Z
Hfexp:round_round_mult_hyp fexp1 fexp2
x, y:R
Zx:x <> 0
Zy:y <> 0
mx:=Ztrunc (scaled_mantissa beta fexp1 x):Z
my:=Ztrunc (scaled_mantissa beta fexp1 y):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
fxy:={| Fnum := mx * my; Fexp := fexp1 (mag x) + fexp1 (mag y) |}:float beta

x * y = IZR (mx * my) * (bpow (fexp1 (mag x)) * bpow (fexp1 (mag y)))
beta:radix
fexp1, fexp2:Z -> Z
Hfexp:round_round_mult_hyp fexp1 fexp2
x, y:R
Zx:x <> 0
Zy:y <> 0
mx:=Ztrunc (scaled_mantissa beta fexp1 x):Z
my:=Ztrunc (scaled_mantissa beta fexp1 y):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
fxy:={| Fnum := mx * my; Fexp := fexp1 (mag x) + fexp1 (mag y) |}:float beta

x * y = IZR mx * IZR my * (bpow (fexp1 (mag x)) * bpow (fexp1 (mag y)))
beta:radix
fexp1, fexp2:Z -> Z
Hfexp:round_round_mult_hyp fexp1 fexp2
x, y:R
Zx:x <> 0
Zy:y <> 0
mx:=Ztrunc (scaled_mantissa beta fexp1 x):Z
my:=Ztrunc (scaled_mantissa beta fexp1 y):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
fxy:={| Fnum := mx * my; Fexp := fexp1 (mag x) + fexp1 (mag y) |}:float beta

IZR mx * bpow (fexp1 (mag x)) * (IZR my * bpow (fexp1 (mag y))) = IZR mx * IZR my * (bpow (fexp1 (mag x)) * bpow (fexp1 (mag y)))
ring.
beta:radix
fexp1, fexp2:Z -> Z
Hfexp:round_round_mult_hyp fexp1 fexp2
x, y:R
Zx:x <> 0
Zy:y <> 0
mx:=Ztrunc (scaled_mantissa beta fexp1 x):Z
my:=Ztrunc (scaled_mantissa beta fexp1 y):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
fxy:={| Fnum := mx * my; Fexp := fexp1 (mag x) + fexp1 (mag y) |}:float beta
Hxy:x * y = F2R fxy

x * y = IZR (Ztrunc (scaled_mantissa beta fexp2 (x * y))) * bpow (fexp2 (mag (x * y)))
beta:radix
fexp1, fexp2:Z -> Z
Hfexp:round_round_mult_hyp fexp1 fexp2
x, y:R
Zx:x <> 0
Zy:y <> 0
mx:=Ztrunc (scaled_mantissa beta fexp1 x):Z
my:=Ztrunc (scaled_mantissa beta fexp1 y):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
fxy:={| Fnum := mx * my; Fexp := fexp1 (mag x) + fexp1 (mag y) |}:float beta
Hxy:x * y = F2R fxy

x * y <> 0 -> (cexp beta fexp2 (x * y) <= Fexp fxy)%Z
beta:radix
fexp1, fexp2:Z -> Z
Hfexp:round_round_mult_hyp fexp1 fexp2
x, y:R
Zx:x <> 0
Zy:y <> 0
mx:=Ztrunc (scaled_mantissa beta fexp1 x):Z
my:=Ztrunc (scaled_mantissa beta fexp1 y):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
fxy:={| Fnum := mx * my; Fexp := fexp1 (mag x) + fexp1 (mag y) |}:float beta
Hxy:x * y = F2R fxy

(cexp beta fexp2 (x * y) <= Fexp fxy)%Z
beta:radix
fexp1, fexp2:Z -> Z
Hfexp:round_round_mult_hyp fexp1 fexp2
x, y:R
Zx:x <> 0
Zy:y <> 0
mx:=Ztrunc (scaled_mantissa beta fexp1 x):Z
my:=Ztrunc (scaled_mantissa beta fexp1 y):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
fxy:={| Fnum := mx * my; Fexp := fexp1 (mag x) + fexp1 (mag y) |}:float beta
Hxy:x * y = F2R fxy

(fexp2 (mag (x * y)) <= fexp1 (mag x) + fexp1 (mag y))%Z
beta:radix
fexp1, fexp2:Z -> Z
Hfexp1:forall ex ey : Z, (fexp2 (ex + ey) <= fexp1 ex + fexp1 ey)%Z
Hfexp2:forall ex ey : Z, (fexp2 (ex + ey - 1) <= fexp1 ex + fexp1 ey)%Z
x, y:R
Zx:x <> 0
Zy:y <> 0
mx:=Ztrunc (scaled_mantissa beta fexp1 x):Z
my:=Ztrunc (scaled_mantissa beta fexp1 y):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
fxy:={| Fnum := mx * my; Fexp := fexp1 (mag x) + fexp1 (mag y) |}:float beta
Hxy:x * y = F2R fxy

(fexp2 (mag (x * y)) <= fexp1 (mag x) + fexp1 (mag y))%Z
now destruct (mag_mult_disj x y Zx Zy) as [Lxy|Lxy]; rewrite Lxy. Qed. Variable rnd : R -> Z. Context { valid_rnd : Valid_rnd rnd }.
beta:radix
rnd:R -> Z
valid_rnd:Valid_rnd rnd

forall fexp1 fexp2 : Z -> Z, round_round_mult_hyp fexp1 fexp2 -> forall x y : R, generic_format beta fexp1 x -> generic_format beta fexp1 y -> round beta fexp1 rnd (round beta fexp2 rnd (x * y)) = round beta fexp1 rnd (x * y)
beta:radix
rnd:R -> Z
valid_rnd:Valid_rnd rnd

forall fexp1 fexp2 : Z -> Z, round_round_mult_hyp fexp1 fexp2 -> forall x y : R, generic_format beta fexp1 x -> generic_format beta fexp1 y -> round beta fexp1 rnd (round beta fexp2 rnd (x * y)) = round beta fexp1 rnd (x * y)
beta:radix
rnd:R -> Z
valid_rnd:Valid_rnd rnd
fexp1, fexp2:Z -> Z
Hfexp:round_round_mult_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y

round beta fexp1 rnd (round beta fexp2 rnd (x * y)) = round beta fexp1 rnd (x * y)
beta:radix
rnd:R -> Z
valid_rnd:Valid_rnd rnd
fexp1, fexp2:Z -> Z
Hfexp:round_round_mult_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y

round beta fexp2 rnd (x * y) = x * y
beta:radix
rnd:R -> Z
valid_rnd:Valid_rnd rnd
fexp1, fexp2:Z -> Z
Hfexp:round_round_mult_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hxy:round beta fexp2 rnd (x * y) = x * y
round beta fexp1 rnd (round beta fexp2 rnd (x * y)) = round beta fexp1 rnd (x * y)
beta:radix
rnd:R -> Z
valid_rnd:Valid_rnd rnd
fexp1, fexp2:Z -> Z
Hfexp:round_round_mult_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y

round beta fexp2 rnd (x * y) = x * y
beta:radix
rnd:R -> Z
valid_rnd:Valid_rnd rnd
fexp1, fexp2:Z -> Z
Hfexp:round_round_mult_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y

generic_format beta fexp2 (x * y)
now apply (round_round_mult_aux fexp1 fexp2).
beta:radix
rnd:R -> Z
valid_rnd:Valid_rnd rnd
fexp1, fexp2:Z -> Z
Hfexp:round_round_mult_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hxy:round beta fexp2 rnd (x * y) = x * y

round beta fexp1 rnd (round beta fexp2 rnd (x * y)) = round beta fexp1 rnd (x * y)
now rewrite Hxy at 1. Qed. Section Double_round_mult_FLX. Variable prec : Z. Variable prec' : Z. Context { prec_gt_0_ : Prec_gt_0 prec }. Context { prec_gt_0_' : Prec_gt_0 prec' }.
beta:radix
rnd:R -> Z
valid_rnd:Valid_rnd rnd
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(2 * prec <= prec')%Z -> forall x y : R, FLX_format beta prec x -> FLX_format beta prec y -> round beta (FLX_exp prec) rnd (round beta (FLX_exp prec') rnd (x * y)) = round beta (FLX_exp prec) rnd (x * y)
beta:radix
rnd:R -> Z
valid_rnd:Valid_rnd rnd
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(2 * prec <= prec')%Z -> forall x y : R, FLX_format beta prec x -> FLX_format beta prec y -> round beta (FLX_exp prec) rnd (round beta (FLX_exp prec') rnd (x * y)) = round beta (FLX_exp prec) rnd (x * y)
beta:radix
rnd:R -> Z
valid_rnd:Valid_rnd rnd
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y

round beta (FLX_exp prec) rnd (round beta (FLX_exp prec') rnd (x * y)) = round beta (FLX_exp prec) rnd (x * y)
beta:radix
rnd:R -> Z
valid_rnd:Valid_rnd rnd
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y

round_round_mult_hyp (FLX_exp prec) (FLX_exp prec')
unfold round_round_mult_hyp; split; intros ex ey; unfold FLX_exp; omega. Qed. End Double_round_mult_FLX. Section Double_round_mult_FLT. Variable emin prec : Z. Variable emin' prec' : Z. Context { prec_gt_0_ : Prec_gt_0 prec }. Context { prec_gt_0_' : Prec_gt_0 prec' }.
beta:radix
rnd:R -> Z
valid_rnd:Valid_rnd rnd
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(emin' <= 2 * emin)%Z -> (2 * prec <= prec')%Z -> forall x y : R, FLT_format beta emin prec x -> FLT_format beta emin prec y -> round beta (FLT_exp emin prec) rnd (round beta (FLT_exp emin' prec') rnd (x * y)) = round beta (FLT_exp emin prec) rnd (x * y)
beta:radix
rnd:R -> Z
valid_rnd:Valid_rnd rnd
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(emin' <= 2 * emin)%Z -> (2 * prec <= prec')%Z -> forall x y : R, FLT_format beta emin prec x -> FLT_format beta emin prec y -> round beta (FLT_exp emin prec) rnd (round beta (FLT_exp emin' prec') rnd (x * y)) = round beta (FLT_exp emin prec) rnd (x * y)
beta:radix
rnd:R -> Z
valid_rnd:Valid_rnd rnd
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= 2 * emin)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y

round beta (FLT_exp emin prec) rnd (round beta (FLT_exp emin' prec') rnd (x * y)) = round beta (FLT_exp emin prec) rnd (x * y)
beta:radix
rnd:R -> Z
valid_rnd:Valid_rnd rnd
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= 2 * emin)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y

round_round_mult_hyp (FLT_exp emin prec) (FLT_exp emin' prec')
unfold round_round_mult_hyp; split; intros ex ey; unfold FLT_exp; generalize (Zmax_spec (ex + ey - prec') emin'); generalize (Zmax_spec (ex + ey - 1 - prec') emin'); generalize (Zmax_spec (ex - prec) emin); generalize (Zmax_spec (ey - prec) emin); omega. Qed. End Double_round_mult_FLT. Section Double_round_mult_FTZ. Variable emin prec : Z. Variable emin' prec' : Z. Context { prec_gt_0_ : Prec_gt_0 prec }. Context { prec_gt_0_' : Prec_gt_0 prec' }.
beta:radix
rnd:R -> Z
valid_rnd:Valid_rnd rnd
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(emin' + prec' <= 2 * emin + prec)%Z -> (2 * prec <= prec')%Z -> forall x y : R, FTZ_format beta emin prec x -> FTZ_format beta emin prec y -> round beta (FTZ_exp emin prec) rnd (round beta (FTZ_exp emin' prec') rnd (x * y)) = round beta (FTZ_exp emin prec) rnd (x * y)
beta:radix
rnd:R -> Z
valid_rnd:Valid_rnd rnd
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(emin' + prec' <= 2 * emin + prec)%Z -> (2 * prec <= prec')%Z -> forall x y : R, FTZ_format beta emin prec x -> FTZ_format beta emin prec y -> round beta (FTZ_exp emin prec) rnd (round beta (FTZ_exp emin' prec') rnd (x * y)) = round beta (FTZ_exp emin prec) rnd (x * y)
beta:radix
rnd:R -> Z
valid_rnd:Valid_rnd rnd
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' + prec' <= 2 * emin + prec)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y

round beta (FTZ_exp emin prec) rnd (round beta (FTZ_exp emin' prec') rnd (x * y)) = round beta (FTZ_exp emin prec) rnd (x * y)
beta:radix
rnd:R -> Z
valid_rnd:Valid_rnd rnd
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' + prec' <= 2 * emin + prec)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y

round_round_mult_hyp (FTZ_exp emin prec) (FTZ_exp emin' prec')
unfold round_round_mult_hyp; split; intros ex ey; unfold FTZ_exp; unfold Prec_gt_0 in *; destruct (Z.ltb_spec (ex + ey - prec') emin'); destruct (Z.ltb_spec (ex - prec) emin); destruct (Z.ltb_spec (ey - prec) emin); destruct (Z.ltb_spec (ex + ey - 1 - prec') emin'); omega. Qed. End Double_round_mult_FTZ. End Double_round_mult. Section Double_round_plus.
beta:radix

forall x y : R, 0 < y -> y <= x -> mag (x + y) = mag x \/ mag (x + y) = (mag x + 1)%Z
beta:radix

forall x y : R, 0 < y -> y <= x -> mag (x + y) = mag x \/ mag (x + y) = (mag x + 1)%Z
beta:radix
x, y:R
Py:0 < y
Hxy:y <= x

mag (x + y) = mag x \/ mag (x + y) = (mag x + 1)%Z
beta:radix
x, y:R
Py:0 < y
Hxy:y <= x
H:(mag x <= mag (x + y))%Z
H0:(mag (x + y) <= mag x + 1)%Z

mag (x + y) = mag x \/ mag (x + y) = (mag x + 1)%Z
omega. Qed.
beta:radix

forall (fexp : Z -> Z) (x y : R), 0 < x -> 0 <= y -> generic_format beta fexp x -> (mag y <= fexp (mag x))%Z -> mag (x + y) = mag x
beta:radix

forall (fexp : Z -> Z) (x y : R), 0 < x -> 0 <= y -> generic_format beta fexp x -> (mag y <= fexp (mag x))%Z -> mag (x + y) = mag x
beta:radix
fexp:Z -> Z
x, y:R
Px:0 < x
Nny:0 <= y
Fx:generic_format beta fexp x
Hsep:(mag y <= fexp (mag x))%Z

mag (x + y) = mag x
beta:radix
fexp:Z -> Z
x, y:R
Px:0 < x
Nny:0 <= y
Fx:generic_format beta fexp x
Hsep:(mag y <= fexp (mag x))%Z

0 <= y < ulp beta fexp x
beta:radix
fexp:Z -> Z
x, y:R
Px:0 < x
Nny:0 <= y
Fx:generic_format beta fexp x
Hsep:(mag y <= fexp (mag x))%Z

y < ulp beta fexp x
beta:radix
fexp:Z -> Z
x, y:R
Px:0 < x
Nny:0 <= y
Fx:generic_format beta fexp x
Hsep:(mag y <= fexp (mag x))%Z

Rabs y < ulp beta fexp x
beta:radix
fexp:Z -> Z
x, y:R
Px:0 < x
Nny:0 <= y
Fx:generic_format beta fexp x
Hsep:(mag y <= fexp (mag x))%Z

bpow (mag y) <= ulp beta fexp x
beta:radix
fexp:Z -> Z
x, y:R
Px:0 < x
Nny:0 <= y
Fx:generic_format beta fexp x
Hsep:(mag y <= fexp (mag x))%Z

bpow (mag y) <= bpow (cexp beta fexp x)
now apply bpow_le. Qed.
beta:radix

forall x y : R, 0 < x -> 0 < y -> (mag y <= mag x - 2)%Z -> mag (x - y) = mag x \/ mag (x - y) = (mag x - 1)%Z
beta:radix

forall x y : R, 0 < x -> 0 < y -> (mag y <= mag x - 2)%Z -> mag (x - y) = mag x \/ mag (x - y) = (mag x - 1)%Z
beta:radix
x, y:R
Px:0 < x
Py:0 < y
Hln:(mag y <= mag x - 2)%Z

mag (x - y) = mag x \/ mag (x - y) = (mag x - 1)%Z
beta:radix
x, y:R
Px:0 < x
Py:0 < y
Hln:(mag y <= mag x - 2)%Z
Hxy:y < x

mag (x - y) = mag x \/ mag (x - y) = (mag x - 1)%Z
beta:radix
x, y:R
Px:0 < x
Py:0 < y
Hln:(mag y <= mag x - 2)%Z
Hxy:y < x
Hln2:(mag (x - y) <= mag x)%Z

mag (x - y) = mag x \/ mag (x - y) = (mag x - 1)%Z
beta:radix
x, y:R
Px:0 < x
Py:0 < y
Hln:(mag y <= mag x - 2)%Z
Hxy:y < x
Hln2:(mag (x - y) <= mag x)%Z
Hln3:(mag x - 1 <= mag (x - y))%Z

mag (x - y) = mag x \/ mag (x - y) = (mag x - 1)%Z
omega. Qed.
beta:radix

forall fexp : Z -> Z, Valid_exp fexp -> forall x y : R, 0 < x -> 0 < y -> y < x -> bpow (mag x - 1) < x -> generic_format beta fexp x -> (mag y <= fexp (mag x))%Z -> mag (x - y) = mag x
beta:radix

forall fexp : Z -> Z, Valid_exp fexp -> forall x y : R, 0 < x -> 0 < y -> y < x -> bpow (mag x - 1) < x -> generic_format beta fexp x -> (mag y <= fexp (mag x))%Z -> mag (x - y) = mag x
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Px:0 < x
Py:0 < y
Yltx:y < x
Xgtpow:bpow (mag x - 1) < x
Fx:generic_format beta fexp x
Ly:(mag y <= fexp (mag x))%Z

mag (x - y) = mag x
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Px:0 < x
Py:0 < y
Yltx:y < x
Xgtpow:bpow (mag x - 1) < x
Fx:generic_format beta fexp x
Ly:(mag y <= fexp (mag x))%Z

bpow (mag x - 1) <= Rabs (x - y) < bpow (mag x)
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Px:0 < x
Py:0 < y
Yltx:y < x
Xgtpow:bpow (mag x - 1) < x
Fx:generic_format beta fexp x
Ly:(mag y <= fexp (mag x))%Z

bpow (mag x - 1) <= Rabs (x - y)
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Px:0 < x
Py:0 < y
Yltx:y < x
Xgtpow:bpow (mag x - 1) < x
Fx:generic_format beta fexp x
Ly:(mag y <= fexp (mag x))%Z
Rabs (x - y) < bpow (mag x)
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Px:0 < x
Py:0 < y
Yltx:y < x
Xgtpow:bpow (mag x - 1) < x
Fx:generic_format beta fexp x
Ly:(mag y <= fexp (mag x))%Z

bpow (mag x - 1) <= Rabs (x - y)
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Px:0 < x
Py:0 < y
Yltx:y < x
Xgtpow:bpow (mag x - 1) < x
Fx:generic_format beta fexp x
Ly:(mag y <= fexp (mag x))%Z

bpow (mag x - 1) <= x - y
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Px:0 < x
Py:0 < y
Yltx:y < x
Xgtpow:bpow (mag x - 1) < x
Fx:generic_format beta fexp x
Ly:(mag y <= fexp (mag x))%Z

y < ulp beta fexp (bpow (mag x - 1))
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Px:0 < x
Py:0 < y
Yltx:y < x
Xgtpow:bpow (mag x - 1) < x
Fx:generic_format beta fexp x
Ly:(mag y <= fexp (mag x))%Z
Hy:y < ulp beta fexp (bpow (mag x - 1))
bpow (mag x - 1) <= x - y
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Px:0 < x
Py:0 < y
Yltx:y < x
Xgtpow:bpow (mag x - 1) < x
Fx:generic_format beta fexp x
Ly:(mag y <= fexp (mag x))%Z

y < ulp beta fexp (bpow (mag x - 1))
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Px:0 < x
Py:0 < y
Yltx:y < x
Xgtpow:bpow (mag x - 1) < x
Fx:generic_format beta fexp x
Ly:(mag y <= fexp (mag x))%Z

y < bpow (fexp (mag x - 1 + 1)%Z)
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Px:0 < x
Py:0 < y
Yltx:y < x
Xgtpow:bpow (mag x - 1) < x
Fx:generic_format beta fexp x
Ly:(mag y <= fexp (mag x))%Z

y < bpow (fexp (mag x))
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Px:0 < x
Py:0 < y
Yltx:y < x
Xgtpow:bpow (mag x - 1) < x
Fx:generic_format beta fexp x
Ly:(mag y <= fexp (mag x))%Z

Rabs y < bpow (fexp (mag x))
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Px:0 < x
Py:0 < y
Yltx:y < x
Xgtpow:bpow (mag x - 1) < x
Fx:generic_format beta fexp x
Ly:(mag y <= fexp (mag x))%Z

Rabs y < bpow (mag y)
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Px:0 < x
Py:0 < y
Yltx:y < x
Xgtpow:bpow (mag x - 1) < x
Fx:generic_format beta fexp x
Ly:(mag y <= fexp (mag x))%Z
bpow (mag y) <= bpow (fexp (mag x))
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Px:0 < x
Py:0 < y
Yltx:y < x
Xgtpow:bpow (mag x - 1) < x
Fx:generic_format beta fexp x
Ly:(mag y <= fexp (mag x))%Z

Rabs y < bpow (mag y)
apply bpow_mag_gt.
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Px:0 < x
Py:0 < y
Yltx:y < x
Xgtpow:bpow (mag x - 1) < x
Fx:generic_format beta fexp x
Ly:(mag y <= fexp (mag x))%Z

bpow (mag y) <= bpow (fexp (mag x))
now apply bpow_le.
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Px:0 < x
Py:0 < y
Yltx:y < x
Xgtpow:bpow (mag x - 1) < x
Fx:generic_format beta fexp x
Ly:(mag y <= fexp (mag x))%Z
Hy:y < ulp beta fexp (bpow (mag x - 1))

bpow (mag x - 1) <= x - y
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Px:0 < x
Py:0 < y
Yltx:y < x
Xgtpow:bpow (mag x - 1) < x
Fx:generic_format beta fexp x
Ly:(mag y <= fexp (mag x))%Z
Hy:y < ulp beta fexp (bpow (mag x - 1))

bpow (mag x - 1) + y <= x
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Px:0 < x
Py:0 < y
Yltx:y < x
Xgtpow:bpow (mag x - 1) < x
Fx:generic_format beta fexp x
Ly:(mag y <= fexp (mag x))%Z
Hy:y < ulp beta fexp (bpow (mag x - 1))

bpow (mag x - 1) + y <= bpow (mag x - 1) + ulp beta fexp (bpow (mag x - 1))
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Px:0 < x
Py:0 < y
Yltx:y < x
Xgtpow:bpow (mag x - 1) < x
Fx:generic_format beta fexp x
Ly:(mag y <= fexp (mag x))%Z
Hy:y < ulp beta fexp (bpow (mag x - 1))
bpow (mag x - 1) + ulp beta fexp (bpow (mag x - 1)) <= x
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Px:0 < x
Py:0 < y
Yltx:y < x
Xgtpow:bpow (mag x - 1) < x
Fx:generic_format beta fexp x
Ly:(mag y <= fexp (mag x))%Z
Hy:y < ulp beta fexp (bpow (mag x - 1))

bpow (mag x - 1) + y <= bpow (mag x - 1) + ulp beta fexp (bpow (mag x - 1))
now apply Rplus_le_compat_l; apply Rlt_le.
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Px:0 < x
Py:0 < y
Yltx:y < x
Xgtpow:bpow (mag x - 1) < x
Fx:generic_format beta fexp x
Ly:(mag y <= fexp (mag x))%Z
Hy:y < ulp beta fexp (bpow (mag x - 1))

bpow (mag x - 1) + ulp beta fexp (bpow (mag x - 1)) <= x
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Px:0 < x
Py:0 < y
Yltx:y < x
Xgtpow:bpow (mag x - 1) < x
Fx:generic_format beta fexp x
Ly:(mag y <= fexp (mag x))%Z
Hy:y < ulp beta fexp (bpow (mag x - 1))

succ beta fexp (bpow (mag x - 1)) <= x
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Px:0 < x
Py:0 < y
Yltx:y < x
Xgtpow:bpow (mag x - 1) < x
Fx:generic_format beta fexp x
Ly:(mag y <= fexp (mag x))%Z
Hy:y < ulp beta fexp (bpow (mag x - 1))

generic_format beta fexp (bpow (mag x - 1))
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Px:0 < x
Py:0 < y
Yltx:y < x
Xgtpow:bpow (mag x - 1) < x
Fx:generic_format beta fexp x
Ly:(mag y <= fexp (mag x))%Z
Hy:y < ulp beta fexp (bpow (mag x - 1))

(fexp (mag x - 1 + 1) <= mag x - 1)%Z
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Px:0 < x
Py:0 < y
Yltx:y < x
Xgtpow:bpow (mag x - 1) < x
Fx:generic_format beta fexp x
Ly:(mag y <= fexp (mag x))%Z
Hy:y < ulp beta fexp (bpow (mag x - 1))

(fexp (mag x) <= mag x - 1)%Z
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Px:0 < x
Py:0 < y
Yltx:y < x
Xgtpow:bpow (mag x - 1) < x
Fx:generic_format beta fexp x
Ly:(mag y <= fexp (mag x))%Z
Hy:y < ulp beta fexp (bpow (mag x - 1))

(fexp (mag x) < mag x)%Z
now apply mag_generic_gt; [|now apply Rgt_not_eq|].
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Px:0 < x
Py:0 < y
Yltx:y < x
Xgtpow:bpow (mag x - 1) < x
Fx:generic_format beta fexp x
Ly:(mag y <= fexp (mag x))%Z

Rabs (x - y) < bpow (mag x)
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Px:0 < x
Py:0 < y
Yltx:y < x
Xgtpow:bpow (mag x - 1) < x
Fx:generic_format beta fexp x
Ly:(mag y <= fexp (mag x))%Z

x - y < bpow (mag x)
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Px:0 < x
Py:0 < y
Yltx:y < x
Xgtpow:bpow (mag x - 1) < x
Fx:generic_format beta fexp x
Ly:(mag y <= fexp (mag x))%Z
x - y >= 0
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Px:0 < x
Py:0 < y
Yltx:y < x
Xgtpow:bpow (mag x - 1) < x
Fx:generic_format beta fexp x
Ly:(mag y <= fexp (mag x))%Z

x - y < bpow (mag x)
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Px:0 < x
Py:0 < y
Yltx:y < x
Xgtpow:bpow (mag x - 1) < x
Fx:generic_format beta fexp x
Ly:(mag y <= fexp (mag x))%Z

x - y < x
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Px:0 < x
Py:0 < y
Yltx:y < x
Xgtpow:bpow (mag x - 1) < x
Fx:generic_format beta fexp x
Ly:(mag y <= fexp (mag x))%Z
x < bpow (mag x)
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Px:0 < x
Py:0 < y
Yltx:y < x
Xgtpow:bpow (mag x - 1) < x
Fx:generic_format beta fexp x
Ly:(mag y <= fexp (mag x))%Z

x - y < x
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Px:0 < x
Py:0 < y
Yltx:y < x
Xgtpow:bpow (mag x - 1) < x
Fx:generic_format beta fexp x
Ly:(mag y <= fexp (mag x))%Z

x - y < x + 0
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Px:0 < x
Py:0 < y
Yltx:y < x
Xgtpow:bpow (mag x - 1) < x
Fx:generic_format beta fexp x
Ly:(mag y <= fexp (mag x))%Z

- y < 0
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Px:0 < x
Py:0 < y
Yltx:y < x
Xgtpow:bpow (mag x - 1) < x
Fx:generic_format beta fexp x
Ly:(mag y <= fexp (mag x))%Z

- y < - 0
now apply Ropp_lt_contravar.
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Px:0 < x
Py:0 < y
Yltx:y < x
Xgtpow:bpow (mag x - 1) < x
Fx:generic_format beta fexp x
Ly:(mag y <= fexp (mag x))%Z

x < bpow (mag x)
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Px:0 < x
Py:0 < y
Yltx:y < x
Xgtpow:bpow (mag x - 1) < x
Fx:generic_format beta fexp x
Ly:(mag y <= fexp (mag x))%Z

Rabs x < bpow (mag x)
apply bpow_mag_gt.
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Px:0 < x
Py:0 < y
Yltx:y < x
Xgtpow:bpow (mag x - 1) < x
Fx:generic_format beta fexp x
Ly:(mag y <= fexp (mag x))%Z

x - y >= 0
lra. Qed. Definition round_round_plus_hyp fexp1 fexp2 := (forall ex ey, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z) /\ (forall ex ey, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z) /\ (forall ex ey, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z) /\ (forall ex ey, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z).
beta:radix

forall (fexp1 fexp2 : Z -> Z) (x y : R), (fexp1 (mag x) <= fexp1 (mag y))%Z -> (fexp2 (mag (x + y)) <= fexp1 (mag x))%Z -> (fexp2 (mag (x + y)) <= fexp1 (mag y))%Z -> generic_format beta fexp1 x -> generic_format beta fexp1 y -> generic_format beta fexp2 (x + y)
beta:radix

forall (fexp1 fexp2 : Z -> Z) (x y : R), (fexp1 (mag x) <= fexp1 (mag y))%Z -> (fexp2 (mag (x + y)) <= fexp1 (mag x))%Z -> (fexp2 (mag (x + y)) <= fexp1 (mag y))%Z -> generic_format beta fexp1 x -> generic_format beta fexp1 y -> generic_format beta fexp2 (x + y)
beta:radix
fexp1, fexp2:Z -> Z
x, y:R
Oxy:(fexp1 (mag x) <= fexp1 (mag y))%Z
Hlnx:(fexp2 (mag (x + y)) <= fexp1 (mag x))%Z
Hlny:(fexp2 (mag (x + y)) <= fexp1 (mag y))%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y

generic_format beta fexp2 (x + y)
beta:radix
fexp1, fexp2:Z -> Z
x, y:R
Oxy:(fexp1 (mag x) <= fexp1 (mag y))%Z
Hlnx:(fexp2 (mag (x + y)) <= fexp1 (mag x))%Z
Hlny:(fexp2 (mag (x + y)) <= fexp1 (mag y))%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0

generic_format beta fexp2 (x + y)
beta:radix
fexp1, fexp2:Z -> Z
x, y:R
Oxy:(fexp1 (mag x) <= fexp1 (mag y))%Z
Hlnx:(fexp2 (mag (x + y)) <= fexp1 (mag x))%Z
Hlny:(fexp2 (mag (x + y)) <= fexp1 (mag y))%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
generic_format beta fexp2 (x + y)
beta:radix
fexp1, fexp2:Z -> Z
x, y:R
Oxy:(fexp1 (mag x) <= fexp1 (mag y))%Z
Hlnx:(fexp2 (mag (x + y)) <= fexp1 (mag x))%Z
Hlny:(fexp2 (mag (x + y)) <= fexp1 (mag y))%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0

generic_format beta fexp2 (x + y)
beta:radix
fexp1, fexp2:Z -> Z
x, y:R
Oxy:(fexp1 (mag x) <= fexp1 (mag y))%Z
Hlnx:(fexp2 (mag (x + y)) <= fexp1 (mag x))%Z
Hlny:(fexp2 (mag y) <= fexp1 (mag y))%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0

generic_format beta fexp2 y
now apply (generic_inclusion_mag beta fexp1).
beta:radix
fexp1, fexp2:Z -> Z
x, y:R
Oxy:(fexp1 (mag x) <= fexp1 (mag y))%Z
Hlnx:(fexp2 (mag (x + y)) <= fexp1 (mag x))%Z
Hlny:(fexp2 (mag (x + y)) <= fexp1 (mag y))%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0

generic_format beta fexp2 (x + y)
beta:radix
fexp1, fexp2:Z -> Z
x, y:R
Oxy:(fexp1 (mag x) <= fexp1 (mag y))%Z
Hlnx:(fexp2 (mag (x + y)) <= fexp1 (mag x))%Z
Hlny:(fexp2 (mag (x + y)) <= fexp1 (mag y))%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0

generic_format beta fexp2 (x + y)
beta:radix
fexp1, fexp2:Z -> Z
x, y:R
Oxy:(fexp1 (mag x) <= fexp1 (mag y))%Z
Hlnx:(fexp2 (mag (x + y)) <= fexp1 (mag x))%Z
Hlny:(fexp2 (mag (x + y)) <= fexp1 (mag y))%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Nzy:y <> 0
generic_format beta fexp2 (x + y)
beta:radix
fexp1, fexp2:Z -> Z
x, y:R
Oxy:(fexp1 (mag x) <= fexp1 (mag y))%Z
Hlnx:(fexp2 (mag (x + y)) <= fexp1 (mag x))%Z
Hlny:(fexp2 (mag (x + y)) <= fexp1 (mag y))%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0

generic_format beta fexp2 (x + y)
beta:radix
fexp1, fexp2:Z -> Z
x, y:R
Oxy:(fexp1 (mag x) <= fexp1 (mag y))%Z
Hlnx:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hlny:(fexp2 (mag (x + y)) <= fexp1 (mag y))%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0

generic_format beta fexp2 x
now apply (generic_inclusion_mag beta fexp1).
beta:radix
fexp1, fexp2:Z -> Z
x, y:R
Oxy:(fexp1 (mag x) <= fexp1 (mag y))%Z
Hlnx:(fexp2 (mag (x + y)) <= fexp1 (mag x))%Z
Hlny:(fexp2 (mag (x + y)) <= fexp1 (mag y))%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Nzy:y <> 0

generic_format beta fexp2 (x + y)
beta:radix
fexp1, fexp2:Z -> Z
x, y:R
Oxy:(fexp1 (mag x) <= fexp1 (mag y))%Z
Hlnx:(fexp2 (mag (x + y)) <= fexp1 (mag x))%Z
Hlny:(fexp2 (mag (x + y)) <= fexp1 (mag y))%Z
Nzx:x <> 0
Nzy:y <> 0

generic_format beta fexp1 x -> generic_format beta fexp1 y -> generic_format beta fexp2 (x + y)
beta:radix
fexp1, fexp2:Z -> Z
x, y:R
Oxy:(fexp1 (mag x) <= fexp1 (mag y))%Z
Hlnx:(fexp2 (mag (x + y)) <= fexp1 (mag x))%Z
Hlny:(fexp2 (mag (x + y)) <= fexp1 (mag y))%Z
Nzx:x <> 0
Nzy:y <> 0

x = IZR (Ztrunc (scaled_mantissa beta fexp1 x)) * bpow (fexp1 (mag x)) -> y = IZR (Ztrunc (scaled_mantissa beta fexp1 y)) * bpow (fexp1 (mag y)) -> generic_format beta fexp2 (x + y)
beta:radix
fexp1, fexp2:Z -> Z
x, y:R
Oxy:(fexp1 (mag x) <= fexp1 (mag y))%Z
Hlnx:(fexp2 (mag (x + y)) <= fexp1 (mag x))%Z
Hlny:(fexp2 (mag (x + y)) <= fexp1 (mag y))%Z
Nzx:x <> 0
Nzy:y <> 0
mx:=Ztrunc (scaled_mantissa beta fexp1 x):Z

x = IZR mx * bpow (fexp1 (mag x)) -> y = IZR (Ztrunc (scaled_mantissa beta fexp1 y)) * bpow (fexp1 (mag y)) -> generic_format beta fexp2 (x + y)
beta:radix
fexp1, fexp2:Z -> Z
x, y:R
Oxy:(fexp1 (mag x) <= fexp1 (mag y))%Z
Hlnx:(fexp2 (mag (x + y)) <= fexp1 (mag x))%Z
Hlny:(fexp2 (mag (x + y)) <= fexp1 (mag y))%Z
Nzx:x <> 0
Nzy:y <> 0
mx:=Ztrunc (scaled_mantissa beta fexp1 x):Z
my:=Ztrunc (scaled_mantissa beta fexp1 y):Z

x = IZR mx * bpow (fexp1 (mag x)) -> y = IZR my * bpow (fexp1 (mag y)) -> generic_format beta fexp2 (x + y)
beta:radix
fexp1, fexp2:Z -> Z
x, y:R
Oxy:(fexp1 (mag x) <= fexp1 (mag y))%Z
Hlnx:(fexp2 (mag (x + y)) <= fexp1 (mag x))%Z
Hlny:(fexp2 (mag (x + y)) <= fexp1 (mag y))%Z
Nzx:x <> 0
Nzy:y <> 0
mx:=Ztrunc (scaled_mantissa beta fexp1 x):Z
my:=Ztrunc (scaled_mantissa beta fexp1 y):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))

generic_format beta fexp2 (x + y)
beta:radix
fexp1, fexp2:Z -> Z
x, y:R
Oxy:(fexp1 (mag x) <= fexp1 (mag y))%Z
Hlnx:(fexp2 (mag (x + y)) <= fexp1 (mag x))%Z
Hlny:(fexp2 (mag (x + y)) <= fexp1 (mag y))%Z
Nzx:x <> 0
Nzy:y <> 0
mx:=Ztrunc (scaled_mantissa beta fexp1 x):Z
my:=Ztrunc (scaled_mantissa beta fexp1 y):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
fxy:={| Fnum := mx + my * beta ^ (fexp1 (mag y) - fexp1 (mag x)); Fexp := fexp1 (mag x) |}:float beta

generic_format beta fexp2 (x + y)
beta:radix
fexp1, fexp2:Z -> Z
x, y:R
Oxy:(fexp1 (mag x) <= fexp1 (mag y))%Z
Hlnx:(fexp2 (mag (x + y)) <= fexp1 (mag x))%Z
Hlny:(fexp2 (mag (x + y)) <= fexp1 (mag y))%Z
Nzx:x <> 0
Nzy:y <> 0
mx:=Ztrunc (scaled_mantissa beta fexp1 x):Z
my:=Ztrunc (scaled_mantissa beta fexp1 y):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
fxy:={| Fnum := mx + my * beta ^ (fexp1 (mag y) - fexp1 (mag x)); Fexp := fexp1 (mag x) |}:float beta

x + y = F2R fxy
beta:radix
fexp1, fexp2:Z -> Z
x, y:R
Oxy:(fexp1 (mag x) <= fexp1 (mag y))%Z
Hlnx:(fexp2 (mag (x + y)) <= fexp1 (mag x))%Z
Hlny:(fexp2 (mag (x + y)) <= fexp1 (mag y))%Z
Nzx:x <> 0
Nzy:y <> 0
mx:=Ztrunc (scaled_mantissa beta fexp1 x):Z
my:=Ztrunc (scaled_mantissa beta fexp1 y):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
fxy:={| Fnum := mx + my * beta ^ (fexp1 (mag y) - fexp1 (mag x)); Fexp := fexp1 (mag x) |}:float beta
Hxy:x + y = F2R fxy
generic_format beta fexp2 (x + y)
beta:radix
fexp1, fexp2:Z -> Z
x, y:R
Oxy:(fexp1 (mag x) <= fexp1 (mag y))%Z
Hlnx:(fexp2 (mag (x + y)) <= fexp1 (mag x))%Z
Hlny:(fexp2 (mag (x + y)) <= fexp1 (mag y))%Z
Nzx:x <> 0
Nzy:y <> 0
mx:=Ztrunc (scaled_mantissa beta fexp1 x):Z
my:=Ztrunc (scaled_mantissa beta fexp1 y):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
fxy:={| Fnum := mx + my * beta ^ (fexp1 (mag y) - fexp1 (mag x)); Fexp := fexp1 (mag x) |}:float beta

x + y = F2R fxy
beta:radix
fexp1, fexp2:Z -> Z
x, y:R
Oxy:(fexp1 (mag x) <= fexp1 (mag y))%Z
Hlnx:(fexp2 (mag (x + y)) <= fexp1 (mag x))%Z
Hlny:(fexp2 (mag (x + y)) <= fexp1 (mag y))%Z
Nzx:x <> 0
Nzy:y <> 0
mx:=Ztrunc (scaled_mantissa beta fexp1 x):Z
my:=Ztrunc (scaled_mantissa beta fexp1 y):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
fxy:={| Fnum := mx + my * beta ^ (fexp1 (mag y) - fexp1 (mag x)); Fexp := fexp1 (mag x) |}:float beta

x + y = IZR (mx + my * beta ^ (fexp1 (mag y) - fexp1 (mag x))) * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
x, y:R
Oxy:(fexp1 (mag x) <= fexp1 (mag y))%Z
Hlnx:(fexp2 (mag (x + y)) <= fexp1 (mag x))%Z
Hlny:(fexp2 (mag (x + y)) <= fexp1 (mag y))%Z
Nzx:x <> 0
Nzy:y <> 0
mx:=Ztrunc (scaled_mantissa beta fexp1 x):Z
my:=Ztrunc (scaled_mantissa beta fexp1 y):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
fxy:={| Fnum := mx + my * beta ^ (fexp1 (mag y) - fexp1 (mag x)); Fexp := fexp1 (mag x) |}:float beta

x + y = (IZR mx + IZR (my * beta ^ (fexp1 (mag y) - fexp1 (mag x)))) * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
x, y:R
Oxy:(fexp1 (mag x) <= fexp1 (mag y))%Z
Hlnx:(fexp2 (mag (x + y)) <= fexp1 (mag x))%Z
Hlny:(fexp2 (mag (x + y)) <= fexp1 (mag y))%Z
Nzx:x <> 0
Nzy:y <> 0
mx:=Ztrunc (scaled_mantissa beta fexp1 x):Z
my:=Ztrunc (scaled_mantissa beta fexp1 y):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
fxy:={| Fnum := mx + my * beta ^ (fexp1 (mag y) - fexp1 (mag x)); Fexp := fexp1 (mag x) |}:float beta

x + y = IZR mx * bpow (fexp1 (mag x)) + IZR (my * beta ^ (fexp1 (mag y) - fexp1 (mag x))) * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
x, y:R
Oxy:(fexp1 (mag x) <= fexp1 (mag y))%Z
Hlnx:(fexp2 (mag (x + y)) <= fexp1 (mag x))%Z
Hlny:(fexp2 (mag (x + y)) <= fexp1 (mag y))%Z
Nzx:x <> 0
Nzy:y <> 0
mx:=Ztrunc (scaled_mantissa beta fexp1 x):Z
my:=Ztrunc (scaled_mantissa beta fexp1 y):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
fxy:={| Fnum := mx + my * beta ^ (fexp1 (mag y) - fexp1 (mag x)); Fexp := fexp1 (mag x) |}:float beta

x + y = x + IZR (my * beta ^ (fexp1 (mag y) - fexp1 (mag x))) * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
x, y:R
Oxy:(fexp1 (mag x) <= fexp1 (mag y))%Z
Hlnx:(fexp2 (mag (x + y)) <= fexp1 (mag x))%Z
Hlny:(fexp2 (mag (x + y)) <= fexp1 (mag y))%Z
Nzx:x <> 0
Nzy:y <> 0
mx:=Ztrunc (scaled_mantissa beta fexp1 x):Z
my:=Ztrunc (scaled_mantissa beta fexp1 y):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
fxy:={| Fnum := mx + my * beta ^ (fexp1 (mag y) - fexp1 (mag x)); Fexp := fexp1 (mag x) |}:float beta

x + y = x + IZR my * IZR (beta ^ (fexp1 (mag y) - fexp1 (mag x))) * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
x, y:R
Oxy:(fexp1 (mag x) <= fexp1 (mag y))%Z
Hlnx:(fexp2 (mag (x + y)) <= fexp1 (mag x))%Z
Hlny:(fexp2 (mag (x + y)) <= fexp1 (mag y))%Z
Nzx:x <> 0
Nzy:y <> 0
mx:=Ztrunc (scaled_mantissa beta fexp1 x):Z
my:=Ztrunc (scaled_mantissa beta fexp1 y):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
fxy:={| Fnum := mx + my * beta ^ (fexp1 (mag y) - fexp1 (mag x)); Fexp := fexp1 (mag x) |}:float beta

x + y = x + IZR my * bpow (fexp1 (mag y) - fexp1 (mag x)) * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
x, y:R
Oxy:(fexp1 (mag x) <= fexp1 (mag y))%Z
Hlnx:(fexp2 (mag (x + y)) <= fexp1 (mag x))%Z
Hlny:(fexp2 (mag (x + y)) <= fexp1 (mag y))%Z
Nzx:x <> 0
Nzy:y <> 0
mx:=Ztrunc (scaled_mantissa beta fexp1 x):Z
my:=Ztrunc (scaled_mantissa beta fexp1 y):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
fxy:={| Fnum := mx + my * beta ^ (fexp1 (mag y) - fexp1 (mag x)); Fexp := fexp1 (mag x) |}:float beta

x + y = x + IZR my * bpow (fexp1 (mag y))
now rewrite <- Fy.
beta:radix
fexp1, fexp2:Z -> Z
x, y:R
Oxy:(fexp1 (mag x) <= fexp1 (mag y))%Z
Hlnx:(fexp2 (mag (x + y)) <= fexp1 (mag x))%Z
Hlny:(fexp2 (mag (x + y)) <= fexp1 (mag y))%Z
Nzx:x <> 0
Nzy:y <> 0
mx:=Ztrunc (scaled_mantissa beta fexp1 x):Z
my:=Ztrunc (scaled_mantissa beta fexp1 y):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
fxy:={| Fnum := mx + my * beta ^ (fexp1 (mag y) - fexp1 (mag x)); Fexp := fexp1 (mag x) |}:float beta
Hxy:x + y = F2R fxy

generic_format beta fexp2 (x + y)
beta:radix
fexp1, fexp2:Z -> Z
x, y:R
Oxy:(fexp1 (mag x) <= fexp1 (mag y))%Z
Hlnx:(fexp2 (mag (x + y)) <= fexp1 (mag x))%Z
Hlny:(fexp2 (mag (x + y)) <= fexp1 (mag y))%Z
Nzx:x <> 0
Nzy:y <> 0
mx:=Ztrunc (scaled_mantissa beta fexp1 x):Z
my:=Ztrunc (scaled_mantissa beta fexp1 y):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
fxy:={| Fnum := mx + my * beta ^ (fexp1 (mag y) - fexp1 (mag x)); Fexp := fexp1 (mag x) |}:float beta
Hxy:x + y = F2R fxy

x + y <> 0 -> (cexp beta fexp2 (x + y) <= Fexp fxy)%Z
beta:radix
fexp1, fexp2:Z -> Z
x, y:R
Oxy:(fexp1 (mag x) <= fexp1 (mag y))%Z
Hlnx:(fexp2 (mag (x + y)) <= fexp1 (mag x))%Z
Hlny:(fexp2 (mag (x + y)) <= fexp1 (mag y))%Z
Nzx:x <> 0
Nzy:y <> 0
mx:=Ztrunc (scaled_mantissa beta fexp1 x):Z
my:=Ztrunc (scaled_mantissa beta fexp1 y):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
fxy:={| Fnum := mx + my * beta ^ (fexp1 (mag y) - fexp1 (mag x)); Fexp := fexp1 (mag x) |}:float beta
Hxy:x + y = F2R fxy

(cexp beta fexp2 (x + y) <= Fexp fxy)%Z
now unfold cexp, fxy; simpl. Qed.
beta:radix

forall (fexp1 fexp2 : Z -> Z) (x y : R), (fexp2 (mag (x + y)) <= fexp1 (mag x))%Z -> (fexp2 (mag (x + y)) <= fexp1 (mag y))%Z -> generic_format beta fexp1 x -> generic_format beta fexp1 y -> generic_format beta fexp2 (x + y)
beta:radix

forall (fexp1 fexp2 : Z -> Z) (x y : R), (fexp2 (mag (x + y)) <= fexp1 (mag x))%Z -> (fexp2 (mag (x + y)) <= fexp1 (mag y))%Z -> generic_format beta fexp1 x -> generic_format beta fexp1 y -> generic_format beta fexp2 (x + y)
beta:radix
fexp1, fexp2:Z -> Z
x, y:R
Hlnx:(fexp2 (mag (x + y)) <= fexp1 (mag x))%Z
Hlny:(fexp2 (mag (x + y)) <= fexp1 (mag y))%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y

generic_format beta fexp2 (x + y)
beta:radix
fexp1, fexp2:Z -> Z
x, y:R
Hlnx:(fexp2 (mag (x + y)) <= fexp1 (mag x))%Z
Hlny:(fexp2 (mag (x + y)) <= fexp1 (mag y))%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hle:(fexp1 (mag x) <= fexp1 (mag y))%Z

generic_format beta fexp2 (x + y)
beta:radix
fexp1, fexp2:Z -> Z
x, y:R
Hlnx:(fexp2 (mag (x + y)) <= fexp1 (mag x))%Z
Hlny:(fexp2 (mag (x + y)) <= fexp1 (mag y))%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hgt:(fexp1 (mag y) < fexp1 (mag x))%Z
generic_format beta fexp2 (x + y)
beta:radix
fexp1, fexp2:Z -> Z
x, y:R
Hlnx:(fexp2 (mag (x + y)) <= fexp1 (mag x))%Z
Hlny:(fexp2 (mag (x + y)) <= fexp1 (mag y))%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hle:(fexp1 (mag x) <= fexp1 (mag y))%Z

generic_format beta fexp2 (x + y)
now apply (round_round_plus_aux0_aux_aux fexp1).
beta:radix
fexp1, fexp2:Z -> Z
x, y:R
Hlnx:(fexp2 (mag (x + y)) <= fexp1 (mag x))%Z
Hlny:(fexp2 (mag (x + y)) <= fexp1 (mag y))%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hgt:(fexp1 (mag y) < fexp1 (mag x))%Z

generic_format beta fexp2 (x + y)
beta:radix
fexp1, fexp2:Z -> Z
x, y:R
Hlnx:(fexp2 (mag (y + x)) <= fexp1 (mag x))%Z
Hlny:(fexp2 (mag (y + x)) <= fexp1 (mag y))%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hgt:(fexp1 (mag y) < fexp1 (mag x))%Z

generic_format beta fexp2 (y + x)
now apply (round_round_plus_aux0_aux_aux fexp1); [omega| | | |]. Qed. (* fexp1 (mag x) - 1 <= mag y : * addition is exact in the largest precision (fexp2). *)
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> round_round_plus_hyp fexp1 fexp2 -> forall x y : R, 0 < x -> 0 < y -> y <= x -> (fexp1 (mag x) - 1 <= mag y)%Z -> generic_format beta fexp1 x -> generic_format beta fexp1 y -> generic_format beta fexp2 (x + y)
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> round_round_plus_hyp fexp1 fexp2 -> forall x y : R, 0 < x -> 0 < y -> y <= x -> (fexp1 (mag x) - 1 <= mag y)%Z -> generic_format beta fexp1 x -> generic_format beta fexp1 y -> generic_format beta fexp2 (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hyx:y <= x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y

generic_format beta fexp2 (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hyx:y <= x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nny:0 <= y

generic_format beta fexp2 (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hyx:y <= x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nny:0 <= y

generic_format beta fexp2 (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hyx:y <= x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nny:0 <= y
Hle:(mag y <= fexp1 (mag x))%Z

generic_format beta fexp2 (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hyx:y <= x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nny:0 <= y
Hgt:(fexp1 (mag x) < mag y)%Z
generic_format beta fexp2 (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hyx:y <= x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nny:0 <= y
Hle:(mag y <= fexp1 (mag x))%Z

generic_format beta fexp2 (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hyx:y <= x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nny:0 <= y
Hle:(mag y <= fexp1 (mag x))%Z
Lxy:mag (x + y) = mag x

generic_format beta fexp2 (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hyx:y <= x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nny:0 <= y
Hle:(mag y <= fexp1 (mag x))%Z
Lxy:mag (x + y) = mag x

(fexp2 (mag x) <= fexp1 (mag x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hyx:y <= x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nny:0 <= y
Hle:(mag y <= fexp1 (mag x))%Z
Lxy:mag (x + y) = mag x
(fexp2 (mag x) <= fexp1 (mag y))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hyx:y <= x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nny:0 <= y
Hle:(mag y <= fexp1 (mag x))%Z
Lxy:mag (x + y) = mag x

(fexp2 (mag x) <= fexp1 (mag x))%Z
now apply Hexp4; omega.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hyx:y <= x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nny:0 <= y
Hle:(mag y <= fexp1 (mag x))%Z
Lxy:mag (x + y) = mag x

(fexp2 (mag x) <= fexp1 (mag y))%Z
now apply Hexp3; omega.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hyx:y <= x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nny:0 <= y
Hgt:(fexp1 (mag x) < mag y)%Z

generic_format beta fexp2 (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hyx:y <= x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nny:0 <= y
Hgt:(fexp1 (mag x) < mag y)%Z

(fexp2 (mag (x + y)) <= fexp1 (mag x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hyx:y <= x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nny:0 <= y
Hgt:(fexp1 (mag x) < mag y)%Z
(fexp2 (mag (x + y)) <= fexp1 (mag y))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hyx:y <= x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nny:0 <= y
Hgt:(fexp1 (mag x) < mag y)%Z
Lxy:mag (x + y) = mag x

(fexp2 (mag x) <= fexp1 (mag x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hyx:y <= x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nny:0 <= y
Hgt:(fexp1 (mag x) < mag y)%Z
Lxy:mag (x + y) = (mag x + 1)%Z
(fexp2 (mag x + 1) <= fexp1 (mag x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hyx:y <= x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nny:0 <= y
Hgt:(fexp1 (mag x) < mag y)%Z
(fexp2 (mag (x + y)) <= fexp1 (mag y))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hyx:y <= x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nny:0 <= y
Hgt:(fexp1 (mag x) < mag y)%Z
Lxy:mag (x + y) = mag x

(fexp2 (mag x) <= fexp1 (mag x))%Z
now apply Hexp4; omega.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hyx:y <= x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nny:0 <= y
Hgt:(fexp1 (mag x) < mag y)%Z
Lxy:mag (x + y) = (mag x + 1)%Z

(fexp2 (mag x + 1) <= fexp1 (mag x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hyx:(mag y <= mag x)%Z
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nny:0 <= y
Hgt:(fexp1 (mag x) < mag y)%Z
Lxy:mag (x + y) = (mag x + 1)%Z

(fexp1 (mag x + 1 - 1) + 1 <= mag x)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hyx:(mag y <= mag x)%Z
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nny:0 <= y
Hgt:(fexp1 (mag x) < mag y)%Z
Lxy:mag (x + y) = (mag x + 1)%Z

(fexp1 (mag x) + 1 <= mag x)%Z
omega.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hyx:y <= x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nny:0 <= y
Hgt:(fexp1 (mag x) < mag y)%Z

(fexp2 (mag (x + y)) <= fexp1 (mag y))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hyx:y <= x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nny:0 <= y
Hgt:(fexp1 (mag x) < mag y)%Z
Lxy:mag (x + y) = mag x

(fexp2 (mag x) <= fexp1 (mag y))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hyx:y <= x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nny:0 <= y
Hgt:(fexp1 (mag x) < mag y)%Z
Lxy:mag (x + y) = (mag x + 1)%Z
(fexp2 (mag x + 1) <= fexp1 (mag y))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hyx:y <= x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nny:0 <= y
Hgt:(fexp1 (mag x) < mag y)%Z
Lxy:mag (x + y) = mag x

(fexp2 (mag x) <= fexp1 (mag y))%Z
now apply Hexp3; omega.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hyx:y <= x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nny:0 <= y
Hgt:(fexp1 (mag x) < mag y)%Z
Lxy:mag (x + y) = (mag x + 1)%Z

(fexp2 (mag x + 1) <= fexp1 (mag y))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hyx:y <= x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nny:0 <= y
Hgt:(fexp1 (mag x) < mag y)%Z
Lxy:mag (x + y) = (mag x + 1)%Z

(fexp1 (mag x + 1 - 1) + 1 <= mag y)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hyx:y <= x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nny:0 <= y
Hgt:(fexp1 (mag x) < mag y)%Z
Lxy:mag (x + y) = (mag x + 1)%Z

(fexp1 (mag x) + 1 <= mag y)%Z
omega. Qed.
beta:radix

forall k : Z, (0 < k)%Z -> forall (fexp : Z -> Z) (x y : R), 0 < x -> 0 < y -> (mag y <= fexp (mag x) - k)%Z -> mag (x + y) = mag x -> generic_format beta fexp x -> 0 < x + y - round beta fexp Zfloor (x + y) < bpow (fexp (mag x) - k)
beta:radix

forall k : Z, (0 < k)%Z -> forall (fexp : Z -> Z) (x y : R), 0 < x -> 0 < y -> (mag y <= fexp (mag x) - k)%Z -> mag (x + y) = mag x -> generic_format beta fexp x -> 0 < x + y - round beta fexp Zfloor (x + y) < bpow (fexp (mag x) - k)
beta:radix

(2 <= beta)%Z
beta:radix
Hbeta:(2 <= beta)%Z
forall k : Z, (0 < k)%Z -> forall (fexp : Z -> Z) (x y : R), 0 < x -> 0 < y -> (mag y <= fexp (mag x) - k)%Z -> mag (x + y) = mag x -> generic_format beta fexp x -> 0 < x + y - round beta fexp Zfloor (x + y) < bpow (fexp (mag x) - k)
beta:radix

(2 <= beta)%Z
beta:radix
beta_val:Z
beta_prop:(2 <=? beta_val)%Z = true

(2 <= {| radix_val := beta_val; radix_prop := beta_prop |})%Z
now apply Zle_bool_imp_le.
beta:radix
Hbeta:(2 <= beta)%Z

forall k : Z, (0 < k)%Z -> forall (fexp : Z -> Z) (x y : R), 0 < x -> 0 < y -> (mag y <= fexp (mag x) - k)%Z -> mag (x + y) = mag x -> generic_format beta fexp x -> 0 < x + y - round beta fexp Zfloor (x + y) < bpow (fexp (mag x) - k)
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
Fx:generic_format beta fexp x

0 < x + y - round beta fexp Zfloor (x + y) < bpow (fexp (mag x) - k)
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x

generic_format beta fexp x -> 0 < x + y - round beta fexp Zfloor (x + y) < bpow (fexp (mag x) - k)
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x

x = IZR (Ztrunc (x * bpow (- fexp (mag x)))) * bpow (fexp (mag x)) -> 0 < x + y - IZR (Zfloor ((x + y) * bpow (- fexp (mag (x + y))))) * bpow (fexp (mag (x + y))) < bpow (fexp (mag x) - k)
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x

x = IZR (Ztrunc (x * bpow (- fexp (mag x)))) * bpow (fexp (mag x)) -> 0 < x + y - IZR (Zfloor ((x + y) * bpow (- fexp (mag x)))) * bpow (fexp (mag x)) < bpow (fexp (mag x) - k)
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z

x = IZR mx * bpow (fexp (mag x)) -> 0 < x + y - IZR (Zfloor ((x + y) * bpow (- fexp (mag x)))) * bpow (fexp (mag x)) < bpow (fexp (mag x) - k)
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))

0 < x + y - IZR (Zfloor ((x + y) * bpow (- fexp (mag x)))) * bpow (fexp (mag x)) < bpow (fexp (mag x) - k)
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))

(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
0 < x + y - IZR (Zfloor ((x + y) * bpow (- fexp (mag x)))) * bpow (fexp (mag x)) < bpow (fexp (mag x) - k)
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))

(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))

(IZR mx * bpow (fexp (mag x)) + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))

IZR mx * bpow (fexp (mag x)) * bpow (- fexp (mag x)) + y * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
now bpow_simplify.
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))

0 < x + y - IZR (Zfloor ((x + y) * bpow (- fexp (mag x)))) * bpow (fexp (mag x)) < bpow (fexp (mag x) - k)
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))

0 < x + y - IZR (Zfloor (IZR mx + y * bpow (- fexp (mag x)))) * bpow (fexp (mag x)) < bpow (fexp (mag x) - k)
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))

0 < y * bpow (- fexp (mag x))
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))
0 < x + y - IZR (Zfloor (IZR mx + y * bpow (- fexp (mag x)))) * bpow (fexp (mag x)) < bpow (fexp (mag x) - k)
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))

0 < y * bpow (- fexp (mag x))
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))

y * 0 < y * bpow (- fexp (mag x))
now apply Rmult_lt_compat_l; [|apply bpow_gt_0].
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))

0 < x + y - IZR (Zfloor (IZR mx + y * bpow (- fexp (mag x)))) * bpow (fexp (mag x)) < bpow (fexp (mag x) - k)
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))

y * bpow (- fexp (mag x)) < / IZR (beta ^ k)
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))
UB:y * bpow (- fexp (mag x)) < / IZR (beta ^ k)
0 < x + y - IZR (Zfloor (IZR mx + y * bpow (- fexp (mag x)))) * bpow (fexp (mag x)) < bpow (fexp (mag x) - k)
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))

y * bpow (- fexp (mag x)) < / IZR (beta ^ k)
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))

y * bpow (- fexp (mag x)) < bpow (mag y) * bpow (- fexp (mag x))
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))
bpow (mag y) * bpow (- fexp (mag x)) <= / IZR (beta ^ k)
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))

y * bpow (- fexp (mag x)) < bpow (mag y) * bpow (- fexp (mag x))
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))

y < bpow (mag y)
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))

Rabs y < bpow (mag y)
apply bpow_mag_gt.
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))

bpow (mag y) * bpow (- fexp (mag x)) <= / IZR (beta ^ k)
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))

bpow (mag y) * bpow (- fexp (mag x)) <= bpow (fexp (mag x) - k) * bpow (- fexp (mag x))
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))
bpow (fexp (mag x) - k) * bpow (- fexp (mag x)) <= / IZR (beta ^ k)
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))

bpow (mag y) * bpow (- fexp (mag x)) <= bpow (fexp (mag x) - k) * bpow (- fexp (mag x))
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))

bpow (mag y) <= bpow (fexp (mag x) - k)
now apply bpow_le.
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))

bpow (fexp (mag x) - k) * bpow (- fexp (mag x)) <= / IZR (beta ^ k)
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))

bpow (- k) <= / IZR (beta ^ k)
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))

/ bpow k <= / IZR (beta ^ k)
beta:radix
Hbeta:(2 <= beta)%Z
Hk:(0 < 0)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - 0)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))

/ bpow 0 <= / IZR (beta ^ 0)
beta:radix
Hbeta:(2 <= beta)%Z
p:positive
Hk:(0 < Z.pos p)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - Z.pos p)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))
/ bpow (Z.pos p) <= / IZR (beta ^ Z.pos p)
beta:radix
Hbeta:(2 <= beta)%Z
p:positive
Hk:(0 < Z.neg p)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - Z.neg p)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))
/ bpow (Z.neg p) <= / IZR (beta ^ Z.neg p)
beta:radix
Hbeta:(2 <= beta)%Z
Hk:(0 < 0)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - 0)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))

/ bpow 0 <= / IZR (beta ^ 0)
omega.
beta:radix
Hbeta:(2 <= beta)%Z
p:positive
Hk:(0 < Z.pos p)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - Z.pos p)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))

/ bpow (Z.pos p) <= / IZR (beta ^ Z.pos p)
beta:radix
Hbeta:(2 <= beta)%Z
p:positive
Hk:(0 < Z.pos p)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - Z.pos p)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))

/ IZR (Pos.iter (Z.mul beta) 1%Z p) <= / IZR (Pos.iter (Z.mul beta) 1%Z p)
now apply Rle_refl.
beta:radix
Hbeta:(2 <= beta)%Z
p:positive
Hk:(0 < Z.neg p)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - Z.neg p)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))

/ bpow (Z.neg p) <= / IZR (beta ^ Z.neg p)
beta:radix
Hbeta:(2 <= beta)%Z
p:positive
Hk:(0 < Z.neg p)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - Z.neg p)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))

(0 < 0)%Z
beta:radix
Hbeta:(2 <= beta)%Z
p:positive
Hk:(0 < Z.neg p)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - Z.neg p)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))

(Z.neg p < 0)%Z
apply Zlt_neg_0.
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))
UB:y * bpow (- fexp (mag x)) < / IZR (beta ^ k)

0 < x + y - IZR (Zfloor (IZR mx + y * bpow (- fexp (mag x)))) * bpow (fexp (mag x)) < bpow (fexp (mag x) - k)
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))
UB:y * bpow (- fexp (mag x)) < / IZR (beta ^ k)

0 < x + y - IZR mx * bpow (fexp (mag x)) < bpow (fexp (mag x) - k)
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))
UB:y * bpow (- fexp (mag x)) < / IZR (beta ^ k)
IZR mx <= IZR mx + y * bpow (- fexp (mag x)) < IZR (mx + 1)
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))
UB:y * bpow (- fexp (mag x)) < / IZR (beta ^ k)

0 < x + y - IZR mx * bpow (fexp (mag x)) < bpow (fexp (mag x) - k)
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))
UB:y * bpow (- fexp (mag x)) < / IZR (beta ^ k)

0 < x + y - IZR mx * bpow (fexp (mag x))
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))
UB:y * bpow (- fexp (mag x)) < / IZR (beta ^ k)
x + y - IZR mx * bpow (fexp (mag x)) < bpow (fexp (mag x) - k)
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))
UB:y * bpow (- fexp (mag x)) < / IZR (beta ^ k)

0 < x + y - IZR mx * bpow (fexp (mag x))
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))
UB:y * bpow (- fexp (mag x)) < / IZR (beta ^ k)

0 * bpow (- fexp (mag x)) < (x + y - IZR mx * bpow (fexp (mag x))) * bpow (- fexp (mag x))
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))
UB:y * bpow (- fexp (mag x)) < / IZR (beta ^ k)

0 < (x + y) * bpow (- fexp (mag x)) - IZR mx * bpow (fexp (mag x)) * bpow (- fexp (mag x))
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))
UB:y * bpow (- fexp (mag x)) < / IZR (beta ^ k)

0 < (x + y) * bpow (- fexp (mag x)) - IZR mx
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))
UB:y * bpow (- fexp (mag x)) < / IZR (beta ^ k)

0 < y * bpow (- fexp (mag x))
now apply Rmult_lt_0_compat; [|apply bpow_gt_0].
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))
UB:y * bpow (- fexp (mag x)) < / IZR (beta ^ k)

x + y - IZR mx * bpow (fexp (mag x)) < bpow (fexp (mag x) - k)
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))
UB:y * bpow (- fexp (mag x)) < / IZR (beta ^ k)

(x + y - IZR mx * bpow (fexp (mag x))) * bpow (- fexp (mag x)) < bpow (fexp (mag x) - k) * bpow (- fexp (mag x))
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))
UB:y * bpow (- fexp (mag x)) < / IZR (beta ^ k)

(x + y) * bpow (- fexp (mag x)) - IZR mx * bpow (fexp (mag x)) * bpow (- fexp (mag x)) < bpow (fexp (mag x) - k) * bpow (- fexp (mag x))
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))
UB:y * bpow (- fexp (mag x)) < / IZR (beta ^ k)

(x + y) * bpow (- fexp (mag x)) - IZR mx < bpow (- k)
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))
UB:y * bpow (- fexp (mag x)) < / IZR (beta ^ k)

y * bpow (- fexp (mag x)) < bpow (- k)
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))
UB:y * bpow (- fexp (mag x)) < / IZR (beta ^ k)

/ IZR (beta ^ k) <= bpow (- k)
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))
UB:y * bpow (- fexp (mag x)) < / IZR (beta ^ k)

/ IZR (beta ^ k) <= / bpow k
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))
UB:y * bpow (- fexp (mag x)) < / IZR (beta ^ k)

bpow k <= IZR (beta ^ k)
now rewrite IZR_Zpower; [right|omega].
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))
UB:y * bpow (- fexp (mag x)) < / IZR (beta ^ k)

IZR mx <= IZR mx + y * bpow (- fexp (mag x)) < IZR (mx + 1)
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))
UB:y * bpow (- fexp (mag x)) < / IZR (beta ^ k)

IZR mx <= IZR mx + y * bpow (- fexp (mag x))
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))
UB:y * bpow (- fexp (mag x)) < / IZR (beta ^ k)
IZR mx + y * bpow (- fexp (mag x)) < IZR (mx + 1)
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))
UB:y * bpow (- fexp (mag x)) < / IZR (beta ^ k)

IZR mx <= IZR mx + y * bpow (- fexp (mag x))
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))
UB:y * bpow (- fexp (mag x)) < / IZR (beta ^ k)

0 <= y * bpow (- fexp (mag x))
now apply Rlt_le.
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))
UB:y * bpow (- fexp (mag x)) < / IZR (beta ^ k)

IZR mx + y * bpow (- fexp (mag x)) < IZR (mx + 1)
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))
UB:y * bpow (- fexp (mag x)) < / IZR (beta ^ k)

y * bpow (- fexp (mag x)) < 1
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))
UB:y * bpow (- fexp (mag x)) < / IZR (beta ^ k)

y * bpow (- fexp (mag x)) * bpow (fexp (mag x)) < 1 * bpow (fexp (mag x))
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))
UB:y * bpow (- fexp (mag x)) < / IZR (beta ^ k)

y * bpow (- fexp (mag x)) * bpow (fexp (mag x)) < bpow (fexp (mag x))
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))
UB:y * bpow (- fexp (mag x)) < / IZR (beta ^ k)

y < bpow (fexp (mag x))
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))
UB:y * bpow (- fexp (mag x)) < / IZR (beta ^ k)

y < bpow (mag y)
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))
UB:y * bpow (- fexp (mag x)) < / IZR (beta ^ k)
bpow (mag y) < bpow (fexp (mag x))
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))
UB:y * bpow (- fexp (mag x)) < / IZR (beta ^ k)

y < bpow (mag y)
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))
UB:y * bpow (- fexp (mag x)) < / IZR (beta ^ k)

Rabs y < bpow (mag y)
apply bpow_mag_gt.
beta:radix
Hbeta:(2 <= beta)%Z
k:Z
Hk:(0 < k)%Z
fexp:Z -> Z
x, y:Rdefinitions.R
Px:0 < x
Py:0 < y
Hln:(mag y <= fexp (mag x) - k)%Z
Hlxy:mag (x + y) = mag x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
R:(x + y) * bpow (- fexp (mag x)) = IZR mx + y * bpow (- fexp (mag x))
LB:0 < y * bpow (- fexp (mag x))
UB:y * bpow (- fexp (mag x)) < / IZR (beta ^ k)

bpow (mag y) < bpow (fexp (mag x))
apply bpow_lt; omega. Qed. (* mag y <= fexp1 (mag x) - 2 : round_round_lt_mid applies. *)
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall choice1 choice2 : Z -> bool, round_round_plus_hyp fexp1 fexp2 -> forall x y : R, 0 < x -> 0 < y -> (mag y <= fexp1 (mag x) - 2)%Z -> generic_format beta fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 (x + y)
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall choice1 choice2 : Z -> bool, round_round_plus_hyp fexp1 fexp2 -> forall x y : R, 0 < x -> 0 < y -> (mag y <= fexp1 (mag x) - 2)%Z -> generic_format beta fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 (x + y)
beta:radix

(2 <= beta)%Z
beta:radix
Hbeta:(2 <= beta)%Z
forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall choice1 choice2 : Z -> bool, round_round_plus_hyp fexp1 fexp2 -> forall x y : R, 0 < x -> 0 < y -> (mag y <= fexp1 (mag x) - 2)%Z -> generic_format beta fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 (x + y)
beta:radix

(2 <= beta)%Z
beta:radix
beta_val:Z
beta_prop:(2 <=? beta_val)%Z = true

(2 <= {| radix_val := beta_val; radix_prop := beta_prop |})%Z
now apply Zle_bool_imp_le.
beta:radix
Hbeta:(2 <= beta)%Z

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall choice1 choice2 : Z -> bool, round_round_plus_hyp fexp1 fexp2 -> forall x y : R, 0 < x -> 0 < y -> (mag y <= fexp1 (mag x) - 2)%Z -> generic_format beta fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 (x + y)
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x

round_round_eq fexp1 fexp2 choice1 choice2 (x + y)
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x

round_round_eq fexp1 fexp2 choice1 choice2 (x + y)
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x

round_round_eq fexp1 fexp2 choice1 choice2 (x + y)
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z

round_round_eq fexp1 fexp2 choice1 choice2 (x + y)
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z

bpow (-2) <= / 2 * / 2
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2
round_round_eq fexp1 fexp2 choice1 choice2 (x + y)
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z

bpow (-2) <= / 2 * / 2
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z

bpow (-2) <= / 4
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z

/ bpow 2 <= / 4
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z

4 <= bpow 2
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z

(2 * 2 <= beta * (beta * 1))%Z
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z

(2 * 2 <= beta * beta)%Z
now apply Zmult_le_compat; omega.
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2

round_round_eq fexp1 fexp2 choice1 choice2 (x + y)
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2
P2:(0 < 2)%Z

round_round_eq fexp1 fexp2 choice1 choice2 (x + y)
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2
P2:(0 < 2)%Z

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2
P2:(0 < 2)%Z

Valid_exp fexp1
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2
P2:(0 < 2)%Z
Valid_exp fexp2
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2
P2:(0 < 2)%Z
0 < x + y
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2
P2:(0 < 2)%Z
(fexp2 (mag (x + y)) <= fexp1 (mag (x + y)))%Z
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2
P2:(0 < 2)%Z
(fexp1 (mag (x + y)) <= mag (x + y))%Z
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2
P2:(0 < 2)%Z
x + y < midp fexp1 (x + y)
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2
P2:(0 < 2)%Z
(fexp2 (mag (x + y)) <= fexp1 (mag (x + y)) - 1)%Z -> x + y < midp fexp1 (x + y) - / 2 * ulp beta fexp2 (x + y)
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2
P2:(0 < 2)%Z

Valid_exp fexp1
exact Vfexp1.
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2
P2:(0 < 2)%Z

Valid_exp fexp2
exact Vfexp2.
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2
P2:(0 < 2)%Z

0 < x + y
lra.
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2
P2:(0 < 2)%Z

(fexp2 (mag (x + y)) <= fexp1 (mag (x + y)))%Z
now rewrite Lxy.
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2
P2:(0 < 2)%Z

(fexp1 (mag (x + y)) <= mag (x + y))%Z
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2
P2:(0 < 2)%Z

(fexp1 (mag x) <= mag x)%Z
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2
P2:(0 < 2)%Z

(fexp1 (mag x) < mag x)%Z
now apply mag_generic_gt; [|apply Rgt_not_eq|].
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2
P2:(0 < 2)%Z

x + y < midp fexp1 (x + y)
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2
P2:(0 < 2)%Z

x + y < round beta fexp1 Zfloor (x + y) + / 2 * ulp beta fexp1 (x + y)
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2
P2:(0 < 2)%Z

x + y + - round beta fexp1 Zfloor (x + y) < round beta fexp1 Zfloor (x + y) + / 2 * ulp beta fexp1 (x + y) + - round beta fexp1 Zfloor (x + y)
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2
P2:(0 < 2)%Z

bpow (fexp1 (mag x) - 2) <= round beta fexp1 Zfloor (x + y) + / 2 * ulp beta fexp1 (x + y) + - round beta fexp1 Zfloor (x + y)
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2
P2:(0 < 2)%Z

bpow (fexp1 (mag x) - 2) <= / 2 * ulp beta fexp1 (x + y)
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2
P2:(0 < 2)%Z

bpow (fexp1 (mag x) - 2) <= / 2 * bpow (cexp beta fexp1 (x + y))
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2
P2:(0 < 2)%Z

bpow (fexp1 (mag x) - 2) <= / 2 * bpow (fexp1 (mag x))
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2
P2:(0 < 2)%Z

bpow (fexp1 (mag x) - 2) * bpow (- fexp1 (mag x)) <= / 2 * bpow (fexp1 (mag x)) * bpow (- fexp1 (mag x))
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2
P2:(0 < 2)%Z

bpow (-2) <= / 2
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2
P2:(0 < 2)%Z

/ 2 * / 2 <= / 2
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2
P2:(0 < 2)%Z

/ 2 * / 2 <= / 2 * 1
apply Rmult_le_compat_l; lra.
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2
P2:(0 < 2)%Z

(fexp2 (mag (x + y)) <= fexp1 (mag (x + y)) - 1)%Z -> x + y < midp fexp1 (x + y) - / 2 * ulp beta fexp2 (x + y)
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2
P2:(0 < 2)%Z

(fexp2 (mag (x + y)) <= fexp1 (mag (x + y)) - 1)%Z -> x + y < midp fexp1 (x + y) - / 2 * bpow (cexp beta fexp2 (x + y))
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2
P2:(0 < 2)%Z

(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z -> x + y < midp fexp1 (x + y) - / 2 * bpow (fexp2 (mag x))
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2
P2:(0 < 2)%Z
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z

x + y < midp fexp1 (x + y) - / 2 * bpow (fexp2 (mag x))
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2
P2:(0 < 2)%Z
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z

(x + y) * bpow (- fexp1 (mag x)) < (midp fexp1 (x + y) - / 2 * bpow (fexp2 (mag x))) * bpow (- fexp1 (mag x))
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2
P2:(0 < 2)%Z
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z

(x + y) * bpow (- fexp1 (mag x)) * bpow (fexp1 (mag x)) < (midp fexp1 (x + y) - / 2 * bpow (fexp2 (mag x))) * bpow (- fexp1 (mag x)) * bpow (fexp1 (mag x))
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2
P2:(0 < 2)%Z
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z

x + y < midp fexp1 (x + y) - / 2 * bpow (fexp2 (mag x))
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2
P2:(0 < 2)%Z
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z

x + y + - round beta fexp1 Zfloor (x + y) < midp fexp1 (x + y) - / 2 * bpow (fexp2 (mag x)) + - round beta fexp1 Zfloor (x + y)
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2
P2:(0 < 2)%Z
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z

x + y - round beta fexp1 Zfloor (x + y) < / 2 * ulp beta fexp1 (x + y) - / 2 * bpow (fexp2 (mag x))
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2
P2:(0 < 2)%Z
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z

bpow (fexp1 (mag x) - 2) <= / 2 * ulp beta fexp1 (x + y) - / 2 * bpow (fexp2 (mag x))
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2
P2:(0 < 2)%Z
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z

bpow (fexp1 (mag x) - 2) * bpow (- fexp1 (mag x)) <= (/ 2 * ulp beta fexp1 (x + y) - / 2 * bpow (fexp2 (mag x))) * bpow (- fexp1 (mag x))
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2
P2:(0 < 2)%Z
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z

bpow (fexp1 (mag x) - 2) * bpow (- fexp1 (mag x)) <= (/ 2 * bpow (cexp beta fexp1 (x + y)) - / 2 * bpow (fexp2 (mag x))) * bpow (- fexp1 (mag x))
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2
P2:(0 < 2)%Z
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z

bpow (-2) <= / 2 - / 2 * bpow (fexp2 (mag x) - fexp1 (mag x))
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2
P2:(0 < 2)%Z
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z

/ 2 * / 2 <= / 2 - / 2 * bpow (fexp2 (mag x) - fexp1 (mag x))
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2
P2:(0 < 2)%Z
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z

/ 2 * / 2 <= / 2 * (1 - bpow (fexp2 (mag x) - fexp1 (mag x)))
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2
P2:(0 < 2)%Z
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z

/ 2 <= 1 - bpow (fexp2 (mag x) - fexp1 (mag x))
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2
P2:(0 < 2)%Z
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z

/ 2 - 1 <= - bpow (fexp2 (mag x) - fexp1 (mag x))
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2
P2:(0 < 2)%Z
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z

- / 2 <= - bpow (fexp2 (mag x) - fexp1 (mag x))
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2
P2:(0 < 2)%Z
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z

bpow (fexp2 (mag x) - fexp1 (mag x)) <= / 2
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2
P2:(0 < 2)%Z
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z

bpow (fexp2 (mag x) - fexp1 (mag x)) <= / 2
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2
P2:(0 < 2)%Z
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z

bpow (fexp2 (mag x) - fexp1 (mag x)) <= bpow (-1)
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2
P2:(0 < 2)%Z
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
bpow (-1) <= / 2
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2
P2:(0 < 2)%Z
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z

bpow (fexp2 (mag x) - fexp1 (mag x)) <= bpow (-1)
apply bpow_le; omega.
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2
P2:(0 < 2)%Z
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z

bpow (-1) <= / 2
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2
P2:(0 < 2)%Z
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z

/ IZR (beta * 1) <= / 2
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow2:bpow (-2) <= / 2 * / 2
P2:(0 < 2)%Z
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z

2 <= IZR (beta * 1)
apply IZR_le; omega. } Qed. (* round_round_plus_aux{0,1} together *)
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall choice1 choice2 : Z -> bool, round_round_plus_hyp fexp1 fexp2 -> forall x y : R, 0 < x -> 0 < y -> y <= x -> generic_format beta fexp1 x -> generic_format beta fexp1 y -> round_round_eq fexp1 fexp2 choice1 choice2 (x + y)
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall choice1 choice2 : Z -> bool, round_round_plus_hyp fexp1 fexp2 -> forall x y : R, 0 < x -> 0 < y -> y <= x -> generic_format beta fexp1 x -> generic_format beta fexp1 y -> round_round_eq fexp1 fexp2 choice1 choice2 (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y

round_round_eq fexp1 fexp2 choice1 choice2 (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hly:(mag y <= fexp1 (mag x) - 2)%Z

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hly:(fexp1 (mag x) - 2 < mag y)%Z
round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hly:(mag y <= fexp1 (mag x) - 2)%Z

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
now apply round_round_plus_aux1.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hly:(fexp1 (mag x) - 2 < mag y)%Z

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hly:(fexp1 (mag x) - 2 < mag y)%Z

round beta fexp1 (Znearest choice1) (x + y) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hly:(fexp1 (mag x) - 2 < mag y)%Z
Valid_rnd (Znearest choice2)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hly:(fexp1 (mag x) - 2 < mag y)%Z
generic_format beta fexp2 (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hly:(fexp1 (mag x) - 2 < mag y)%Z

round beta fexp1 (Znearest choice1) (x + y) = round beta fexp1 (Znearest choice1) (x + y)
reflexivity.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hly:(fexp1 (mag x) - 2 < mag y)%Z

Valid_rnd (Znearest choice2)
now apply valid_rnd_N.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hly:(fexp1 (mag x) - 2 < mag y)%Z

generic_format beta fexp2 (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hly:(fexp1 (mag x) - 2 < mag y)%Z
Hf1:(fexp1 (mag x) - 1 <= mag y)%Z

generic_format beta fexp2 (x + y)
now apply (round_round_plus_aux0 fexp1). Qed.
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall choice1 choice2 : Z -> bool, round_round_plus_hyp fexp1 fexp2 -> forall x y : R, 0 <= x -> 0 <= y -> generic_format beta fexp1 x -> generic_format beta fexp1 y -> round_round_eq fexp1 fexp2 choice1 choice2 (x + y)
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall choice1 choice2 : Z -> bool, round_round_plus_hyp fexp1 fexp2 -> forall x y : R, 0 <= x -> 0 <= y -> generic_format beta fexp1 x -> generic_format beta fexp1 y -> round_round_eq fexp1 fexp2 choice1 choice2 (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y

round_round_eq fexp1 fexp2 choice1 choice2 (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) y) = round beta fexp1 (Znearest choice1) y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0

round beta fexp1 (Znearest choice1) y = round beta fexp1 (Znearest choice1) y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0
Valid_rnd (Znearest choice2)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0
generic_format beta fexp2 y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0

round beta fexp1 (Znearest choice1) y = round beta fexp1 (Znearest choice1) y
reflexivity.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0

Valid_rnd (Znearest choice2)
now apply valid_rnd_N.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0

generic_format beta fexp2 y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0

y <> 0 -> (fexp2 (mag y) <= fexp1 (mag y))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0
generic_format beta fexp1 y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0

generic_format beta fexp1 y
exact Fy.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Nzy:y <> 0
round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0

round beta fexp1 (Znearest choice1) x = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0
Valid_rnd (Znearest choice2)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0
generic_format beta fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0

round beta fexp1 (Znearest choice1) x = round beta fexp1 (Znearest choice1) x
reflexivity.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0

Valid_rnd (Znearest choice2)
now apply valid_rnd_N.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0

generic_format beta fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0

x <> 0 -> (fexp2 (mag x) <= fexp1 (mag x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0
generic_format beta fexp1 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0

generic_format beta fexp1 x
exact Fx.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Nzy:y <> 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Nzy:y <> 0
Px:0 < x

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Nzy:y <> 0
Px:0 < x
Py:0 < y

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Nzy:y <> 0
Px:0 < x
Py:0 < y
H:x < y

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Nzy:y <> 0
Px:0 < x
Py:0 < y
H:y <= x
round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Nzy:y <> 0
Px:0 < x
Py:0 < y
H:x < y

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Nzy:y <> 0
Px:0 < x
Py:0 < y
H:x <= y

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Nzy:y <> 0
Px:0 < x
Py:0 < y
H:x <= y

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (y + x)) = round beta fexp1 (Znearest choice1) (y + x)
now apply round_round_plus_aux2.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Nzy:y <> 0
Px:0 < x
Py:0 < y
H:y <= x

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
now apply round_round_plus_aux2. Qed.
beta:radix

forall (fexp1 fexp2 : Z -> Z) (x y : R), (fexp2 (mag (x - y)) <= fexp1 (mag x))%Z -> (fexp2 (mag (x - y)) <= fexp1 (mag y))%Z -> generic_format beta fexp1 x -> generic_format beta fexp1 y -> generic_format beta fexp2 (x - y)
beta:radix

forall (fexp1 fexp2 : Z -> Z) (x y : R), (fexp2 (mag (x - y)) <= fexp1 (mag x))%Z -> (fexp2 (mag (x - y)) <= fexp1 (mag y))%Z -> generic_format beta fexp1 x -> generic_format beta fexp1 y -> generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
x, y:R

(fexp2 (mag (x - y)) <= fexp1 (mag x))%Z -> (fexp2 (mag (x - y)) <= fexp1 (mag y))%Z -> generic_format beta fexp1 x -> generic_format beta fexp1 y -> generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
x, y:R

(fexp2 (mag (x + - y)) <= fexp1 (mag x))%Z -> (fexp2 (mag (x + - y)) <= fexp1 (mag y))%Z -> generic_format beta fexp1 x -> generic_format beta fexp1 y -> generic_format beta fexp2 (x + - y)
beta:radix
fexp1, fexp2:Z -> Z
x, y:R
Hlnx:(fexp2 (mag (x + - y)) <= fexp1 (mag x))%Z
Hlny:(fexp2 (mag (x + - y)) <= fexp1 (mag y))%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y

generic_format beta fexp2 (x + - y)
beta:radix
fexp1, fexp2:Z -> Z
x, y:R
Hlnx:(fexp2 (mag (x + - y)) <= fexp1 (mag x))%Z
Hlny:(fexp2 (mag (x + - y)) <= fexp1 (mag (- y)))%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y

generic_format beta fexp2 (x + - y)
beta:radix
fexp1, fexp2:Z -> Z
x, y:R
Hlnx:(fexp2 (mag (x + - y)) <= fexp1 (mag x))%Z
Hlny:(fexp2 (mag (x + - y)) <= fexp1 (mag (- y)))%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 (- y)

generic_format beta fexp2 (x + - y)
now apply (round_round_plus_aux0_aux fexp1). Qed. (* fexp1 (mag x) - 1 <= mag y : * substraction is exact in the largest precision (fexp2). *)
beta:radix

forall fexp1 fexp2 : Z -> Z, round_round_plus_hyp fexp1 fexp2 -> forall x y : R, 0 < y -> y < x -> (fexp1 (mag x) - 1 <= mag y)%Z -> generic_format beta fexp1 x -> generic_format beta fexp1 y -> generic_format beta fexp2 (x - y)
beta:radix

forall fexp1 fexp2 : Z -> Z, round_round_plus_hyp fexp1 fexp2 -> forall x y : R, 0 < y -> y < x -> (fexp1 (mag x) - 1 <= mag y)%Z -> generic_format beta fexp1 x -> generic_format beta fexp1 y -> generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y

generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x

generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x

generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z

generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hlt:(mag x - 2 < mag y)%Z

generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hge:(mag y <= mag x - 2)%Z
generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hlt:(mag x - 2 < mag y)%Z

generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hlt:(mag x - 2 < mag y)%Z
Hor:mag y = mag x \/ mag y = (mag x - 1)%Z

generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hlt:(mag x - 2 < mag y)%Z
Heq:mag y = mag x

generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hlt:(mag x - 2 < mag y)%Z
Heqm1:mag y = (mag x - 1)%Z
generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hlt:(mag x - 2 < mag y)%Z
Heq:mag y = mag x

generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hlt:(mag x - 2 < mag y)%Z
Heq:mag y = mag x

(fexp2 (mag (x - y)) <= fexp1 (mag x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hlt:(mag x - 2 < mag y)%Z
Heq:mag y = mag x
(fexp2 (mag (x - y)) <= fexp1 (mag y))%Z
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hlt:(mag x - 2 < mag y)%Z
Heq:mag y = mag x

(fexp2 (mag (x - y)) <= fexp1 (mag x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hlt:(mag x - 2 < mag y)%Z
Heq:mag y = mag x

(mag (x - y) - 1 <= mag x)%Z
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hlt:(mag x - 2 < mag y)%Z
Heq:mag y = mag x

(mag (x - y) <= mag x)%Z
now apply mag_minus.
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hlt:(mag x - 2 < mag y)%Z
Heq:mag y = mag x

(fexp2 (mag (x - y)) <= fexp1 (mag y))%Z
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hlt:(mag x - 2 < mag y)%Z
Heq:mag y = mag x

(fexp2 (mag (x - y)) <= fexp1 (mag x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hlt:(mag x - 2 < mag y)%Z
Heq:mag y = mag x

(mag (x - y) - 1 <= mag x)%Z
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hlt:(mag x - 2 < mag y)%Z
Heq:mag y = mag x

(mag (x - y) <= mag x)%Z
now apply mag_minus.
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hlt:(mag x - 2 < mag y)%Z
Heqm1:mag y = (mag x - 1)%Z

generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hlt:(mag x - 2 < mag y)%Z
Heqm1:mag y = (mag x - 1)%Z

(fexp2 (mag (x - y)) <= fexp1 (mag x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hlt:(mag x - 2 < mag y)%Z
Heqm1:mag y = (mag x - 1)%Z
(fexp2 (mag (x - y)) <= fexp1 (mag y))%Z
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hlt:(mag x - 2 < mag y)%Z
Heqm1:mag y = (mag x - 1)%Z

(fexp2 (mag (x - y)) <= fexp1 (mag x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hlt:(mag x - 2 < mag y)%Z
Heqm1:mag y = (mag x - 1)%Z

(mag (x - y) - 1 <= mag x)%Z
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hlt:(mag x - 2 < mag y)%Z
Heqm1:mag y = (mag x - 1)%Z

(mag (x - y) <= mag x)%Z
now apply mag_minus.
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hlt:(mag x - 2 < mag y)%Z
Heqm1:mag y = (mag x - 1)%Z

(fexp2 (mag (x - y)) <= fexp1 (mag y))%Z
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hlt:(mag x - 2 < mag y)%Z
Heqm1:mag y = (mag x - 1)%Z

(fexp2 (mag (x - y)) <= fexp1 (mag x - 1))%Z
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hlt:(mag x - 2 < mag y)%Z
Heqm1:mag y = (mag x - 1)%Z

(mag (x - y) - 1 <= mag x - 1)%Z
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hlt:(mag x - 2 < mag y)%Z
Heqm1:mag y = (mag x - 1)%Z

(mag (x - y) <= mag x)%Z
now apply mag_minus.
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hge:(mag y <= mag x - 2)%Z

generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hge:(mag y <= mag x - 2)%Z
Lxmy:mag (x - y) = mag x

generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hge:(mag y <= mag x - 2)%Z
Lxmy:mag (x - y) = (mag x - 1)%Z
generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hge:(mag y <= mag x - 2)%Z
Lxmy:mag (x - y) = mag x

generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hge:(mag y <= mag x - 2)%Z
Lxmy:mag (x - y) = mag x

(fexp2 (mag (x - y)) <= fexp1 (mag x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hge:(mag y <= mag x - 2)%Z
Lxmy:mag (x - y) = mag x
(fexp2 (mag (x - y)) <= fexp1 (mag y))%Z
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hge:(mag y <= mag x - 2)%Z
Lxmy:mag (x - y) = mag x

(fexp2 (mag (x - y)) <= fexp1 (mag x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hge:(mag y <= mag x - 2)%Z
Lxmy:mag (x - y) = mag x

(mag (x - y) - 1 <= mag x)%Z
omega.
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hge:(mag y <= mag x - 2)%Z
Lxmy:mag (x - y) = mag x

(fexp2 (mag (x - y)) <= fexp1 (mag y))%Z
now rewrite Lxmy; apply Hexp3.
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hge:(mag y <= mag x - 2)%Z
Lxmy:mag (x - y) = (mag x - 1)%Z

generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hge:(mag y <= mag x - 2)%Z
Lxmy:mag (x - y) = (mag x - 1)%Z

(fexp2 (mag x - 1) <= fexp1 (mag x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hge:(mag y <= mag x - 2)%Z
Lxmy:mag (x - y) = (mag x - 1)%Z
(fexp2 (mag x - 1) <= fexp1 (mag y))%Z
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hge:(mag y <= mag x - 2)%Z
Lxmy:mag (x - y) = (mag x - 1)%Z

(fexp2 (mag x - 1) <= fexp1 (mag x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hge:(mag y <= mag x - 2)%Z
Lxmy:mag (x - y) = (mag x - 1)%Z

(fexp1 (mag x - 1 + 1) - 1 <= mag x)%Z
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hge:(mag y <= mag x - 2)%Z
Lxmy:mag (x - y) = (mag x - 1)%Z

(fexp1 (mag x) - 1 <= mag x)%Z
now apply Z.le_trans with (mag y).
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hge:(mag y <= mag x - 2)%Z
Lxmy:mag (x - y) = (mag x - 1)%Z

(fexp2 (mag x - 1) <= fexp1 (mag y))%Z
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hge:(mag y <= mag x - 2)%Z
Lxmy:mag (x - y) = (mag x - 1)%Z

(fexp1 (mag x - 1 + 1) - 1 <= mag y)%Z
now replace (_ + _)%Z with (mag x : Z); [|ring]. Qed. (* mag y <= fexp1 (mag x) - 2, * fexp1 (mag (x - y)) - 1 <= mag y : * substraction is exact in the largest precision (fexp2). *)
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> round_round_plus_hyp fexp1 fexp2 -> forall x y : R, 0 < y -> y < x -> (mag y <= fexp1 (mag x) - 2)%Z -> (fexp1 (mag (x - y)) - 1 <= mag y)%Z -> generic_format beta fexp1 x -> generic_format beta fexp1 y -> generic_format beta fexp2 (x - y)
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> round_round_plus_hyp fexp1 fexp2 -> forall x y : R, 0 < y -> y < x -> (mag y <= fexp1 (mag x) - 2)%Z -> (fexp1 (mag (x - y)) - 1 <= mag y)%Z -> generic_format beta fexp1 x -> generic_format beta fexp1 y -> generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y < x
Hln:(mag y <= fexp1 (mag x) - 2)%Z
Hln':(fexp1 (mag (x - y)) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y

generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y < x
Hln:(mag y <= fexp1 (mag x) - 2)%Z
Hln':(fexp1 (mag (x - y)) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x

generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(mag y <= fexp1 (mag x) - 2)%Z
Hln':(fexp1 (mag (x - y)) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x

generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(mag y <= fexp1 (mag x) - 2)%Z
Hln':(fexp1 (mag (x - y)) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z

generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(mag y <= fexp1 (mag x) - 2)%Z
Hln':(fexp1 (mag (x - y)) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hfx:(fexp1 (mag x) < mag x)%Z

generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(mag y <= fexp1 (mag x) - 2)%Z
Hln':(fexp1 (mag (x - y)) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z

generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(mag y <= fexp1 (mag x) - 2)%Z
Hln':(fexp1 (mag (x - y)) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z

(fexp2 (mag (x - y)) <= fexp1 (mag x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(mag y <= fexp1 (mag x) - 2)%Z
Hln':(fexp1 (mag (x - y)) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
(fexp2 (mag (x - y)) <= fexp1 (mag y))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(mag y <= fexp1 (mag x) - 2)%Z
Hln':(fexp1 (mag (x - y)) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z

(fexp2 (mag (x - y)) <= fexp1 (mag x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(mag y <= fexp1 (mag x) - 2)%Z
Hln':(fexp1 (mag (x - y)) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z

(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(mag y <= fexp1 (mag x) - 2)%Z
Hln':(fexp1 (mag (x - y)) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
(fexp1 (mag (x - y)) <= fexp1 (mag x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(mag y <= fexp1 (mag x) - 2)%Z
Hln':(fexp1 (mag (x - y)) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z

(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)))%Z
apply Hexp4; omega.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(mag y <= fexp1 (mag x) - 2)%Z
Hln':(fexp1 (mag (x - y)) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z

(fexp1 (mag (x - y)) <= fexp1 (mag x))%Z
omega.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(mag y <= fexp1 (mag x) - 2)%Z
Hln':(fexp1 (mag (x - y)) - 1 <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z

(fexp2 (mag (x - y)) <= fexp1 (mag y))%Z
now apply Hexp3. Qed.
beta:radix

forall fexp : Z -> Z, Valid_exp fexp -> forall x y : R, 0 < y -> y < x -> (mag y <= fexp (mag x) - 1)%Z -> generic_format beta fexp x -> generic_format beta fexp y -> round beta fexp Zceil (x - y) - (x - y) <= y
beta:radix

forall fexp : Z -> Z, Valid_exp fexp -> forall x y : R, 0 < y -> y < x -> (mag y <= fexp (mag x) - 1)%Z -> generic_format beta fexp x -> generic_format beta fexp y -> round beta fexp Zceil (x - y) - (x - y) <= y
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fx:generic_format beta fexp x
Fy:generic_format beta fexp y

round beta fexp Zceil (x - y) - (x - y) <= y
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fx:generic_format beta fexp x
Fy:generic_format beta fexp y
Px:0 < x

round beta fexp Zceil (x - y) - (x - y) <= y
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x

generic_format beta fexp x -> round beta fexp Zceil (x - y) - (x - y) <= y
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x

x = IZR (Ztrunc (x * bpow (- fexp (mag x)))) * bpow (fexp (mag x)) -> round beta fexp Zceil (x - y) - (x - y) <= y
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z

x = IZR mx * bpow (fexp (mag x)) -> round beta fexp Zceil (x - y) - (x - y) <= y
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))

round beta fexp Zceil (x - y) - (x - y) <= y
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z

round beta fexp Zceil (x - y) - (x - y) <= y
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z

round beta fexp Zceil (x - y) - (x - y) <= y
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:bpow (mag x - 1) < x

round beta fexp Zceil (x - y) - (x - y) <= y
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:x <= bpow (mag x - 1)
round beta fexp Zceil (x - y) - (x - y) <= y
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:bpow (mag x - 1) < x

round beta fexp Zceil (x - y) - (x - y) <= y
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:bpow (mag x - 1) < x
Lxy:mag (x - y) = mag x

round beta fexp Zceil (x - y) - (x - y) <= y
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:bpow (mag x - 1) < x
Lxy:mag (x - y) = mag x

round beta fexp Zceil (x - y) = x
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:bpow (mag x - 1) < x
Lxy:mag (x - y) = mag x
Rxy:round beta fexp Zceil (x - y) = x
round beta fexp Zceil (x - y) - (x - y) <= y
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:bpow (mag x - 1) < x
Lxy:mag (x - y) = mag x

round beta fexp Zceil (x - y) = x
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:bpow (mag x - 1) < x
Lxy:mag (x - y) = mag x

IZR (Zceil ((x - y) * bpow (- fexp (mag (x - y))))) * bpow (fexp (mag (x - y))) = x
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:bpow (mag x - 1) < x
Lxy:mag (x - y) = mag x

IZR (Zceil ((x - y) * bpow (- fexp (mag x)))) * bpow (fexp (mag x)) = x
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:bpow (mag x - 1) < x
Lxy:mag (x - y) = mag x

IZR (Zceil ((x - y) * bpow (- fexp (mag x)))) * bpow (fexp (mag x)) = IZR mx * bpow (fexp (mag x))
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:bpow (mag x - 1) < x
Lxy:mag (x - y) = mag x

IZR (Zceil ((x - y) * bpow (- fexp (mag x)))) = IZR mx
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:bpow (mag x - 1) < x
Lxy:mag (x - y) = mag x

Zceil ((x - y) * bpow (- fexp (mag x))) = mx
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:bpow (mag x - 1) < x
Lxy:mag (x - y) = mag x

Zceil ((IZR mx * bpow (fexp (mag x)) - y) * bpow (- fexp (mag x))) = mx
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:bpow (mag x - 1) < x
Lxy:mag (x - y) = mag x

Zceil (IZR mx * bpow (fexp (mag x)) * bpow (- fexp (mag x)) - y * bpow (- fexp (mag x))) = mx
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:bpow (mag x - 1) < x
Lxy:mag (x - y) = mag x

Zceil (IZR mx - y * bpow (- fexp (mag x))) = mx
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:bpow (mag x - 1) < x
Lxy:mag (x - y) = mag x

IZR (mx - 1) < IZR mx - y * bpow (- fexp (mag x)) <= IZR mx
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:bpow (mag x - 1) < x
Lxy:mag (x - y) = mag x

IZR (mx - 1) < IZR mx - y * bpow (- fexp (mag x))
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:bpow (mag x - 1) < x
Lxy:mag (x - y) = mag x
IZR mx - y * bpow (- fexp (mag x)) <= IZR mx
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:bpow (mag x - 1) < x
Lxy:mag (x - y) = mag x

IZR (mx - 1) < IZR mx - y * bpow (- fexp (mag x))
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:bpow (mag x - 1) < x
Lxy:mag (x - y) = mag x

IZR mx + IZR (- (1)) < IZR mx - y * bpow (- fexp (mag x))
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:bpow (mag x - 1) < x
Lxy:mag (x - y) = mag x

IZR (- (1)) < - (y * bpow (- fexp (mag x)))
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:bpow (mag x - 1) < x
Lxy:mag (x - y) = mag x

y * bpow (- fexp (mag x)) < IPR 1
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:bpow (mag x - 1) < x
Lxy:mag (x - y) = mag x

y * bpow (- fexp (mag x)) * bpow (fexp (mag x)) < IPR 1 * bpow (fexp (mag x))
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:bpow (mag x - 1) < x
Lxy:mag (x - y) = mag x

y < bpow (fexp (mag x))
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:bpow (mag x - 1) < x
Lxy:mag (x - y) = mag x

y < bpow (mag y)
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:bpow (mag x - 1) < x
Lxy:mag (x - y) = mag x
bpow (mag y) <= bpow (fexp (mag x))
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:bpow (mag x - 1) < x
Lxy:mag (x - y) = mag x

y < bpow (mag y)
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:bpow (mag x - 1) < x
Lxy:mag (x - y) = mag x

Rabs y < bpow (mag y)
apply bpow_mag_gt.
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:bpow (mag x - 1) < x
Lxy:mag (x - y) = mag x

bpow (mag y) <= bpow (fexp (mag x))
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:bpow (mag x - 1) < x
Lxy:mag (x - y) = mag x

(mag y <= fexp (mag x))%Z
omega.
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:bpow (mag x - 1) < x
Lxy:mag (x - y) = mag x

IZR mx - y * bpow (- fexp (mag x)) <= IZR mx
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:bpow (mag x - 1) < x
Lxy:mag (x - y) = mag x

IZR mx - y * bpow (- fexp (mag x)) <= IZR mx + 0
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:bpow (mag x - 1) < x
Lxy:mag (x - y) = mag x

- (y * bpow (- fexp (mag x))) <= 0
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:bpow (mag x - 1) < x
Lxy:mag (x - y) = mag x

0 <= y * bpow (- fexp (mag x))
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:bpow (mag x - 1) < x
Lxy:mag (x - y) = mag x

y * 0 <= y * bpow (- fexp (mag x))
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:bpow (mag x - 1) < x
Lxy:mag (x - y) = mag x

0 <= bpow (- fexp (mag x))
now apply bpow_ge_0.
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:bpow (mag x - 1) < x
Lxy:mag (x - y) = mag x
Rxy:round beta fexp Zceil (x - y) = x

round beta fexp Zceil (x - y) - (x - y) <= y
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:bpow (mag x - 1) < x
Lxy:mag (x - y) = mag x
Rxy:round beta fexp Zceil (x - y) = x

y <= y
apply Rle_refl.
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:x <= bpow (mag x - 1)

round beta fexp Zceil (x - y) - (x - y) <= y
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:x <= bpow (mag x - 1)

x = bpow (mag x - 1)
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:x <= bpow (mag x - 1)
Xpow:x = bpow (mag x - 1)
round beta fexp Zceil (x - y) - (x - y) <= y
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:x <= bpow (mag x - 1)

x = bpow (mag x - 1)
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:x <= bpow (mag x - 1)

bpow (mag x - 1) <= x
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
ex:Z
Hex:x <> 0 -> bpow (ex - 1) <= Rabs x < bpow ex
Hly:(mag y <= fexp (Build_mag_prop beta x ex Hex) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (Build_mag_prop beta x ex Hex))):Z
Fx:x = IZR mx * bpow (fexp (Build_mag_prop beta x ex Hex))
Hfx:(fexp (Build_mag_prop beta x ex Hex) < Build_mag_prop beta x ex Hex)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:x <= bpow (Build_mag_prop beta x ex Hex - 1)

bpow (ex - 1) <= x
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
ex:Z
Hex:x <> 0 -> bpow (ex - 1) <= Rabs x < bpow ex
Hly:(mag y <= fexp (Build_mag_prop beta x ex Hex) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (Build_mag_prop beta x ex Hex))):Z
Fx:x = IZR mx * bpow (fexp (Build_mag_prop beta x ex Hex))
Hfx:(fexp (Build_mag_prop beta x ex Hex) < Build_mag_prop beta x ex Hex)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:x <= bpow (Build_mag_prop beta x ex Hex - 1)

bpow (ex - 1) <= Rabs x
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
ex:Z
Hex:x <> 0 -> bpow (ex - 1) <= Rabs x < bpow ex
Hly:(mag y <= fexp (Build_mag_prop beta x ex Hex) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (Build_mag_prop beta x ex Hex))):Z
Fx:x = IZR mx * bpow (fexp (Build_mag_prop beta x ex Hex))
Hfx:(fexp (Build_mag_prop beta x ex Hex) < Build_mag_prop beta x ex Hex)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:x <= bpow (Build_mag_prop beta x ex Hex - 1)

x <> 0
now apply Rgt_not_eq.
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:x <= bpow (mag x - 1)
Xpow:x = bpow (mag x - 1)

round beta fexp Zceil (x - y) - (x - y) <= y
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:x <= bpow (mag x - 1)
Xpow:x = bpow (mag x - 1)

mag (x - y) = (mag x - 1)%Z
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:x <= bpow (mag x - 1)
Xpow:x = bpow (mag x - 1)
Lxy:mag (x - y) = (mag x - 1)%Z
round beta fexp Zceil (x - y) - (x - y) <= y
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:x <= bpow (mag x - 1)
Xpow:x = bpow (mag x - 1)

mag (x - y) = (mag x - 1)%Z
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:x <= bpow (mag x - 1)
Xpow:x = bpow (mag x - 1)

(mag (x - y) <= mag x - 1)%Z
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:x <= bpow (mag x - 1)
Xpow:x = bpow (mag x - 1)
(mag x - 1 <= mag (x - y))%Z
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:x <= bpow (mag x - 1)
Xpow:x = bpow (mag x - 1)

(mag (x - y) <= mag x - 1)%Z
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:x <= bpow (mag x - 1)
Xpow:x = bpow (mag x - 1)

x - y <> 0
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:x <= bpow (mag x - 1)
Xpow:x = bpow (mag x - 1)
Rabs (x - y) < bpow (mag x - 1)
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:x <= bpow (mag x - 1)
Xpow:x = bpow (mag x - 1)

x - y <> 0
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:x <= bpow (mag x - 1)
Xpow:x = bpow (mag x - 1)

x <> y
now intro Hx'; rewrite Hx' in Hxy; apply (Rlt_irrefl y).
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:x <= bpow (mag x - 1)
Xpow:x = bpow (mag x - 1)

Rabs (x - y) < bpow (mag x - 1)
rewrite Rabs_right; lra.
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:x <= bpow (mag x - 1)
Xpow:x = bpow (mag x - 1)

(mag x - 1 <= mag (x - y))%Z
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:x <= bpow (mag x - 1)
Xpow:x = bpow (mag x - 1)

(mag y <= mag x - 2)%Z
omega.
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:x <= bpow (mag x - 1)
Xpow:x = bpow (mag x - 1)
Lxy:mag (x - y) = (mag x - 1)%Z

round beta fexp Zceil (x - y) - (x - y) <= y
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:x <= bpow (mag x - 1)
Xpow:x = bpow (mag x - 1)
Lxy:mag (x - y) = (mag x - 1)%Z
Hfx1:(fexp (mag x - 1) < mag x - 1)%Z

round beta fexp Zceil (x - y) - (x - y) <= y
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:x <= bpow (mag x - 1)
Xpow:x = bpow (mag x - 1)
Lxy:mag (x - y) = (mag x - 1)%Z
Hfx1:(fexp (mag x - 1) < mag x - 1)%Z

round beta fexp Zceil (x - y) <= x
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:x <= bpow (mag x - 1)
Xpow:x = bpow (mag x - 1)
Lxy:mag (x - y) = (mag x - 1)%Z
Hfx1:(fexp (mag x - 1) < mag x - 1)%Z
Rxy:round beta fexp Zceil (x - y) <= x
round beta fexp Zceil (x - y) - (x - y) <= y
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:x <= bpow (mag x - 1)
Xpow:x = bpow (mag x - 1)
Lxy:mag (x - y) = (mag x - 1)%Z
Hfx1:(fexp (mag x - 1) < mag x - 1)%Z

round beta fexp Zceil (x - y) <= x
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:x <= bpow (mag x - 1)
Xpow:x = bpow (mag x - 1)
Lxy:mag (x - y) = (mag x - 1)%Z
Hfx1:(fexp (mag x - 1) < mag x - 1)%Z

round beta fexp Zceil (x - y) <= bpow (mag x - 1)
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:x <= bpow (mag x - 1)
Xpow:x = bpow (mag x - 1)
Lxy:mag (x - y) = (mag x - 1)%Z
Hfx1:(fexp (mag x - 1) < mag x - 1)%Z

IZR (Zceil ((x - y) * bpow (- fexp (mag (x - y))))) * bpow (fexp (mag (x - y))) <= bpow (mag x - 1)
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:x <= bpow (mag x - 1)
Xpow:x = bpow (mag x - 1)
Lxy:mag (x - y) = (mag x - 1)%Z
Hfx1:(fexp (mag x - 1) < mag x - 1)%Z

IZR (Zceil ((x - y) * bpow (- fexp (mag x - 1)%Z))) * bpow (fexp (mag x - 1)%Z) <= bpow (mag x - 1)
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:x <= bpow (mag x - 1)
Xpow:x = bpow (mag x - 1)
Lxy:mag (x - y) = (mag x - 1)%Z
Hfx1:(fexp (mag x - 1) < mag x - 1)%Z

IZR (Zceil ((x - y) * bpow (- fexp (mag x - 1)%Z))) * bpow (fexp (mag x - 1)%Z) * bpow (- fexp (mag x - 1)%Z) <= bpow (mag x - 1) * bpow (- fexp (mag x - 1)%Z)
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:x <= bpow (mag x - 1)
Xpow:x = bpow (mag x - 1)
Lxy:mag (x - y) = (mag x - 1)%Z
Hfx1:(fexp (mag x - 1) < mag x - 1)%Z

IZR (Zceil ((x - y) * bpow (- fexp (mag x - 1)%Z))) <= bpow (mag x - fexp (mag x - 1)%Z - 1)
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:x <= bpow (mag x - 1)
Xpow:x = bpow (mag x - 1)
Lxy:mag (x - y) = (mag x - 1)%Z
Hfx1:(fexp (mag x - 1) < mag x - 1)%Z

IZR (Zceil ((x - y) * bpow (- fexp (mag x - 1)%Z))) <= IZR (beta ^ (mag x - fexp (mag x - 1)%Z - 1))
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:x <= bpow (mag x - 1)
Xpow:x = bpow (mag x - 1)
Lxy:mag (x - y) = (mag x - 1)%Z
Hfx1:(fexp (mag x - 1) < mag x - 1)%Z

(Zceil ((x - y) * bpow (- fexp (mag x - 1))) <= beta ^ (mag x - fexp (mag x - 1) - 1))%Z
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:x <= bpow (mag x - 1)
Xpow:x = bpow (mag x - 1)
Lxy:mag (x - y) = (mag x - 1)%Z
Hfx1:(fexp (mag x - 1) < mag x - 1)%Z

(x - y) * bpow (- fexp (mag x - 1)%Z) <= IZR (beta ^ (mag x - fexp (mag x - 1)%Z - 1))
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:x <= bpow (mag x - 1)
Xpow:x = bpow (mag x - 1)
Lxy:mag (x - y) = (mag x - 1)%Z
Hfx1:(fexp (mag x - 1) < mag x - 1)%Z

(x - y) * bpow (- fexp (mag x - 1)%Z) <= bpow (mag x - fexp (mag x - 1)%Z - 1)
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:x <= bpow (mag x - 1)
Xpow:x = bpow (mag x - 1)
Lxy:mag (x - y) = (mag x - 1)%Z
Hfx1:(fexp (mag x - 1) < mag x - 1)%Z

(bpow (mag x - 1) - y) * bpow (- fexp (mag x - 1)%Z) <= bpow (mag x - fexp (mag x - 1)%Z - 1)
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:x <= bpow (mag x - 1)
Xpow:x = bpow (mag x - 1)
Lxy:mag (x - y) = (mag x - 1)%Z
Hfx1:(fexp (mag x - 1) < mag x - 1)%Z

bpow (mag x - 1) * bpow (- fexp (mag x - 1)%Z) - y * bpow (- fexp (mag x - 1)%Z) <= bpow (mag x - fexp (mag x - 1)%Z - 1)
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:x <= bpow (mag x - 1)
Xpow:x = bpow (mag x - 1)
Lxy:mag (x - y) = (mag x - 1)%Z
Hfx1:(fexp (mag x - 1) < mag x - 1)%Z

bpow (mag x - fexp (mag x - 1)%Z - 1) - y * bpow (- fexp (mag x - 1)%Z) <= bpow (mag x - fexp (mag x - 1)%Z - 1)
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:x <= bpow (mag x - 1)
Xpow:x = bpow (mag x - 1)
Lxy:mag (x - y) = (mag x - 1)%Z
Hfx1:(fexp (mag x - 1) < mag x - 1)%Z

bpow (mag x - fexp (mag x - 1)%Z - 1) - y * bpow (- fexp (mag x - 1)%Z) <= bpow (mag x - fexp (mag x - 1)%Z - 1) + 0
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:x <= bpow (mag x - 1)
Xpow:x = bpow (mag x - 1)
Lxy:mag (x - y) = (mag x - 1)%Z
Hfx1:(fexp (mag x - 1) < mag x - 1)%Z

- (y * bpow (- fexp (mag x - 1)%Z)) <= 0
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:x <= bpow (mag x - 1)
Xpow:x = bpow (mag x - 1)
Lxy:mag (x - y) = (mag x - 1)%Z
Hfx1:(fexp (mag x - 1) < mag x - 1)%Z

0 <= y * bpow (- fexp (mag x - 1)%Z)
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:x <= bpow (mag x - 1)
Xpow:x = bpow (mag x - 1)
Lxy:mag (x - y) = (mag x - 1)%Z
Hfx1:(fexp (mag x - 1) < mag x - 1)%Z

y * 0 <= y * bpow (- fexp (mag x - 1)%Z)
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:x <= bpow (mag x - 1)
Xpow:x = bpow (mag x - 1)
Lxy:mag (x - y) = (mag x - 1)%Z
Hfx1:(fexp (mag x - 1) < mag x - 1)%Z

0 <= bpow (- fexp (mag x - 1)%Z)
now apply bpow_ge_0.
beta:radix
fexp:Z -> Z
Vfexp:Valid_exp fexp
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp (mag x) - 1)%Z
Fy:generic_format beta fexp y
Px:0 < x
mx:=Ztrunc (x * bpow (- fexp (mag x))):Z
Fx:x = IZR mx * bpow (fexp (mag x))
Hfx:(fexp (mag x) < mag x)%Z
Hfy:(fexp (mag y) < mag y)%Z
Hx:x <= bpow (mag x - 1)
Xpow:x = bpow (mag x - 1)
Lxy:mag (x - y) = (mag x - 1)%Z
Hfx1:(fexp (mag x - 1) < mag x - 1)%Z
Rxy:round beta fexp Zceil (x - y) <= x

round beta fexp Zceil (x - y) - (x - y) <= y
lra. Qed. (* mag y <= fexp1 (mag x) - 2 : * mag y <= fexp1 (mag (x - y)) - 2 : * round_round_gt_mid applies. *)
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall choice1 choice2 : Z -> bool, round_round_plus_hyp fexp1 fexp2 -> forall x y : R, 0 < y -> y < x -> (mag y <= fexp1 (mag x) - 2)%Z -> (mag y <= fexp1 (mag (x - y)) - 2)%Z -> generic_format beta fexp1 x -> generic_format beta fexp1 y -> round_round_eq fexp1 fexp2 choice1 choice2 (x - y)
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall choice1 choice2 : Z -> bool, round_round_plus_hyp fexp1 fexp2 -> forall x y : R, 0 < y -> y < x -> (mag y <= fexp1 (mag x) - 2)%Z -> (mag y <= fexp1 (mag (x - y)) - 2)%Z -> generic_format beta fexp1 x -> generic_format beta fexp1 y -> round_round_eq fexp1 fexp2 choice1 choice2 (x - y)
beta:radix

(2 <= beta)%Z
beta:radix
Hbeta:(2 <= beta)%Z
forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall choice1 choice2 : Z -> bool, round_round_plus_hyp fexp1 fexp2 -> forall x y : R, 0 < y -> y < x -> (mag y <= fexp1 (mag x) - 2)%Z -> (mag y <= fexp1 (mag (x - y)) - 2)%Z -> generic_format beta fexp1 x -> generic_format beta fexp1 y -> round_round_eq fexp1 fexp2 choice1 choice2 (x - y)
beta:radix

(2 <= beta)%Z
beta:radix
beta_val:Z
beta_prop:(2 <=? beta_val)%Z = true

(2 <= {| radix_val := beta_val; radix_prop := beta_prop |})%Z
now apply Zle_bool_imp_le.
beta:radix
Hbeta:(2 <= beta)%Z

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall choice1 choice2 : Z -> bool, round_round_plus_hyp fexp1 fexp2 -> forall x y : R, 0 < y -> y < x -> (mag y <= fexp1 (mag x) - 2)%Z -> (mag y <= fexp1 (mag (x - y)) - 2)%Z -> generic_format beta fexp1 x -> generic_format beta fexp1 y -> round_round_eq fexp1 fexp2 choice1 choice2 (x - y)
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y

round_round_eq fexp1 fexp2 choice1 choice2 (x - y)
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x

round_round_eq fexp1 fexp2 choice1 choice2 (x - y)
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x

round_round_eq fexp1 fexp2 choice1 choice2 (x - y)
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z

round_round_eq fexp1 fexp2 choice1 choice2 (x - y)
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z

round_round_eq fexp1 fexp2 choice1 choice2 (x - y)
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z

bpow (-2) <= / 2 * / 2
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
round_round_eq fexp1 fexp2 choice1 choice2 (x - y)
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z

bpow (-2) <= / 2 * / 2
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z

bpow (-2) <= / 4
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z

/ bpow 2 <= / 4
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z

4 <= bpow 2
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z

(2 * 2 <= beta * (beta * 1))%Z
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z

(2 * 2 <= beta * beta)%Z
now apply Zmult_le_compat; omega.
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2

round_round_eq fexp1 fexp2 choice1 choice2 (x - y)
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2

y < bpow (mag y)
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)
round_round_eq fexp1 fexp2 choice1 choice2 (x - y)
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2

y < bpow (mag y)
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2

Rabs y < bpow (mag y)
apply bpow_mag_gt.
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)

round_round_eq fexp1 fexp2 choice1 choice2 (x - y)
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)

Valid_exp fexp1
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)
Valid_exp fexp2
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)
0 < x - y
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)
(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)))%Z
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)
(fexp1 (mag (x - y)) <= mag (x - y))%Z
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)
midp' fexp1 (x - y) < x - y
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)
(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z -> midp' fexp1 (x - y) + / 2 * ulp beta fexp2 (x - y) < x - y
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)

Valid_exp fexp1
exact Vfexp1.
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)

Valid_exp fexp2
exact Vfexp2.
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)

0 < x - y
lra.
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)

(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)))%Z
apply Hexp4; omega.
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)

(fexp1 (mag (x - y)) <= mag (x - y))%Z
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)

(fexp1 (mag (x - y)) < mag (x - y))%Z
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)

(fexp1 (mag x - 1) < mag x - 1)%Z
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)
(mag x - 1 <= mag (x - y))%Z
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)

(fexp1 (mag x - 1) < mag x - 1)%Z
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)

(fexp1 (mag y) < mag y)%Z
now apply mag_generic_gt; [|apply Rgt_not_eq|].
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)

(mag x - 1 <= mag (x - y))%Z
now apply mag_minus_lb; [| |omega].
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)

midp' fexp1 (x - y) < x - y
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)

round beta fexp1 Zceil (x - y) - / 2 * ulp beta fexp1 (x - y) < x - y
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)

round beta fexp1 Zceil (x - y) - / 2 * ulp beta fexp1 (x - y) + (/ 2 * ulp beta fexp1 (x - y) - (x - y)) < x - y + (/ 2 * ulp beta fexp1 (x - y) - (x - y))
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)

round beta fexp1 Zceil (x - y) - x + y < / 2 * ulp beta fexp1 (x - y)
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)

round beta fexp1 Zceil (x - y) - (x - y) < / 2 * ulp beta fexp1 (x - y)
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)

round beta fexp1 Zceil (x - y) - (x - y) < bpow (fexp1 (mag (x - y)) - 2)
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)
bpow (fexp1 (mag (x - y)) - 2) <= / 2 * ulp beta fexp1 (x - y)
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)

round beta fexp1 Zceil (x - y) - (x - y) < bpow (fexp1 (mag (x - y)) - 2)
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)

y < bpow (fexp1 (mag (x - y)) - 2)
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)

bpow (mag y) <= bpow (fexp1 (mag (x - y)) - 2)
now apply bpow_le.
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)

bpow (fexp1 (mag (x - y)) - 2) <= / 2 * ulp beta fexp1 (x - y)
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)

bpow (fexp1 (mag (x - y)) - 2) <= / 2 * bpow (cexp beta fexp1 (x - y))
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)

bpow (fexp1 (mag (x - y)) - 2) <= / 2 * bpow (fexp1 (mag (x - y)))
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)

bpow (fexp1 (mag (x - y)) - 1 - 1) <= / 2 * bpow (fexp1 (mag (x - y)))
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)

bpow (fexp1 (mag (x - y)) - 1) * bpow (- (1)) <= / 2 * bpow (fexp1 (mag (x - y)))
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)

bpow (- (1)) * bpow (fexp1 (mag (x - y)) - 1) <= / 2 * bpow (fexp1 (mag (x - y)))
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)

0 <= bpow (- (1))
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)
0 <= bpow (fexp1 (mag (x - y)) - 1)
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)
bpow (- (1)) <= / 2
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)
bpow (fexp1 (mag (x - y)) - 1) <= bpow (fexp1 (mag (x - y)))
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)

0 <= bpow (- (1))
now apply bpow_ge_0.
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)

0 <= bpow (fexp1 (mag (x - y)) - 1)
now apply bpow_ge_0.
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)

bpow (- (1)) <= / 2
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)

/ IZR (beta * 1) <= / 2
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)

0 < 2
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)
2 <= IZR beta
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)

2 <= IZR beta
now apply IZR_le.
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)

bpow (fexp1 (mag (x - y)) - 1) <= bpow (fexp1 (mag (x - y)))
apply bpow_le; omega.
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)

(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z -> midp' fexp1 (x - y) + / 2 * ulp beta fexp2 (x - y) < x - y
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

midp' fexp1 (x - y) + / 2 * ulp beta fexp2 (x - y) < x - y
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

round beta fexp1 Zceil (x - y) - / 2 * ulp beta fexp1 (x - y) + / 2 * ulp beta fexp2 (x - y) < x - y
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

round beta fexp1 Zceil (x - y) - / 2 * ulp beta fexp1 (x - y) + / 2 * ulp beta fexp2 (x - y) + (/ 2 * ulp beta fexp1 (x - y) - (x - y) - / 2 * ulp beta fexp2 (x - y)) < x - y + (/ 2 * ulp beta fexp1 (x - y) - (x - y) - / 2 * ulp beta fexp2 (x - y))
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

round beta fexp1 Zceil (x - y) - x + y < / 2 * ulp beta fexp1 (x - y) - / 2 * ulp beta fexp2 (x - y)
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

round beta fexp1 Zceil (x - y) - (x - y) < / 2 * ulp beta fexp1 (x - y) - / 2 * ulp beta fexp2 (x - y)
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

y < / 2 * ulp beta fexp1 (x - y) - / 2 * ulp beta fexp2 (x - y)
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

bpow (mag y) <= / 2 * ulp beta fexp1 (x - y) - / 2 * ulp beta fexp2 (x - y)
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

bpow (fexp1 (mag (x - y)) - 2) <= / 2 * ulp beta fexp1 (x - y) - / 2 * ulp beta fexp2 (x - y)
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

bpow (fexp1 (mag (x - y)) - 1 - 1) <= / 2 * ulp beta fexp1 (x - y) - / 2 * ulp beta fexp2 (x - y)
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

bpow (fexp1 (mag (x - y)) - 1) * bpow (- (1)) <= / 2 * ulp beta fexp1 (x - y) - / 2 * ulp beta fexp2 (x - y)
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

bpow (fexp1 (mag (x - y)) - 1) * bpow (- (1)) <= / 2 * (ulp beta fexp1 (x - y) - ulp beta fexp2 (x - y))
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

0 <= bpow (- (1))
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z
0 <= bpow (fexp1 (mag (x - y)) - 1)
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z
bpow (- (1)) <= / 2
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z
bpow (fexp1 (mag (x - y)) - 1) <= ulp beta fexp1 (x - y) - ulp beta fexp2 (x - y)
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

0 <= bpow (- (1))
apply bpow_ge_0.
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

0 <= bpow (fexp1 (mag (x - y)) - 1)
apply bpow_ge_0.
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

bpow (- (1)) <= / 2
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

/ IZR (beta * 1) <= / 2
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

2 <= IZR beta
now apply IZR_le.
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

bpow (fexp1 (mag (x - y)) - 1) <= ulp beta fexp1 (x - y) - ulp beta fexp2 (x - y)
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

bpow (fexp1 (mag (x - y)) - 1) <= bpow (cexp beta fexp1 (x - y)) - bpow (cexp beta fexp2 (x - y))
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

bpow (fexp1 (mag (x - y)) - 1) <= bpow (fexp1 (mag (x - y))) - bpow (fexp2 (mag (x - y)))
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

bpow (fexp1 (mag (x - y)) - 1) + bpow (fexp2 (mag (x - y))) <= bpow (fexp1 (mag (x - y)))
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

bpow (fexp1 (mag (x - y)) - 1) + bpow (fexp2 (mag (x - y))) <= 2 * bpow (fexp1 (mag (x - y)) - 1)
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z
2 * bpow (fexp1 (mag (x - y)) - 1) <= bpow (fexp1 (mag (x - y)))
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

bpow (fexp1 (mag (x - y)) - 1) + bpow (fexp2 (mag (x - y))) <= 2 * bpow (fexp1 (mag (x - y)) - 1)
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

bpow (fexp1 (mag (x - y)) - 1) + bpow (fexp2 (mag (x - y))) <= bpow (fexp1 (mag (x - y)) - 1) + bpow (fexp1 (mag (x - y)) - 1)
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

bpow (fexp2 (mag (x - y))) <= bpow (fexp1 (mag (x - y)) - 1)
now apply bpow_le.
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

2 * bpow (fexp1 (mag (x - y)) - 1) <= bpow (fexp1 (mag (x - y)))
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

2 * (bpow (fexp1 (mag (x - y))) * bpow (- (1))) <= bpow (fexp1 (mag (x - y)))
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

bpow (fexp1 (mag (x - y))) * (bpow (- (1)) * 2) <= bpow (fexp1 (mag (x - y)))
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

bpow (fexp1 (mag (x - y))) * (bpow (- (1)) * 2) <= bpow (fexp1 (mag (x - y))) * 1
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

bpow (- (1)) * 2 <= 1
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

/ IZR (beta * 1) * 2 <= 1
beta:radix
Hbeta:(2 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

/ IZR beta * 2 <= 1
beta:radix
Hbeta:/ IZR beta <= / 2
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

/ IZR beta * 2 <= 1
beta:radix
Hbeta:2 <= IZR beta
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z
0 < 2
beta:radix
Hbeta:/ IZR beta <= / 2
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

/ IZR beta * 2 <= 1
beta:radix
Hbeta:2 <= IZR beta
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z
0 < 2
beta:radix
Hbeta:2 <= IZR beta
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow2:bpow (-2) <= / 2 * / 2
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

0 < 2
apply Rlt_0_2. Qed. (* round_round_minus_aux{0,1,2} together *)
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall choice1 choice2 : Z -> bool, round_round_plus_hyp fexp1 fexp2 -> forall x y : R, 0 < y -> y <= x -> generic_format beta fexp1 x -> generic_format beta fexp1 y -> round_round_eq fexp1 fexp2 choice1 choice2 (x - y)
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall choice1 choice2 : Z -> bool, round_round_plus_hyp fexp1 fexp2 -> forall x y : R, 0 < y -> y <= x -> generic_format beta fexp1 x -> generic_format beta fexp1 y -> round_round_eq fexp1 fexp2 choice1 choice2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y

round_round_eq fexp1 fexp2 choice1 choice2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x

round_round_eq fexp1 fexp2 choice1 choice2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y = x

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y <> x
round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y = x

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y = x

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) 0) = round beta fexp1 (Znearest choice1) 0
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y = x

round beta fexp1 (Znearest choice1) 0 = round beta fexp1 (Znearest choice1) 0
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y = x
Valid_rnd (Znearest choice2)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y = x

round beta fexp1 (Znearest choice1) 0 = round beta fexp1 (Znearest choice1) 0
reflexivity.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y = x

Valid_rnd (Znearest choice2)
now apply valid_rnd_N.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y <> x

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y <> x
Hyx':y < x

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y <> x
Hyx':y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y <> x
Hyx':y < x
Hly:(fexp1 (mag x) - 2 < mag y)%Z
round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y <> x
Hyx':y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y <> x
Hyx':y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y <> x
Hyx':y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(fexp1 (mag (x - y)) - 2 < mag y)%Z
round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y <> x
Hyx':y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 2)%Z

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
now apply round_round_minus_aux2.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y <> x
Hyx':y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(fexp1 (mag (x - y)) - 2 < mag y)%Z

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y <> x
Hyx':y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(fexp1 (mag (x - y)) - 2 < mag y)%Z

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y <> x
Hyx':y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(fexp1 (mag (x - y)) - 2 < mag y)%Z

round beta fexp1 (Znearest choice1) (x - y) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y <> x
Hyx':y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(fexp1 (mag (x - y)) - 2 < mag y)%Z
Valid_rnd (Znearest choice2)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y <> x
Hyx':y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(fexp1 (mag (x - y)) - 2 < mag y)%Z
generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y <> x
Hyx':y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(fexp1 (mag (x - y)) - 2 < mag y)%Z

round beta fexp1 (Znearest choice1) (x - y) = round beta fexp1 (Znearest choice1) (x - y)
reflexivity.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y <> x
Hyx':y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(fexp1 (mag (x - y)) - 2 < mag y)%Z

Valid_rnd (Znearest choice2)
now apply valid_rnd_N.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y <> x
Hyx':y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(fexp1 (mag (x - y)) - 2 < mag y)%Z

generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y <> x
Hyx':y < x
Hly:(mag y <= fexp1 (mag x) - 2)%Z
Hly':(fexp1 (mag (x - y)) - 2 < mag y)%Z
Hf1:(fexp1 (mag (x - y)) - 1 <= mag y)%Z

generic_format beta fexp2 (x - y)
now apply (round_round_minus_aux1 fexp1). }
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y <> x
Hyx':y < x
Hly:(fexp1 (mag x) - 2 < mag y)%Z

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y <> x
Hyx':y < x
Hly:(fexp1 (mag x) - 2 < mag y)%Z

round beta fexp1 (Znearest choice1) (x - y) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y <> x
Hyx':y < x
Hly:(fexp1 (mag x) - 2 < mag y)%Z
Valid_rnd (Znearest choice2)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y <> x
Hyx':y < x
Hly:(fexp1 (mag x) - 2 < mag y)%Z
generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y <> x
Hyx':y < x
Hly:(fexp1 (mag x) - 2 < mag y)%Z

round beta fexp1 (Znearest choice1) (x - y) = round beta fexp1 (Znearest choice1) (x - y)
reflexivity.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y <> x
Hyx':y < x
Hly:(fexp1 (mag x) - 2 < mag y)%Z

Valid_rnd (Znearest choice2)
now apply valid_rnd_N.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y <> x
Hyx':y < x
Hly:(fexp1 (mag x) - 2 < mag y)%Z

generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y <> x
Hyx':y < x
Hly:(fexp1 (mag x) - 2 < mag y)%Z
Hf1:(fexp1 (mag x) - 1 <= mag y)%Z

generic_format beta fexp2 (x - y)
now apply (round_round_minus_aux0 fexp1). Qed.
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall choice1 choice2 : Z -> bool, round_round_plus_hyp fexp1 fexp2 -> forall x y : R, 0 <= x -> 0 <= y -> generic_format beta fexp1 x -> generic_format beta fexp1 y -> round_round_eq fexp1 fexp2 choice1 choice2 (x - y)
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall choice1 choice2 : Z -> bool, round_round_plus_hyp fexp1 fexp2 -> forall x y : R, 0 <= x -> 0 <= y -> generic_format beta fexp1 x -> generic_format beta fexp1 y -> round_round_eq fexp1 fexp2 choice1 choice2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y

round_round_eq fexp1 fexp2 choice1 choice2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (- y)) = round beta fexp1 (Znearest choice1) (- y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0

- round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) (round beta fexp2 (Znearest (fun t : Z => negb (choice2 (- (t + 1))%Z))) y) = - round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0

- round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) y = - round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0
Valid_rnd (Znearest (fun t : Z => negb (choice2 (- (t + 1))%Z)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0
generic_format beta fexp2 y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0

- round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) y = - round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) y
reflexivity.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0

Valid_rnd (Znearest (fun t : Z => negb (choice2 (- (t + 1))%Z)))
now apply valid_rnd_N.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0

generic_format beta fexp2 y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0

y <> 0 -> (fexp2 (mag y) <= fexp1 (mag y))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0
generic_format beta fexp1 y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0

y <> 0 -> (fexp2 (mag y) <= fexp1 (mag y))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0
generic_format beta fexp1 y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0

generic_format beta fexp1 y
exact Fy.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Nzy:y <> 0
round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0

round beta fexp1 (Znearest choice1) x = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0
Valid_rnd (Znearest choice2)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0
generic_format beta fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0

round beta fexp1 (Znearest choice1) x = round beta fexp1 (Znearest choice1) x
reflexivity.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0

Valid_rnd (Znearest choice2)
now apply valid_rnd_N.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0

generic_format beta fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0

x <> 0 -> (fexp2 (mag x) <= fexp1 (mag x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0
generic_format beta fexp1 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0

x <> 0 -> (fexp2 (mag x) <= fexp1 (mag x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0
generic_format beta fexp1 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0

generic_format beta fexp1 x
exact Fx.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Nzy:y <> 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Nzy:y <> 0
Px:0 < x

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Nzy:y <> 0
Px:0 < x
Py:0 < y

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Nzy:y <> 0
Px:0 < x
Py:0 < y
H:x < y

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Nzy:y <> 0
Px:0 < x
Py:0 < y
H:y <= x
round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Nzy:y <> 0
Px:0 < x
Py:0 < y
H:x < y

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Nzy:y <> 0
Px:0 < x
Py:0 < y
H:x <= y

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Nzy:y <> 0
Px:0 < x
Py:0 < y
H:x <= y

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (- (y - x))) = round beta fexp1 (Znearest choice1) (- (y - x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Nzy:y <> 0
Px:0 < x
Py:0 < y
H:x <= y

- round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) (round beta fexp2 (Znearest (fun t : Z => negb (choice2 (- (t + 1))%Z))) (y - x)) = - round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) (y - x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Nzy:y <> 0
Px:0 < x
Py:0 < y
H:x <= y

round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) (round beta fexp2 (Znearest (fun t : Z => negb (choice2 (- (t + 1))%Z))) (y - x)) = round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) (y - x)
now apply round_round_minus_aux3.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Nzy:y <> 0
Px:0 < x
Py:0 < y
H:y <= x

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
now apply round_round_minus_aux3. Qed.
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall choice1 choice2 : Z -> bool, round_round_plus_hyp fexp1 fexp2 -> forall x y : R, generic_format beta fexp1 x -> generic_format beta fexp1 y -> round_round_eq fexp1 fexp2 choice1 choice2 (x + y)
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall choice1 choice2 : Z -> bool, round_round_plus_hyp fexp1 fexp2 -> forall x y : R, generic_format beta fexp1 x -> generic_format beta fexp1 y -> round_round_eq fexp1 fexp2 choice1 choice2 (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y

round_round_eq fexp1 fexp2 choice1 choice2 (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Sx:x < 0
Sy:y < 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Sx:x < 0
Sy:0 <= y
round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Sx:0 <= x
Sy:y < 0
round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Sx:0 <= x
Sy:0 <= y
round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Sx:x < 0
Sy:y < 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Sx:x < 0
Sy:y < 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (- (- x - y))) = round beta fexp1 (Znearest choice1) (- (- x - y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Sx:x < 0
Sy:y < 0

- round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) (round beta fexp2 (Znearest (fun t : Z => negb (choice2 (- (t + 1))%Z))) (- x - y)) = - round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) (- x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Sx:x < 0
Sy:y < 0

round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) (round beta fexp2 (Znearest (fun t : Z => negb (choice2 (- (t + 1))%Z))) (- x - y)) = round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) (- x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Sx:x < 0
Sy:y < 0
Px:0 <= - x

round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) (round beta fexp2 (Znearest (fun t : Z => negb (choice2 (- (t + 1))%Z))) (- x - y)) = round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) (- x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Sx:x < 0
Sy:y < 0
Px:0 <= - x
Py:0 <= - y

round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) (round beta fexp2 (Znearest (fun t : Z => negb (choice2 (- (t + 1))%Z))) (- x - y)) = round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) (- x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 (- x)
Fy:generic_format beta fexp1 y
Sx:x < 0
Sy:y < 0
Px:0 <= - x
Py:0 <= - y

round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) (round beta fexp2 (Znearest (fun t : Z => negb (choice2 (- (t + 1))%Z))) (- x - y)) = round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) (- x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 (- x)
Fy:generic_format beta fexp1 (- y)
Sx:x < 0
Sy:y < 0
Px:0 <= - x
Py:0 <= - y

round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) (round beta fexp2 (Znearest (fun t : Z => negb (choice2 (- (t + 1))%Z))) (- x - y)) = round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) (- x - y)
now apply round_round_plus_aux.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Sx:x < 0
Sy:0 <= y

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Sx:x < 0
Sy:0 <= y

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (y - - x)) = round beta fexp1 (Znearest choice1) (y - - x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Sx:x < 0
Sy:0 <= y
Px:0 <= - x

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (y - - x)) = round beta fexp1 (Znearest choice1) (y - - x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 (- x)
Fy:generic_format beta fexp1 y
Sx:x < 0
Sy:0 <= y
Px:0 <= - x

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (y - - x)) = round beta fexp1 (Znearest choice1) (y - - x)
now apply round_round_minus_aux.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Sx:0 <= x
Sy:y < 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Sx:0 <= x
Sy:y < 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - - y)) = round beta fexp1 (Znearest choice1) (x - - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Sx:0 <= x
Sy:y < 0
Py:0 <= - y

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - - y)) = round beta fexp1 (Znearest choice1) (x - - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 (- y)
Sx:0 <= x
Sy:y < 0
Py:0 <= - y

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - - y)) = round beta fexp1 (Znearest choice1) (x - - y)
now apply round_round_minus_aux.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Sx:0 <= x
Sy:0 <= y

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
now apply round_round_plus_aux. Qed.
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall choice1 choice2 : Z -> bool, round_round_plus_hyp fexp1 fexp2 -> forall x y : R, generic_format beta fexp1 x -> generic_format beta fexp1 y -> round_round_eq fexp1 fexp2 choice1 choice2 (x - y)
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall choice1 choice2 : Z -> bool, round_round_plus_hyp fexp1 fexp2 -> forall x y : R, generic_format beta fexp1 x -> generic_format beta fexp1 y -> round_round_eq fexp1 fexp2 choice1 choice2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y

round_round_eq fexp1 fexp2 choice1 choice2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y

round_round_eq fexp1 fexp2 choice1 choice2 (x + - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 (- y)

round_round_eq fexp1 fexp2 choice1 choice2 (x + - y)
now apply round_round_plus. Qed. Section Double_round_plus_FLX. Variable prec : Z. Variable prec' : Z. Context { prec_gt_0_ : Prec_gt_0 prec }. Context { prec_gt_0_' : Prec_gt_0 prec' }.
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(2 * prec + 1 <= prec')%Z -> round_round_plus_hyp (FLX_exp prec) (FLX_exp prec')
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(2 * prec + 1 <= prec')%Z -> round_round_plus_hyp (FLX_exp prec) (FLX_exp prec')
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hprec:(2 * prec + 1 <= prec')%Z

round_round_plus_hyp (FLX_exp prec) (FLX_exp prec')
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hprec:(2 * prec + 1 <= prec')%Z

round_round_plus_hyp (fun e : Z => (e - prec)%Z) (fun e : Z => (e - prec')%Z)
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hprec:(2 * prec + 1 <= prec')%Z
ex, ey:Z

(ex - 1 <= ey)%Z -> (ex - prec' <= ey - prec)%Z
beta:radix
prec, prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hprec:(2 * prec + 1 <= prec')%Z
ex, ey:Z

(ex - 1 <= ey)%Z -> (ex - prec' <= ey - prec)%Z
omega. Qed.
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

forall choice1 choice2 : Z -> bool, (2 * prec + 1 <= prec')%Z -> forall x y : R, FLX_format beta prec x -> FLX_format beta prec y -> round_round_eq (FLX_exp prec) (FLX_exp prec') choice1 choice2 (x + y)
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

forall choice1 choice2 : Z -> bool, (2 * prec + 1 <= prec')%Z -> forall x y : R, FLX_format beta prec x -> FLX_format beta prec y -> round_round_eq (FLX_exp prec) (FLX_exp prec') choice1 choice2 (x + y)
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y

round_round_eq (FLX_exp prec) (FLX_exp prec') choice1 choice2 (x + y)
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y

Valid_exp (FLX_exp prec)
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y
Valid_exp (FLX_exp prec')
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y
round_round_plus_hyp (FLX_exp prec) (FLX_exp prec')
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y
generic_format beta (FLX_exp prec) x
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y
generic_format beta (FLX_exp prec) y
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y

Valid_exp (FLX_exp prec)
now apply FLX_exp_valid.
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y

Valid_exp (FLX_exp prec')
now apply FLX_exp_valid.
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y

round_round_plus_hyp (FLX_exp prec) (FLX_exp prec')
now apply FLX_round_round_plus_hyp.
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y

generic_format beta (FLX_exp prec) x
now apply generic_format_FLX.
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y

generic_format beta (FLX_exp prec) y
now apply generic_format_FLX. Qed.
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

forall choice1 choice2 : Z -> bool, (2 * prec + 1 <= prec')%Z -> forall x y : R, FLX_format beta prec x -> FLX_format beta prec y -> round_round_eq (FLX_exp prec) (FLX_exp prec') choice1 choice2 (x - y)
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

forall choice1 choice2 : Z -> bool, (2 * prec + 1 <= prec')%Z -> forall x y : R, FLX_format beta prec x -> FLX_format beta prec y -> round_round_eq (FLX_exp prec) (FLX_exp prec') choice1 choice2 (x - y)
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y

round_round_eq (FLX_exp prec) (FLX_exp prec') choice1 choice2 (x - y)
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y

Valid_exp (FLX_exp prec)
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y
Valid_exp (FLX_exp prec')
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y
round_round_plus_hyp (FLX_exp prec) (FLX_exp prec')
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y
generic_format beta (FLX_exp prec) x
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y
generic_format beta (FLX_exp prec) y
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y

Valid_exp (FLX_exp prec)
now apply FLX_exp_valid.
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y

Valid_exp (FLX_exp prec')
now apply FLX_exp_valid.
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y

round_round_plus_hyp (FLX_exp prec) (FLX_exp prec')
now apply FLX_round_round_plus_hyp.
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y

generic_format beta (FLX_exp prec) x
now apply generic_format_FLX.
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y

generic_format beta (FLX_exp prec) y
now apply generic_format_FLX. Qed. End Double_round_plus_FLX. Section Double_round_plus_FLT. Variable emin prec : Z. Variable emin' prec' : Z. Context { prec_gt_0_ : Prec_gt_0 prec }. Context { prec_gt_0_' : Prec_gt_0 prec' }.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(emin' <= emin)%Z -> (2 * prec + 1 <= prec')%Z -> round_round_plus_hyp (FLT_exp emin prec) (FLT_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(emin' <= emin)%Z -> (2 * prec + 1 <= prec')%Z -> round_round_plus_hyp (FLT_exp emin prec) (FLT_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec + 1 <= prec')%Z

round_round_plus_hyp (FLT_exp emin prec) (FLT_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec + 1 <= prec')%Z

round_round_plus_hyp (fun e : Z => Z.max (e - prec) emin) (fun e : Z => Z.max (e - prec') emin')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec + 1 <= prec')%Z
ex, ey:Z

(Z.max (ex + 1 - prec) emin - 1 <= ey)%Z -> (Z.max (ex - prec') emin' <= Z.max (ey - prec) emin)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec + 1 <= prec')%Z
ex, ey:Z
(Z.max (ex - 1 - prec) emin + 1 <= ey)%Z -> (Z.max (ex - prec') emin' <= Z.max (ey - prec) emin)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec + 1 <= prec')%Z
ex, ey:Z
(Z.max (ex - prec) emin - 1 <= ey)%Z -> (Z.max (ex - prec') emin' <= Z.max (ey - prec) emin)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec + 1 <= prec')%Z
ex, ey:Z
(ex - 1 <= ey)%Z -> (Z.max (ex - prec') emin' <= Z.max (ey - prec) emin)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec + 1 <= prec')%Z
ex, ey:Z

(Z.max (ex + 1 - prec) emin - 1 <= ey)%Z -> (Z.max (ex - prec') emin' <= Z.max (ey - prec) emin)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec + 1 <= prec')%Z
ex, ey:Z

(ex + 1 - prec >= emin)%Z /\ Z.max (ex + 1 - prec) emin = (ex + 1 - prec)%Z \/ (ex + 1 - prec < emin)%Z /\ Z.max (ex + 1 - prec) emin = emin -> (Z.max (ex + 1 - prec) emin - 1 <= ey)%Z -> (Z.max (ex - prec') emin' <= Z.max (ey - prec) emin)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec + 1 <= prec')%Z
ex, ey:Z

(ex - prec' >= emin')%Z /\ Z.max (ex - prec') emin' = (ex - prec')%Z \/ (ex - prec' < emin')%Z /\ Z.max (ex - prec') emin' = emin' -> (ex + 1 - prec >= emin)%Z /\ Z.max (ex + 1 - prec) emin = (ex + 1 - prec)%Z \/ (ex + 1 - prec < emin)%Z /\ Z.max (ex + 1 - prec) emin = emin -> (Z.max (ex + 1 - prec) emin - 1 <= ey)%Z -> (Z.max (ex - prec') emin' <= Z.max (ey - prec) emin)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec + 1 <= prec')%Z
ex, ey:Z

(ey - prec >= emin)%Z /\ Z.max (ey - prec) emin = (ey - prec)%Z \/ (ey - prec < emin)%Z /\ Z.max (ey - prec) emin = emin -> (ex - prec' >= emin')%Z /\ Z.max (ex - prec') emin' = (ex - prec')%Z \/ (ex - prec' < emin')%Z /\ Z.max (ex - prec') emin' = emin' -> (ex + 1 - prec >= emin)%Z /\ Z.max (ex + 1 - prec) emin = (ex + 1 - prec)%Z \/ (ex + 1 - prec < emin)%Z /\ Z.max (ex + 1 - prec) emin = emin -> (Z.max (ex + 1 - prec) emin - 1 <= ey)%Z -> (Z.max (ex - prec') emin' <= Z.max (ey - prec) emin)%Z
omega.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec + 1 <= prec')%Z
ex, ey:Z

(Z.max (ex - 1 - prec) emin + 1 <= ey)%Z -> (Z.max (ex - prec') emin' <= Z.max (ey - prec) emin)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec + 1 <= prec')%Z
ex, ey:Z

(ex - 1 - prec >= emin)%Z /\ Z.max (ex - 1 - prec) emin = (ex - 1 - prec)%Z \/ (ex - 1 - prec < emin)%Z /\ Z.max (ex - 1 - prec) emin = emin -> (Z.max (ex - 1 - prec) emin + 1 <= ey)%Z -> (Z.max (ex - prec') emin' <= Z.max (ey - prec) emin)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec + 1 <= prec')%Z
ex, ey:Z

(ex - prec' >= emin')%Z /\ Z.max (ex - prec') emin' = (ex - prec')%Z \/ (ex - prec' < emin')%Z /\ Z.max (ex - prec') emin' = emin' -> (ex - 1 - prec >= emin)%Z /\ Z.max (ex - 1 - prec) emin = (ex - 1 - prec)%Z \/ (ex - 1 - prec < emin)%Z /\ Z.max (ex - 1 - prec) emin = emin -> (Z.max (ex - 1 - prec) emin + 1 <= ey)%Z -> (Z.max (ex - prec') emin' <= Z.max (ey - prec) emin)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec + 1 <= prec')%Z
ex, ey:Z

(ey - prec >= emin)%Z /\ Z.max (ey - prec) emin = (ey - prec)%Z \/ (ey - prec < emin)%Z /\ Z.max (ey - prec) emin = emin -> (ex - prec' >= emin')%Z /\ Z.max (ex - prec') emin' = (ex - prec')%Z \/ (ex - prec' < emin')%Z /\ Z.max (ex - prec') emin' = emin' -> (ex - 1 - prec >= emin)%Z /\ Z.max (ex - 1 - prec) emin = (ex - 1 - prec)%Z \/ (ex - 1 - prec < emin)%Z /\ Z.max (ex - 1 - prec) emin = emin -> (Z.max (ex - 1 - prec) emin + 1 <= ey)%Z -> (Z.max (ex - prec') emin' <= Z.max (ey - prec) emin)%Z
omega.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec + 1 <= prec')%Z
ex, ey:Z

(Z.max (ex - prec) emin - 1 <= ey)%Z -> (Z.max (ex - prec') emin' <= Z.max (ey - prec) emin)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec + 1 <= prec')%Z
ex, ey:Z

(ex - prec >= emin)%Z /\ Z.max (ex - prec) emin = (ex - prec)%Z \/ (ex - prec < emin)%Z /\ Z.max (ex - prec) emin = emin -> (Z.max (ex - prec) emin - 1 <= ey)%Z -> (Z.max (ex - prec') emin' <= Z.max (ey - prec) emin)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec + 1 <= prec')%Z
ex, ey:Z

(ex - prec' >= emin')%Z /\ Z.max (ex - prec') emin' = (ex - prec')%Z \/ (ex - prec' < emin')%Z /\ Z.max (ex - prec') emin' = emin' -> (ex - prec >= emin)%Z /\ Z.max (ex - prec) emin = (ex - prec)%Z \/ (ex - prec < emin)%Z /\ Z.max (ex - prec) emin = emin -> (Z.max (ex - prec) emin - 1 <= ey)%Z -> (Z.max (ex - prec') emin' <= Z.max (ey - prec) emin)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec + 1 <= prec')%Z
ex, ey:Z

(ey - prec >= emin)%Z /\ Z.max (ey - prec) emin = (ey - prec)%Z \/ (ey - prec < emin)%Z /\ Z.max (ey - prec) emin = emin -> (ex - prec' >= emin')%Z /\ Z.max (ex - prec') emin' = (ex - prec')%Z \/ (ex - prec' < emin')%Z /\ Z.max (ex - prec') emin' = emin' -> (ex - prec >= emin)%Z /\ Z.max (ex - prec) emin = (ex - prec)%Z \/ (ex - prec < emin)%Z /\ Z.max (ex - prec) emin = emin -> (Z.max (ex - prec) emin - 1 <= ey)%Z -> (Z.max (ex - prec') emin' <= Z.max (ey - prec) emin)%Z
omega.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec + 1 <= prec')%Z
ex, ey:Z

(ex - 1 <= ey)%Z -> (Z.max (ex - prec') emin' <= Z.max (ey - prec) emin)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec + 1 <= prec')%Z
ex, ey:Z

(ex - 1 <= ey)%Z -> (Z.max (ex - prec') emin' <= Z.max (ey - prec) emin)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec + 1 <= prec')%Z
ex, ey:Z

(ex - prec' >= emin')%Z /\ Z.max (ex - prec') emin' = (ex - prec')%Z \/ (ex - prec' < emin')%Z /\ Z.max (ex - prec') emin' = emin' -> (ex - 1 <= ey)%Z -> (Z.max (ex - prec') emin' <= Z.max (ey - prec) emin)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec + 1 <= prec')%Z
ex, ey:Z

(ey - prec >= emin)%Z /\ Z.max (ey - prec) emin = (ey - prec)%Z \/ (ey - prec < emin)%Z /\ Z.max (ey - prec) emin = emin -> (ex - prec' >= emin')%Z /\ Z.max (ex - prec') emin' = (ex - prec')%Z \/ (ex - prec' < emin')%Z /\ Z.max (ex - prec') emin' = emin' -> (ex - 1 <= ey)%Z -> (Z.max (ex - prec') emin' <= Z.max (ey - prec) emin)%Z
omega. Qed.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

forall choice1 choice2 : Z -> bool, (emin' <= emin)%Z -> (2 * prec + 1 <= prec')%Z -> forall x y : R, FLT_format beta emin prec x -> FLT_format beta emin prec y -> round_round_eq (FLT_exp emin prec) (FLT_exp emin' prec') choice1 choice2 (x + y)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

forall choice1 choice2 : Z -> bool, (emin' <= emin)%Z -> (2 * prec + 1 <= prec')%Z -> forall x y : R, FLT_format beta emin prec x -> FLT_format beta emin prec y -> round_round_eq (FLT_exp emin prec) (FLT_exp emin' prec') choice1 choice2 (x + y)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y

round_round_eq (FLT_exp emin prec) (FLT_exp emin' prec') choice1 choice2 (x + y)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y

Valid_exp (FLT_exp emin prec)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y
Valid_exp (FLT_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y
round_round_plus_hyp (FLT_exp emin prec) (FLT_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y
generic_format beta (FLT_exp emin prec) x
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y
generic_format beta (FLT_exp emin prec) y
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y

Valid_exp (FLT_exp emin prec)
now apply FLT_exp_valid.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y

Valid_exp (FLT_exp emin' prec')
now apply FLT_exp_valid.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y

round_round_plus_hyp (FLT_exp emin prec) (FLT_exp emin' prec')
now apply FLT_round_round_plus_hyp.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y

generic_format beta (FLT_exp emin prec) x
now apply generic_format_FLT.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y

generic_format beta (FLT_exp emin prec) y
now apply generic_format_FLT. Qed.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

forall choice1 choice2 : Z -> bool, (emin' <= emin)%Z -> (2 * prec + 1 <= prec')%Z -> forall x y : R, FLT_format beta emin prec x -> FLT_format beta emin prec y -> round_round_eq (FLT_exp emin prec) (FLT_exp emin' prec') choice1 choice2 (x - y)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

forall choice1 choice2 : Z -> bool, (emin' <= emin)%Z -> (2 * prec + 1 <= prec')%Z -> forall x y : R, FLT_format beta emin prec x -> FLT_format beta emin prec y -> round_round_eq (FLT_exp emin prec) (FLT_exp emin' prec') choice1 choice2 (x - y)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y

round_round_eq (FLT_exp emin prec) (FLT_exp emin' prec') choice1 choice2 (x - y)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y

Valid_exp (FLT_exp emin prec)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y
Valid_exp (FLT_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y
round_round_plus_hyp (FLT_exp emin prec) (FLT_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y
generic_format beta (FLT_exp emin prec) x
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y
generic_format beta (FLT_exp emin prec) y
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y

Valid_exp (FLT_exp emin prec)
now apply FLT_exp_valid.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y

Valid_exp (FLT_exp emin' prec')
now apply FLT_exp_valid.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y

round_round_plus_hyp (FLT_exp emin prec) (FLT_exp emin' prec')
now apply FLT_round_round_plus_hyp.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y

generic_format beta (FLT_exp emin prec) x
now apply generic_format_FLT.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y

generic_format beta (FLT_exp emin prec) y
now apply generic_format_FLT. Qed. End Double_round_plus_FLT. Section Double_round_plus_FTZ. Variable emin prec : Z. Variable emin' prec' : Z. Context { prec_gt_0_ : Prec_gt_0 prec }. Context { prec_gt_0_' : Prec_gt_0 prec' }.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(emin' + prec' <= emin + 1)%Z -> (2 * prec + 1 <= prec')%Z -> round_round_plus_hyp (FTZ_exp emin prec) (FTZ_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(emin' + prec' <= emin + 1)%Z -> (2 * prec + 1 <= prec')%Z -> round_round_plus_hyp (FTZ_exp emin prec) (FTZ_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z

round_round_plus_hyp (FTZ_exp emin prec) (FTZ_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z

round_round_plus_hyp (fun e : Z => if (e - prec <? emin)%Z then (emin + prec - 1)%Z else (e - prec)%Z) (fun e : Z => if (e - prec' <? emin')%Z then (emin' + prec' - 1)%Z else (e - prec')%Z)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':(0 < prec')%Z
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z

round_round_plus_hyp (fun e : Z => if (e - prec <? emin)%Z then (emin + prec - 1)%Z else (e - prec)%Z) (fun e : Z => if (e - prec' <? emin')%Z then (emin' + prec' - 1)%Z else (e - prec')%Z)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':(0 < prec')%Z
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z
ex, ey:Z

((if ex + 1 - prec <? emin then emin + prec - 1 else ex + 1 - prec) - 1 <= ey)%Z -> ((if ex - prec' <? emin' then emin' + prec' - 1 else ex - prec') <= (if ey - prec <? emin then emin + prec - 1 else ey - prec))%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':(0 < prec')%Z
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z
ex, ey:Z
((if ex - 1 - prec <? emin then emin + prec - 1 else ex - 1 - prec) + 1 <= ey)%Z -> ((if ex - prec' <? emin' then emin' + prec' - 1 else ex - prec') <= (if ey - prec <? emin then emin + prec - 1 else ey - prec))%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':(0 < prec')%Z
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z
ex, ey:Z
((if ex - prec <? emin then emin + prec - 1 else ex - prec) - 1 <= ey)%Z -> ((if ex - prec' <? emin' then emin' + prec' - 1 else ex - prec') <= (if ey - prec <? emin then emin + prec - 1 else ey - prec))%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':(0 < prec')%Z
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z
ex, ey:Z
(ex - 1 <= ey)%Z -> ((if ex - prec' <? emin' then emin' + prec' - 1 else ex - prec') <= (if ey - prec <? emin then emin + prec - 1 else ey - prec))%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':(0 < prec')%Z
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z
ex, ey:Z

((if ex + 1 - prec <? emin then emin + prec - 1 else ex + 1 - prec) - 1 <= ey)%Z -> ((if ex - prec' <? emin' then emin' + prec' - 1 else ex - prec') <= (if ey - prec <? emin then emin + prec - 1 else ey - prec))%Z
destruct (Z.ltb_spec (ex + 1 - prec) emin); destruct (Z.ltb_spec (ex - prec') emin'); destruct (Z.ltb_spec (ey - prec) emin); omega.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':(0 < prec')%Z
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z
ex, ey:Z

((if ex - 1 - prec <? emin then emin + prec - 1 else ex - 1 - prec) + 1 <= ey)%Z -> ((if ex - prec' <? emin' then emin' + prec' - 1 else ex - prec') <= (if ey - prec <? emin then emin + prec - 1 else ey - prec))%Z
destruct (Z.ltb_spec (ex - 1 - prec) emin); destruct (Z.ltb_spec (ex - prec') emin'); destruct (Z.ltb_spec (ey - prec) emin); omega.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':(0 < prec')%Z
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z
ex, ey:Z

((if ex - prec <? emin then emin + prec - 1 else ex - prec) - 1 <= ey)%Z -> ((if ex - prec' <? emin' then emin' + prec' - 1 else ex - prec') <= (if ey - prec <? emin then emin + prec - 1 else ey - prec))%Z
destruct (Z.ltb_spec (ex - prec) emin); destruct (Z.ltb_spec (ex - prec') emin'); destruct (Z.ltb_spec (ey - prec) emin); omega.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':(0 < prec')%Z
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z
ex, ey:Z

(ex - 1 <= ey)%Z -> ((if ex - prec' <? emin' then emin' + prec' - 1 else ex - prec') <= (if ey - prec <? emin then emin + prec - 1 else ey - prec))%Z
destruct (Z.ltb_spec (ex - prec') emin'); destruct (Z.ltb_spec (ey - prec) emin); omega. Qed.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

forall choice1 choice2 : Z -> bool, (emin' + prec' <= emin + 1)%Z -> (2 * prec + 1 <= prec')%Z -> forall x y : R, FTZ_format beta emin prec x -> FTZ_format beta emin prec y -> round_round_eq (FTZ_exp emin prec) (FTZ_exp emin' prec') choice1 choice2 (x + y)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

forall choice1 choice2 : Z -> bool, (emin' + prec' <= emin + 1)%Z -> (2 * prec + 1 <= prec')%Z -> forall x y : R, FTZ_format beta emin prec x -> FTZ_format beta emin prec y -> round_round_eq (FTZ_exp emin prec) (FTZ_exp emin' prec') choice1 choice2 (x + y)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y

round_round_eq (FTZ_exp emin prec) (FTZ_exp emin' prec') choice1 choice2 (x + y)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y

Valid_exp (FTZ_exp emin prec)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y
Valid_exp (FTZ_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y
round_round_plus_hyp (FTZ_exp emin prec) (FTZ_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y
generic_format beta (FTZ_exp emin prec) x
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y
generic_format beta (FTZ_exp emin prec) y
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y

Valid_exp (FTZ_exp emin prec)
now apply FTZ_exp_valid.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y

Valid_exp (FTZ_exp emin' prec')
now apply FTZ_exp_valid.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y

round_round_plus_hyp (FTZ_exp emin prec) (FTZ_exp emin' prec')
now apply FTZ_round_round_plus_hyp.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y

generic_format beta (FTZ_exp emin prec) x
now apply generic_format_FTZ.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y

generic_format beta (FTZ_exp emin prec) y
now apply generic_format_FTZ. Qed.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

forall choice1 choice2 : Z -> bool, (emin' + prec' <= emin + 1)%Z -> (2 * prec + 1 <= prec')%Z -> forall x y : R, FTZ_format beta emin prec x -> FTZ_format beta emin prec y -> round_round_eq (FTZ_exp emin prec) (FTZ_exp emin' prec') choice1 choice2 (x - y)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

forall choice1 choice2 : Z -> bool, (emin' + prec' <= emin + 1)%Z -> (2 * prec + 1 <= prec')%Z -> forall x y : R, FTZ_format beta emin prec x -> FTZ_format beta emin prec y -> round_round_eq (FTZ_exp emin prec) (FTZ_exp emin' prec') choice1 choice2 (x - y)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y

round_round_eq (FTZ_exp emin prec) (FTZ_exp emin' prec') choice1 choice2 (x - y)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y

Valid_exp (FTZ_exp emin prec)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y
Valid_exp (FTZ_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y
round_round_plus_hyp (FTZ_exp emin prec) (FTZ_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y
generic_format beta (FTZ_exp emin prec) x
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y
generic_format beta (FTZ_exp emin prec) y
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y

Valid_exp (FTZ_exp emin prec)
now apply FTZ_exp_valid.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y

Valid_exp (FTZ_exp emin' prec')
now apply FTZ_exp_valid.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y

round_round_plus_hyp (FTZ_exp emin prec) (FTZ_exp emin' prec')
now apply FTZ_round_round_plus_hyp.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y

generic_format beta (FTZ_exp emin prec) x
now apply generic_format_FTZ.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x, y:R
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y

generic_format beta (FTZ_exp emin prec) y
now apply generic_format_FTZ. Qed. End Double_round_plus_FTZ. Section Double_round_plus_radix_ge_3. Definition round_round_plus_radix_ge_3_hyp fexp1 fexp2 := (forall ex ey, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z) /\ (forall ex ey, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z) /\ (forall ex ey, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z) /\ (forall ex ey, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z). (* fexp1 (mag x) <= mag y : * addition is exact in the largest precision (fexp2). *)
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> round_round_plus_radix_ge_3_hyp fexp1 fexp2 -> forall x y : R, 0 < y -> y <= x -> (fexp1 (mag x) <= mag y)%Z -> generic_format beta fexp1 x -> generic_format beta fexp1 y -> generic_format beta fexp2 (x + y)
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> round_round_plus_radix_ge_3_hyp fexp1 fexp2 -> forall x y : R, 0 < y -> y <= x -> (fexp1 (mag x) <= mag y)%Z -> generic_format beta fexp1 x -> generic_format beta fexp1 y -> generic_format beta fexp2 (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y

generic_format beta fexp2 (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x

generic_format beta fexp2 (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Nny:0 <= y

generic_format beta fexp2 (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y <= x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Nny:0 <= y

generic_format beta fexp2 (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y <= x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Nny:0 <= y
Hle:(mag y <= fexp1 (mag x))%Z

generic_format beta fexp2 (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y <= x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Nny:0 <= y
Hgt:(fexp1 (mag x) < mag y)%Z
generic_format beta fexp2 (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y <= x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Nny:0 <= y
Hle:(mag y <= fexp1 (mag x))%Z

generic_format beta fexp2 (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y <= x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Nny:0 <= y
Hle:(mag y <= fexp1 (mag x))%Z
Lxy:mag (x + y) = mag x

generic_format beta fexp2 (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y <= x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Nny:0 <= y
Hle:(mag y <= fexp1 (mag x))%Z
Lxy:mag (x + y) = mag x

(fexp2 (mag x) <= fexp1 (mag x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y <= x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Nny:0 <= y
Hle:(mag y <= fexp1 (mag x))%Z
Lxy:mag (x + y) = mag x
(fexp2 (mag x) <= fexp1 (mag y))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y <= x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Nny:0 <= y
Hle:(mag y <= fexp1 (mag x))%Z
Lxy:mag (x + y) = mag x

(fexp2 (mag x) <= fexp1 (mag x))%Z
now apply Hexp4; omega.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y <= x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Nny:0 <= y
Hle:(mag y <= fexp1 (mag x))%Z
Lxy:mag (x + y) = mag x

(fexp2 (mag x) <= fexp1 (mag y))%Z
now apply Hexp3; omega.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y <= x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Nny:0 <= y
Hgt:(fexp1 (mag x) < mag y)%Z

generic_format beta fexp2 (x + y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y <= x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Nny:0 <= y
Hgt:(fexp1 (mag x) < mag y)%Z

(fexp2 (mag (x + y)) <= fexp1 (mag x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y <= x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Nny:0 <= y
Hgt:(fexp1 (mag x) < mag y)%Z
(fexp2 (mag (x + y)) <= fexp1 (mag y))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y <= x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Nny:0 <= y
Hgt:(fexp1 (mag x) < mag y)%Z
Lxy:mag (x + y) = mag x

(fexp2 (mag x) <= fexp1 (mag x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y <= x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Nny:0 <= y
Hgt:(fexp1 (mag x) < mag y)%Z
Lxy:mag (x + y) = (mag x + 1)%Z
(fexp2 (mag x + 1) <= fexp1 (mag x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y <= x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Nny:0 <= y
Hgt:(fexp1 (mag x) < mag y)%Z
(fexp2 (mag (x + y)) <= fexp1 (mag y))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y <= x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Nny:0 <= y
Hgt:(fexp1 (mag x) < mag y)%Z
Lxy:mag (x + y) = mag x

(fexp2 (mag x) <= fexp1 (mag x))%Z
now apply Hexp4; omega.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y <= x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Nny:0 <= y
Hgt:(fexp1 (mag x) < mag y)%Z
Lxy:mag (x + y) = (mag x + 1)%Z

(fexp2 (mag x + 1) <= fexp1 (mag x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:(mag y <= mag x)%Z
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Nny:0 <= y
Hgt:(fexp1 (mag x) < mag y)%Z
Lxy:mag (x + y) = (mag x + 1)%Z

(fexp1 (mag x + 1 - 1) + 1 <= mag x)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:(mag y <= mag x)%Z
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Nny:0 <= y
Hgt:(fexp1 (mag x) < mag y)%Z
Lxy:mag (x + y) = (mag x + 1)%Z

(fexp1 (mag x) + 1 <= mag x)%Z
omega.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y <= x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Nny:0 <= y
Hgt:(fexp1 (mag x) < mag y)%Z

(fexp2 (mag (x + y)) <= fexp1 (mag y))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y <= x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Nny:0 <= y
Hgt:(fexp1 (mag x) < mag y)%Z
Lxy:mag (x + y) = mag x

(fexp2 (mag x) <= fexp1 (mag y))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y <= x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Nny:0 <= y
Hgt:(fexp1 (mag x) < mag y)%Z
Lxy:mag (x + y) = (mag x + 1)%Z
(fexp2 (mag x + 1) <= fexp1 (mag y))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y <= x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Nny:0 <= y
Hgt:(fexp1 (mag x) < mag y)%Z
Lxy:mag (x + y) = mag x

(fexp2 (mag x) <= fexp1 (mag y))%Z
now apply Hexp3; omega.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y <= x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Nny:0 <= y
Hgt:(fexp1 (mag x) < mag y)%Z
Lxy:mag (x + y) = (mag x + 1)%Z

(fexp2 (mag x + 1) <= fexp1 (mag y))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y <= x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Nny:0 <= y
Hgt:(fexp1 (mag x) < mag y)%Z
Lxy:mag (x + y) = (mag x + 1)%Z

(fexp1 (mag x + 1 - 1) + 1 <= mag y)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y <= x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Nny:0 <= y
Hgt:(fexp1 (mag x) < mag y)%Z
Lxy:mag (x + y) = (mag x + 1)%Z

(fexp1 (mag x) + 1 <= mag y)%Z
omega. Qed. (* mag y <= fexp1 (mag x) - 1 : round_round_lt_mid applies. *)
beta:radix

(3 <= beta)%Z -> forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall choice1 choice2 : Z -> bool, round_round_plus_radix_ge_3_hyp fexp1 fexp2 -> forall x y : R, 0 < x -> 0 < y -> (mag y <= fexp1 (mag x) - 1)%Z -> generic_format beta fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 (x + y)
beta:radix

(3 <= beta)%Z -> forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall choice1 choice2 : Z -> bool, round_round_plus_radix_ge_3_hyp fexp1 fexp2 -> forall x y : R, 0 < x -> 0 < y -> (mag y <= fexp1 (mag x) - 1)%Z -> generic_format beta fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x

round_round_eq fexp1 fexp2 choice1 choice2 (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x

round_round_eq fexp1 fexp2 choice1 choice2 (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x

round_round_eq fexp1 fexp2 choice1 choice2 (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z

round_round_eq fexp1 fexp2 choice1 choice2 (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z

bpow (-1) <= / 3
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow3:bpow (-1) <= / 3
round_round_eq fexp1 fexp2 choice1 choice2 (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z

bpow (-1) <= / 3
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z

/ IZR (beta * 1) <= / 3
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z

/ IZR beta <= / 3
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z

3 <= IZR beta
now apply IZR_le.
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow3:bpow (-1) <= / 3

round_round_eq fexp1 fexp2 choice1 choice2 (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow3:bpow (-1) <= / 3
P1:(0 < 1)%Z

round_round_eq fexp1 fexp2 choice1 choice2 (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow3:bpow (-1) <= / 3
P1:(0 < 1)%Z

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow3:bpow (-1) <= / 3
P1:(0 < 1)%Z

Valid_exp fexp1
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow3:bpow (-1) <= / 3
P1:(0 < 1)%Z
Valid_exp fexp2
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow3:bpow (-1) <= / 3
P1:(0 < 1)%Z
0 < x + y
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow3:bpow (-1) <= / 3
P1:(0 < 1)%Z
(fexp2 (mag (x + y)) <= fexp1 (mag (x + y)))%Z
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow3:bpow (-1) <= / 3
P1:(0 < 1)%Z
(fexp1 (mag (x + y)) <= mag (x + y))%Z
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow3:bpow (-1) <= / 3
P1:(0 < 1)%Z
x + y < midp fexp1 (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow3:bpow (-1) <= / 3
P1:(0 < 1)%Z
(fexp2 (mag (x + y)) <= fexp1 (mag (x + y)) - 1)%Z -> x + y < midp fexp1 (x + y) - / 2 * ulp beta fexp2 (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow3:bpow (-1) <= / 3
P1:(0 < 1)%Z

Valid_exp fexp1
exact Vfexp1.
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow3:bpow (-1) <= / 3
P1:(0 < 1)%Z

Valid_exp fexp2
exact Vfexp2.
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow3:bpow (-1) <= / 3
P1:(0 < 1)%Z

0 < x + y
lra.
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow3:bpow (-1) <= / 3
P1:(0 < 1)%Z

(fexp2 (mag (x + y)) <= fexp1 (mag (x + y)))%Z
now rewrite Lxy.
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow3:bpow (-1) <= / 3
P1:(0 < 1)%Z

(fexp1 (mag (x + y)) <= mag (x + y))%Z
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow3:bpow (-1) <= / 3
P1:(0 < 1)%Z

(fexp1 (mag x) <= mag x)%Z
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow3:bpow (-1) <= / 3
P1:(0 < 1)%Z

(fexp1 (mag x) < mag x)%Z
now apply mag_generic_gt; [|apply Rgt_not_eq|].
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow3:bpow (-1) <= / 3
P1:(0 < 1)%Z

x + y < midp fexp1 (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow3:bpow (-1) <= / 3
P1:(0 < 1)%Z

x + y < round beta fexp1 Zfloor (x + y) + / 2 * ulp beta fexp1 (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow3:bpow (-1) <= / 3
P1:(0 < 1)%Z

x + y + - round beta fexp1 Zfloor (x + y) < round beta fexp1 Zfloor (x + y) + / 2 * ulp beta fexp1 (x + y) + - round beta fexp1 Zfloor (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow3:bpow (-1) <= / 3
P1:(0 < 1)%Z

bpow (fexp1 (mag x) - 1) <= round beta fexp1 Zfloor (x + y) + / 2 * ulp beta fexp1 (x + y) + - round beta fexp1 Zfloor (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow3:bpow (-1) <= / 3
P1:(0 < 1)%Z

bpow (fexp1 (mag x) - 1) <= / 2 * ulp beta fexp1 (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow3:bpow (-1) <= / 3
P1:(0 < 1)%Z

bpow (fexp1 (mag x) - 1) <= / 2 * bpow (cexp beta fexp1 (x + y))
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow3:bpow (-1) <= / 3
P1:(0 < 1)%Z

bpow (fexp1 (mag x) - 1) <= / 2 * bpow (fexp1 (mag x))
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow3:bpow (-1) <= / 3
P1:(0 < 1)%Z

bpow (fexp1 (mag x) - 1) * bpow (- fexp1 (mag x)) <= / 2 * bpow (fexp1 (mag x)) * bpow (- fexp1 (mag x))
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow3:bpow (-1) <= / 3
P1:(0 < 1)%Z

bpow (- (1)) <= / 2
apply (Rle_trans _ _ _ Bpow3); lra.
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow3:bpow (-1) <= / 3
P1:(0 < 1)%Z

(fexp2 (mag (x + y)) <= fexp1 (mag (x + y)) - 1)%Z -> x + y < midp fexp1 (x + y) - / 2 * ulp beta fexp2 (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow3:bpow (-1) <= / 3
P1:(0 < 1)%Z

(fexp2 (mag (x + y)) <= fexp1 (mag (x + y)) - 1)%Z -> x + y < midp fexp1 (x + y) - / 2 * bpow (cexp beta fexp2 (x + y))
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow3:bpow (-1) <= / 3
P1:(0 < 1)%Z

(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z -> x + y < midp fexp1 (x + y) - / 2 * bpow (fexp2 (mag x))
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow3:bpow (-1) <= / 3
P1:(0 < 1)%Z
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z

x + y < midp fexp1 (x + y) - / 2 * bpow (fexp2 (mag x))
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow3:bpow (-1) <= / 3
P1:(0 < 1)%Z
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z

x + y < round beta fexp1 Zfloor (x + y) + / 2 * ulp beta fexp1 (x + y) - / 2 * bpow (fexp2 (mag x))
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow3:bpow (-1) <= / 3
P1:(0 < 1)%Z
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z

x + y - round beta fexp1 Zfloor (x + y) < / 2 * ulp beta fexp1 (x + y) - / 2 * bpow (fexp2 (mag x))
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow3:bpow (-1) <= / 3
P1:(0 < 1)%Z
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z

x + y - round beta fexp1 Zfloor (x + y) < / 2 * (ulp beta fexp1 (x + y) - bpow (fexp2 (mag x)))
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow3:bpow (-1) <= / 3
P1:(0 < 1)%Z
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z

bpow (fexp1 (mag x) - 1) <= / 2 * (ulp beta fexp1 (x + y) - bpow (fexp2 (mag x)))
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow3:bpow (-1) <= / 3
P1:(0 < 1)%Z
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z

bpow (fexp1 (mag x) - 1) <= / 2 * (bpow (cexp beta fexp1 (x + y)) - bpow (fexp2 (mag x)))
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow3:bpow (-1) <= / 3
P1:(0 < 1)%Z
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z

bpow (fexp1 (mag x) - 1) <= / 2 * (bpow (fexp1 (mag x)) - bpow (fexp2 (mag x)))
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow3:bpow (-1) <= / 3
P1:(0 < 1)%Z
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z

bpow (fexp1 (mag x) - 1) * bpow (- fexp1 (mag x)) <= / 2 * (bpow (fexp1 (mag x)) - bpow (fexp2 (mag x))) * bpow (- fexp1 (mag x))
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow3:bpow (-1) <= / 3
P1:(0 < 1)%Z
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z

bpow (fexp1 (mag x) - 1) * bpow (- fexp1 (mag x)) <= / 2 * ((bpow (fexp1 (mag x)) - bpow (fexp2 (mag x))) * bpow (- fexp1 (mag x)))
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow3:bpow (-1) <= / 3
P1:(0 < 1)%Z
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z

bpow (fexp1 (mag x) - 1) * bpow (- fexp1 (mag x)) <= / 2 * (bpow (fexp1 (mag x)) * bpow (- fexp1 (mag x)) - bpow (fexp2 (mag x)) * bpow (- fexp1 (mag x)))
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow3:bpow (-1) <= / 3
P1:(0 < 1)%Z
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z

bpow (- (1)) <= / 2 * (1 - bpow (fexp2 (mag x) - fexp1 (mag x)))
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow3:bpow (-1) <= / 3
P1:(0 < 1)%Z
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z

/ 3 <= / 2 * (1 - bpow (fexp2 (mag x) - fexp1 (mag x)))
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow3:bpow (-1) <= / 3
P1:(0 < 1)%Z
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z

/ 2 * (2 / 3) <= / 2 * (1 - bpow (fexp2 (mag x) - fexp1 (mag x)))
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow3:bpow (-1) <= / 3
P1:(0 < 1)%Z
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z

2 / 3 <= 1 - bpow (fexp2 (mag x) - fexp1 (mag x))
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow3:bpow (-1) <= / 3
P1:(0 < 1)%Z
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z

2 / 3 - 1 <= - bpow (fexp2 (mag x) - fexp1 (mag x))
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow3:bpow (-1) <= / 3
P1:(0 < 1)%Z
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z

- / 3 <= - bpow (fexp2 (mag x) - fexp1 (mag x))
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Px:0 < x
Py:0 < y
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Fx:generic_format beta fexp1 x
Lxy:mag (x + y) = mag x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Bpow3:bpow (-1) <= / 3
P1:(0 < 1)%Z
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z

bpow (fexp2 (mag x) - fexp1 (mag x)) <= / 3
now apply Rle_trans with (bpow (- 1)); [apply bpow_le; omega|]. Qed. (* round_round_plus_radix_ge_3_aux{0,1} together *)
beta:radix

(3 <= beta)%Z -> forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall choice1 choice2 : Z -> bool, round_round_plus_radix_ge_3_hyp fexp1 fexp2 -> forall x y : R, 0 < y -> y <= x -> generic_format beta fexp1 x -> generic_format beta fexp1 y -> round_round_eq fexp1 fexp2 choice1 choice2 (x + y)
beta:radix

(3 <= beta)%Z -> forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall choice1 choice2 : Z -> bool, round_round_plus_radix_ge_3_hyp fexp1 fexp2 -> forall x y : R, 0 < y -> y <= x -> generic_format beta fexp1 x -> generic_format beta fexp1 y -> round_round_eq fexp1 fexp2 choice1 choice2 (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y

round_round_eq fexp1 fexp2 choice1 choice2 (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x

round_round_eq fexp1 fexp2 choice1 choice2 (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hly:(fexp1 (mag x) - 1 < mag y)%Z
round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
now apply round_round_plus_radix_ge_3_aux1.
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hly:(fexp1 (mag x) - 1 < mag y)%Z

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hly:(fexp1 (mag x) - 1 < mag y)%Z

round beta fexp1 (Znearest choice1) (x + y) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hly:(fexp1 (mag x) - 1 < mag y)%Z
Valid_rnd (Znearest choice2)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hly:(fexp1 (mag x) - 1 < mag y)%Z
generic_format beta fexp2 (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hly:(fexp1 (mag x) - 1 < mag y)%Z

round beta fexp1 (Znearest choice1) (x + y) = round beta fexp1 (Znearest choice1) (x + y)
reflexivity.
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hly:(fexp1 (mag x) - 1 < mag y)%Z

Valid_rnd (Znearest choice2)
now apply valid_rnd_N.
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hly:(fexp1 (mag x) - 1 < mag y)%Z

generic_format beta fexp2 (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hly:(fexp1 (mag x) - 1 < mag y)%Z
Hf1:(fexp1 (mag x) <= mag y)%Z

generic_format beta fexp2 (x + y)
now apply (round_round_plus_radix_ge_3_aux0 fexp1). Qed.
beta:radix

(3 <= beta)%Z -> forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall choice1 choice2 : Z -> bool, round_round_plus_radix_ge_3_hyp fexp1 fexp2 -> forall x y : R, 0 <= x -> 0 <= y -> generic_format beta fexp1 x -> generic_format beta fexp1 y -> round_round_eq fexp1 fexp2 choice1 choice2 (x + y)
beta:radix

(3 <= beta)%Z -> forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall choice1 choice2 : Z -> bool, round_round_plus_radix_ge_3_hyp fexp1 fexp2 -> forall x y : R, 0 <= x -> 0 <= y -> generic_format beta fexp1 x -> generic_format beta fexp1 y -> round_round_eq fexp1 fexp2 choice1 choice2 (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y

round_round_eq fexp1 fexp2 choice1 choice2 (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) y) = round beta fexp1 (Znearest choice1) y
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0

round beta fexp1 (Znearest choice1) y = round beta fexp1 (Znearest choice1) y
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0
Valid_rnd (Znearest choice2)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0
generic_format beta fexp2 y
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0

round beta fexp1 (Znearest choice1) y = round beta fexp1 (Znearest choice1) y
reflexivity.
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0

Valid_rnd (Znearest choice2)
now apply valid_rnd_N.
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0

generic_format beta fexp2 y
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0

y <> 0 -> (fexp2 (mag y) <= fexp1 (mag y))%Z
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0
generic_format beta fexp1 y
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0

generic_format beta fexp1 y
exact Fy.
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Nzy:y <> 0
round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0

round beta fexp1 (Znearest choice1) x = round beta fexp1 (Znearest choice1) x
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0
Valid_rnd (Znearest choice2)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0
generic_format beta fexp2 x
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0

round beta fexp1 (Znearest choice1) x = round beta fexp1 (Znearest choice1) x
reflexivity.
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0

Valid_rnd (Znearest choice2)
now apply valid_rnd_N.
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0

generic_format beta fexp2 x
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0

x <> 0 -> (fexp2 (mag x) <= fexp1 (mag x))%Z
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0
generic_format beta fexp1 x
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0

generic_format beta fexp1 x
exact Fx.
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Nzy:y <> 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Nzy:y <> 0
Px:0 < x

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Nzy:y <> 0
Px:0 < x
Py:0 < y

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Nzy:y <> 0
Px:0 < x
Py:0 < y
H:x < y

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Nzy:y <> 0
Px:0 < x
Py:0 < y
H:y <= x
round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Nzy:y <> 0
Px:0 < x
Py:0 < y
H:x < y

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Nzy:y <> 0
Px:0 < x
Py:0 < y
H:x <= y

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Nzy:y <> 0
Px:0 < x
Py:0 < y
H:x <= y

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (y + x)) = round beta fexp1 (Znearest choice1) (y + x)
now apply round_round_plus_radix_ge_3_aux2.
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Nzy:y <> 0
Px:0 < x
Py:0 < y
H:y <= x

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
now apply round_round_plus_radix_ge_3_aux2. Qed. (* fexp1 (mag x) <= mag y : * substraction is exact in the largest precision (fexp2). *)
beta:radix

forall fexp1 fexp2 : Z -> Z, round_round_plus_radix_ge_3_hyp fexp1 fexp2 -> forall x y : R, 0 < y -> y < x -> (fexp1 (mag x) <= mag y)%Z -> generic_format beta fexp1 x -> generic_format beta fexp1 y -> generic_format beta fexp2 (x - y)
beta:radix

forall fexp1 fexp2 : Z -> Z, round_round_plus_radix_ge_3_hyp fexp1 fexp2 -> forall x y : R, 0 < y -> y < x -> (fexp1 (mag x) <= mag y)%Z -> generic_format beta fexp1 x -> generic_format beta fexp1 y -> generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y

generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x

generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x

generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z

generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hlt:(mag x - 2 < mag y)%Z

generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hge:(mag y <= mag x - 2)%Z
generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hlt:(mag x - 2 < mag y)%Z

generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hlt:(mag x - 2 < mag y)%Z
Hor:mag y = mag x \/ mag y = (mag x - 1)%Z

generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hlt:(mag x - 2 < mag y)%Z
Heq:mag y = mag x

generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hlt:(mag x - 2 < mag y)%Z
Heqm1:mag y = (mag x - 1)%Z
generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hlt:(mag x - 2 < mag y)%Z
Heq:mag y = mag x

generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hlt:(mag x - 2 < mag y)%Z
Heq:mag y = mag x

(fexp2 (mag (x - y)) <= fexp1 (mag x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hlt:(mag x - 2 < mag y)%Z
Heq:mag y = mag x
(fexp2 (mag (x - y)) <= fexp1 (mag y))%Z
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hlt:(mag x - 2 < mag y)%Z
Heq:mag y = mag x

(fexp2 (mag (x - y)) <= fexp1 (mag x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hlt:(mag x - 2 < mag y)%Z
Heq:mag y = mag x

(mag (x - y) - 1 <= mag x)%Z
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hlt:(mag x - 2 < mag y)%Z
Heq:mag y = mag x

(mag (x - y) <= mag x)%Z
now apply mag_minus.
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hlt:(mag x - 2 < mag y)%Z
Heq:mag y = mag x

(fexp2 (mag (x - y)) <= fexp1 (mag y))%Z
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hlt:(mag x - 2 < mag y)%Z
Heq:mag y = mag x

(fexp2 (mag (x - y)) <= fexp1 (mag x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hlt:(mag x - 2 < mag y)%Z
Heq:mag y = mag x

(mag (x - y) - 1 <= mag x)%Z
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hlt:(mag x - 2 < mag y)%Z
Heq:mag y = mag x

(mag (x - y) <= mag x)%Z
now apply mag_minus.
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hlt:(mag x - 2 < mag y)%Z
Heqm1:mag y = (mag x - 1)%Z

generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hlt:(mag x - 2 < mag y)%Z
Heqm1:mag y = (mag x - 1)%Z

(fexp2 (mag (x - y)) <= fexp1 (mag x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hlt:(mag x - 2 < mag y)%Z
Heqm1:mag y = (mag x - 1)%Z
(fexp2 (mag (x - y)) <= fexp1 (mag y))%Z
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hlt:(mag x - 2 < mag y)%Z
Heqm1:mag y = (mag x - 1)%Z

(fexp2 (mag (x - y)) <= fexp1 (mag x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hlt:(mag x - 2 < mag y)%Z
Heqm1:mag y = (mag x - 1)%Z

(mag (x - y) - 1 <= mag x)%Z
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hlt:(mag x - 2 < mag y)%Z
Heqm1:mag y = (mag x - 1)%Z

(mag (x - y) <= mag x)%Z
now apply mag_minus.
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hlt:(mag x - 2 < mag y)%Z
Heqm1:mag y = (mag x - 1)%Z

(fexp2 (mag (x - y)) <= fexp1 (mag y))%Z
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hlt:(mag x - 2 < mag y)%Z
Heqm1:mag y = (mag x - 1)%Z

(fexp2 (mag (x - y)) <= fexp1 (mag x - 1))%Z
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hlt:(mag x - 2 < mag y)%Z
Heqm1:mag y = (mag x - 1)%Z

(mag (x - y) - 1 <= mag x - 1)%Z
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hlt:(mag x - 2 < mag y)%Z
Heqm1:mag y = (mag x - 1)%Z

(mag (x - y) <= mag x)%Z
now apply mag_minus.
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hge:(mag y <= mag x - 2)%Z

generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hge:(mag y <= mag x - 2)%Z
Lxmy:mag (x - y) = mag x

generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hge:(mag y <= mag x - 2)%Z
Lxmy:mag (x - y) = (mag x - 1)%Z
generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hge:(mag y <= mag x - 2)%Z
Lxmy:mag (x - y) = mag x

generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hge:(mag y <= mag x - 2)%Z
Lxmy:mag (x - y) = mag x

(fexp2 (mag (x - y)) <= fexp1 (mag x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hge:(mag y <= mag x - 2)%Z
Lxmy:mag (x - y) = mag x
(fexp2 (mag (x - y)) <= fexp1 (mag y))%Z
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hge:(mag y <= mag x - 2)%Z
Lxmy:mag (x - y) = mag x

(fexp2 (mag (x - y)) <= fexp1 (mag x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hge:(mag y <= mag x - 2)%Z
Lxmy:mag (x - y) = mag x

(mag (x - y) - 1 <= mag x)%Z
omega.
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hge:(mag y <= mag x - 2)%Z
Lxmy:mag (x - y) = mag x

(fexp2 (mag (x - y)) <= fexp1 (mag y))%Z
now rewrite Lxmy; apply Hexp3.
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hge:(mag y <= mag x - 2)%Z
Lxmy:mag (x - y) = (mag x - 1)%Z

generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hge:(mag y <= mag x - 2)%Z
Lxmy:mag (x - y) = (mag x - 1)%Z

(fexp2 (mag x - 1) <= fexp1 (mag x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hge:(mag y <= mag x - 2)%Z
Lxmy:mag (x - y) = (mag x - 1)%Z
(fexp2 (mag x - 1) <= fexp1 (mag y))%Z
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hge:(mag y <= mag x - 2)%Z
Lxmy:mag (x - y) = (mag x - 1)%Z

(fexp2 (mag x - 1) <= fexp1 (mag x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hge:(mag y <= mag x - 2)%Z
Lxmy:mag (x - y) = (mag x - 1)%Z

(fexp1 (mag x - 1 + 1) <= mag x)%Z
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hge:(mag y <= mag x - 2)%Z
Lxmy:mag (x - y) = (mag x - 1)%Z

(fexp1 (mag x) <= mag x)%Z
now apply Z.le_trans with (mag y).
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hge:(mag y <= mag x - 2)%Z
Lxmy:mag (x - y) = (mag x - 1)%Z

(fexp2 (mag x - 1) <= fexp1 (mag y))%Z
beta:radix
fexp1, fexp2:Z -> Z
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(fexp1 (mag x) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hge:(mag y <= mag x - 2)%Z
Lxmy:mag (x - y) = (mag x - 1)%Z

(fexp1 (mag x - 1 + 1) <= mag y)%Z
now replace (_ + _)%Z with (mag x : Z); [|ring]. Qed. (* mag y <= fexp1 (mag x) - 1, * fexp1 (mag (x - y)) <= mag y : * substraction is exact in the largest precision (fexp2). *)
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> round_round_plus_radix_ge_3_hyp fexp1 fexp2 -> forall x y : R, 0 < y -> y < x -> (mag y <= fexp1 (mag x) - 1)%Z -> (fexp1 (mag (x - y)) <= mag y)%Z -> generic_format beta fexp1 x -> generic_format beta fexp1 y -> generic_format beta fexp2 (x - y)
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> round_round_plus_radix_ge_3_hyp fexp1 fexp2 -> forall x y : R, 0 < y -> y < x -> (mag y <= fexp1 (mag x) - 1)%Z -> (fexp1 (mag (x - y)) <= mag y)%Z -> generic_format beta fexp1 x -> generic_format beta fexp1 y -> generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y < x
Hln:(mag y <= fexp1 (mag x) - 1)%Z
Hln':(fexp1 (mag (x - y)) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y

generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y < x
Hln:(mag y <= fexp1 (mag x) - 1)%Z
Hln':(fexp1 (mag (x - y)) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x

generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(mag y <= fexp1 (mag x) - 1)%Z
Hln':(fexp1 (mag (x - y)) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x

generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(mag y <= fexp1 (mag x) - 1)%Z
Hln':(fexp1 (mag (x - y)) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z

generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(mag y <= fexp1 (mag x) - 1)%Z
Hln':(fexp1 (mag (x - y)) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hfx:(fexp1 (mag x) < mag x)%Z

generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(mag y <= fexp1 (mag x) - 1)%Z
Hln':(fexp1 (mag (x - y)) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z

generic_format beta fexp2 (x - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(mag y <= fexp1 (mag x) - 1)%Z
Hln':(fexp1 (mag (x - y)) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z

(fexp2 (mag (x - y)) <= fexp1 (mag x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(mag y <= fexp1 (mag x) - 1)%Z
Hln':(fexp1 (mag (x - y)) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
(fexp2 (mag (x - y)) <= fexp1 (mag y))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(mag y <= fexp1 (mag x) - 1)%Z
Hln':(fexp1 (mag (x - y)) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z

(fexp2 (mag (x - y)) <= fexp1 (mag x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(mag y <= fexp1 (mag x) - 1)%Z
Hln':(fexp1 (mag (x - y)) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z

(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(mag y <= fexp1 (mag x) - 1)%Z
Hln':(fexp1 (mag (x - y)) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
(fexp1 (mag (x - y)) <= fexp1 (mag x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(mag y <= fexp1 (mag x) - 1)%Z
Hln':(fexp1 (mag (x - y)) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z

(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)))%Z
apply Hexp4; omega.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(mag y <= fexp1 (mag x) - 1)%Z
Hln':(fexp1 (mag (x - y)) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z

(fexp1 (mag (x - y)) <= fexp1 (mag x))%Z
omega.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp1:forall ex ey : Z, (fexp1 (ex + 1) <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp2:forall ex ey : Z, (fexp1 (ex - 1) + 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp3:forall ex ey : Z, (fexp1 ex <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hyx:y < x
Hln:(mag y <= fexp1 (mag x) - 1)%Z
Hln':(fexp1 (mag (x - y)) <= mag y)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Lyx:(mag y <= mag x)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z

(fexp2 (mag (x - y)) <= fexp1 (mag y))%Z
now apply Hexp3. Qed. (* mag y <= fexp1 (mag x) - 1 : * mag y <= fexp1 (mag (x - y)) - 1 : * round_round_gt_mid applies. *)
beta:radix

(3 <= beta)%Z -> forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall choice1 choice2 : Z -> bool, round_round_plus_radix_ge_3_hyp fexp1 fexp2 -> forall x y : R, 0 < y -> y < x -> (mag y <= fexp1 (mag x) - 1)%Z -> (mag y <= fexp1 (mag (x - y)) - 1)%Z -> generic_format beta fexp1 x -> generic_format beta fexp1 y -> round_round_eq fexp1 fexp2 choice1 choice2 (x - y)
beta:radix

(3 <= beta)%Z -> forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall choice1 choice2 : Z -> bool, round_round_plus_radix_ge_3_hyp fexp1 fexp2 -> forall x y : R, 0 < y -> y < x -> (mag y <= fexp1 (mag x) - 1)%Z -> (mag y <= fexp1 (mag (x - y)) - 1)%Z -> generic_format beta fexp1 x -> generic_format beta fexp1 y -> round_round_eq fexp1 fexp2 choice1 choice2 (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y

round_round_eq fexp1 fexp2 choice1 choice2 (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x

round_round_eq fexp1 fexp2 choice1 choice2 (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x

round_round_eq fexp1 fexp2 choice1 choice2 (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z

round_round_eq fexp1 fexp2 choice1 choice2 (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z

round_round_eq fexp1 fexp2 choice1 choice2 (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z

bpow (-1) <= / 3
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
round_round_eq fexp1 fexp2 choice1 choice2 (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z

bpow (-1) <= / 3
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z

/ IZR (beta * 1) <= / 3
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z

/ IZR beta <= / 3
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z

3 <= IZR beta
now apply IZR_le.
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3

round_round_eq fexp1 fexp2 choice1 choice2 (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3

y < bpow (mag y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)
round_round_eq fexp1 fexp2 choice1 choice2 (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3

y < bpow (mag y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3

Rabs y < bpow (mag y)
apply bpow_mag_gt.
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)

round_round_eq fexp1 fexp2 choice1 choice2 (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)

Valid_exp fexp1
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)
Valid_exp fexp2
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)
0 < x - y
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)
(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)))%Z
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)
(fexp1 (mag (x - y)) <= mag (x - y))%Z
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)
midp' fexp1 (x - y) < x - y
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)
(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z -> midp' fexp1 (x - y) + / 2 * ulp beta fexp2 (x - y) < x - y
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)

Valid_exp fexp1
exact Vfexp1.
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)

Valid_exp fexp2
exact Vfexp2.
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)

0 < x - y
lra.
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)

(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)))%Z
apply Hexp4; omega.
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)

(fexp1 (mag (x - y)) <= mag (x - y))%Z
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)

(fexp1 (mag (x - y)) < mag (x - y))%Z
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)

(fexp1 (mag x - 1) < mag x - 1)%Z
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)
(mag x - 1 <= mag (x - y))%Z
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)

(fexp1 (mag x - 1) < mag x - 1)%Z
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)

(fexp1 (mag y) < mag y)%Z
now apply mag_generic_gt; [|apply Rgt_not_eq|].
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)

(mag x - 1 <= mag (x - y))%Z
now apply mag_minus_lb; [| |omega].
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)

midp' fexp1 (x - y) < x - y
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)

round beta fexp1 Zceil (x - y) - / 2 * ulp beta fexp1 (x - y) < x - y
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)

round beta fexp1 Zceil (x - y) - / 2 * ulp beta fexp1 (x - y) + (/ 2 * ulp beta fexp1 (x - y) - (x - y)) < x - y + (/ 2 * ulp beta fexp1 (x - y) - (x - y))
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)

round beta fexp1 Zceil (x - y) - x + y < / 2 * ulp beta fexp1 (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)

round beta fexp1 Zceil (x - y) - (x - y) < / 2 * ulp beta fexp1 (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)

round beta fexp1 Zceil (x - y) - (x - y) < bpow (fexp1 (mag (x - y)) - 1)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)
bpow (fexp1 (mag (x - y)) - 1) <= / 2 * ulp beta fexp1 (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)

round beta fexp1 Zceil (x - y) - (x - y) < bpow (fexp1 (mag (x - y)) - 1)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)

y < bpow (fexp1 (mag (x - y)) - 1)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)

bpow (mag y) <= bpow (fexp1 (mag (x - y)) - 1)
now apply bpow_le.
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)

bpow (fexp1 (mag (x - y)) - 1) <= / 2 * ulp beta fexp1 (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)

bpow (fexp1 (mag (x - y)) - 1) <= / 2 * bpow (cexp beta fexp1 (x - y))
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)

bpow (fexp1 (mag (x - y)) - 1) <= / 2 * bpow (fexp1 (mag (x - y)))
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)

bpow (fexp1 (mag (x - y))) * bpow (- (1)) <= / 2 * bpow (fexp1 (mag (x - y)))
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)

bpow (- (1)) * bpow (fexp1 (mag (x - y))) <= / 2 * bpow (fexp1 (mag (x - y)))
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)

bpow (- (1)) <= / 2
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)

/ IZR (beta * 1) <= / 2
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)

2 <= IZR beta
now apply IZR_le; omega.
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)

(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z -> midp' fexp1 (x - y) + / 2 * ulp beta fexp2 (x - y) < x - y
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

midp' fexp1 (x - y) + / 2 * ulp beta fexp2 (x - y) < x - y
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

round beta fexp1 Zceil (x - y) - / 2 * ulp beta fexp1 (x - y) + / 2 * ulp beta fexp2 (x - y) < x - y
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

round beta fexp1 Zceil (x - y) - / 2 * ulp beta fexp1 (x - y) + / 2 * ulp beta fexp2 (x - y) + (/ 2 * (ulp beta fexp1 (x - y) - ulp beta fexp2 (x - y)) - (x - y)) < x - y + (/ 2 * (ulp beta fexp1 (x - y) - ulp beta fexp2 (x - y)) - (x - y))
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

round beta fexp1 Zceil (x - y) - x + y < / 2 * (ulp beta fexp1 (x - y) - ulp beta fexp2 (x - y))
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

round beta fexp1 Zceil (x - y) - (x - y) < / 2 * (ulp beta fexp1 (x - y) - ulp beta fexp2 (x - y))
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

y < / 2 * (ulp beta fexp1 (x - y) - ulp beta fexp2 (x - y))
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

bpow (mag y) <= / 2 * (ulp beta fexp1 (x - y) - ulp beta fexp2 (x - y))
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

bpow (fexp1 (mag (x - y)) - 1) <= / 2 * (ulp beta fexp1 (x - y) - ulp beta fexp2 (x - y))
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

bpow (fexp1 (mag (x - y)) - 1) <= / 2 * (bpow (cexp beta fexp1 (x - y)) - bpow (cexp beta fexp2 (x - y)))
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

bpow (fexp1 (mag (x - y)) - 1) <= / 2 * (bpow (fexp1 (mag (x - y))) - bpow (fexp2 (mag (x - y))))
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

bpow (fexp1 (mag (x - y)) - 1) * bpow (- fexp1 (mag (x - y))) <= / 2 * (bpow (fexp1 (mag (x - y))) - bpow (fexp2 (mag (x - y)))) * bpow (- fexp1 (mag (x - y)))
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

bpow (fexp1 (mag (x - y)) - 1) * bpow (- fexp1 (mag (x - y))) <= / 2 * ((bpow (fexp1 (mag (x - y))) - bpow (fexp2 (mag (x - y)))) * bpow (- fexp1 (mag (x - y))))
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

bpow (fexp1 (mag (x - y)) - 1) * bpow (- fexp1 (mag (x - y))) <= / 2 * (bpow (fexp1 (mag (x - y))) * bpow (- fexp1 (mag (x - y))) - bpow (fexp2 (mag (x - y))) * bpow (- fexp1 (mag (x - y))))
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

bpow (- (1)) <= / 2 * (1 - bpow (fexp2 (mag (x - y)) - fexp1 (mag (x - y))))
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

bpow (- (1)) <= / 3
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z
/ 3 <= / 2 * (1 - bpow (fexp2 (mag (x - y)) - fexp1 (mag (x - y))))
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

bpow (- (1)) <= / 3
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

/ IZR (beta * 1) <= / 3
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

3 <= IZR beta
now apply IZR_le.
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

/ 3 <= / 2 * (1 - bpow (fexp2 (mag (x - y)) - fexp1 (mag (x - y))))
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

/ 2 * (2 / 3) <= / 2 * (1 - bpow (fexp2 (mag (x - y)) - fexp1 (mag (x - y))))
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

2 / 3 <= 1 - bpow (fexp2 (mag (x - y)) - fexp1 (mag (x - y)))
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

2 / 3 - 1 <= - bpow (fexp2 (mag (x - y)) - fexp1 (mag (x - y)))
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

- / 3 <= - bpow (fexp2 (mag (x - y)) - fexp1 (mag (x - y)))
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

bpow (fexp2 (mag (x - y)) - fexp1 (mag (x - y))) <= / 3
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

bpow (fexp2 (mag (x - y)) - fexp1 (mag (x - y))) <= bpow (-1)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z
bpow (-1) <= / 3
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

bpow (fexp2 (mag (x - y)) - fexp1 (mag (x - y))) <= bpow (-1)
apply bpow_le; omega.
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

bpow (-1) <= / 3
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

/ IZR (beta * 1) <= / 3
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Py:0 < y
Hxy:y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Bpow3:bpow (-1) <= / 3
Ly:y < bpow (mag y)
Hf2':(fexp2 (mag (x - y)) <= fexp1 (mag (x - y)) - 1)%Z

3 <= IZR beta
now apply IZR_le. Qed. (* round_round_minus_aux{0,1,2} together *)
beta:radix

(3 <= beta)%Z -> forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall choice1 choice2 : Z -> bool, round_round_plus_radix_ge_3_hyp fexp1 fexp2 -> forall x y : R, 0 < y -> y <= x -> generic_format beta fexp1 x -> generic_format beta fexp1 y -> round_round_eq fexp1 fexp2 choice1 choice2 (x - y)
beta:radix

(3 <= beta)%Z -> forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall choice1 choice2 : Z -> bool, round_round_plus_radix_ge_3_hyp fexp1 fexp2 -> forall x y : R, 0 < y -> y <= x -> generic_format beta fexp1 x -> generic_format beta fexp1 y -> round_round_eq fexp1 fexp2 choice1 choice2 (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y

round_round_eq fexp1 fexp2 choice1 choice2 (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x

round_round_eq fexp1 fexp2 choice1 choice2 (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y = x

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y <> x
round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y = x

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y = x

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) 0) = round beta fexp1 (Znearest choice1) 0
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y = x

round beta fexp1 (Znearest choice1) 0 = round beta fexp1 (Znearest choice1) 0
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y = x
Valid_rnd (Znearest choice2)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y = x

round beta fexp1 (Znearest choice1) 0 = round beta fexp1 (Znearest choice1) 0
reflexivity.
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y = x

Valid_rnd (Znearest choice2)
now apply valid_rnd_N.
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y <> x

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y <> x
Hyx':y < x

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y <> x
Hyx':y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y <> x
Hyx':y < x
Hly:(fexp1 (mag x) - 1 < mag y)%Z
round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y <> x
Hyx':y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y <> x
Hyx':y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y <> x
Hyx':y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(fexp1 (mag (x - y)) - 1 < mag y)%Z
round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y <> x
Hyx':y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(mag y <= fexp1 (mag (x - y)) - 1)%Z

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
now apply round_round_minus_radix_ge_3_aux2.
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y <> x
Hyx':y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(fexp1 (mag (x - y)) - 1 < mag y)%Z

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y <> x
Hyx':y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(fexp1 (mag (x - y)) - 1 < mag y)%Z

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y <> x
Hyx':y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(fexp1 (mag (x - y)) - 1 < mag y)%Z

round beta fexp1 (Znearest choice1) (x - y) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y <> x
Hyx':y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(fexp1 (mag (x - y)) - 1 < mag y)%Z
Valid_rnd (Znearest choice2)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y <> x
Hyx':y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(fexp1 (mag (x - y)) - 1 < mag y)%Z
generic_format beta fexp2 (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y <> x
Hyx':y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(fexp1 (mag (x - y)) - 1 < mag y)%Z

round beta fexp1 (Znearest choice1) (x - y) = round beta fexp1 (Znearest choice1) (x - y)
reflexivity.
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y <> x
Hyx':y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(fexp1 (mag (x - y)) - 1 < mag y)%Z

Valid_rnd (Znearest choice2)
now apply valid_rnd_N.
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y <> x
Hyx':y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(fexp1 (mag (x - y)) - 1 < mag y)%Z

generic_format beta fexp2 (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y <> x
Hyx':y < x
Hly:(mag y <= fexp1 (mag x) - 1)%Z
Hly':(fexp1 (mag (x - y)) - 1 < mag y)%Z
Hf1:(fexp1 (mag (x - y)) <= mag y)%Z

generic_format beta fexp2 (x - y)
now apply (round_round_minus_radix_ge_3_aux1 fexp1). }
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y <> x
Hyx':y < x
Hly:(fexp1 (mag x) - 1 < mag y)%Z

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y <> x
Hyx':y < x
Hly:(fexp1 (mag x) - 1 < mag y)%Z

round beta fexp1 (Znearest choice1) (x - y) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y <> x
Hyx':y < x
Hly:(fexp1 (mag x) - 1 < mag y)%Z
Valid_rnd (Znearest choice2)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y <> x
Hyx':y < x
Hly:(fexp1 (mag x) - 1 < mag y)%Z
generic_format beta fexp2 (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y <> x
Hyx':y < x
Hly:(fexp1 (mag x) - 1 < mag y)%Z

round beta fexp1 (Znearest choice1) (x - y) = round beta fexp1 (Znearest choice1) (x - y)
reflexivity.
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y <> x
Hyx':y < x
Hly:(fexp1 (mag x) - 1 < mag y)%Z

Valid_rnd (Znearest choice2)
now apply valid_rnd_N.
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y <> x
Hyx':y < x
Hly:(fexp1 (mag x) - 1 < mag y)%Z

generic_format beta fexp2 (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Py:0 < y
Hyx:y <= x
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:0 < x
Hy:y <> x
Hyx':y < x
Hly:(fexp1 (mag x) - 1 < mag y)%Z
Hf1:(fexp1 (mag x) <= mag y)%Z

generic_format beta fexp2 (x - y)
now apply (round_round_minus_radix_ge_3_aux0 fexp1). Qed.
beta:radix

(3 <= beta)%Z -> forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall choice1 choice2 : Z -> bool, round_round_plus_radix_ge_3_hyp fexp1 fexp2 -> forall x y : R, 0 <= x -> 0 <= y -> generic_format beta fexp1 x -> generic_format beta fexp1 y -> round_round_eq fexp1 fexp2 choice1 choice2 (x - y)
beta:radix

(3 <= beta)%Z -> forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall choice1 choice2 : Z -> bool, round_round_plus_radix_ge_3_hyp fexp1 fexp2 -> forall x y : R, 0 <= x -> 0 <= y -> generic_format beta fexp1 x -> generic_format beta fexp1 y -> round_round_eq fexp1 fexp2 choice1 choice2 (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y

round_round_eq fexp1 fexp2 choice1 choice2 (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (- y)) = round beta fexp1 (Znearest choice1) (- y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0

- round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) (round beta fexp2 (Znearest (fun t : Z => negb (choice2 (- (t + 1))%Z))) y) = - round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) y
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0

- round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) y = - round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) y
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0
Valid_rnd (Znearest (fun t : Z => negb (choice2 (- (t + 1))%Z)))
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0
generic_format beta fexp2 y
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0

- round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) y = - round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) y
reflexivity.
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0

Valid_rnd (Znearest (fun t : Z => negb (choice2 (- (t + 1))%Z)))
now apply valid_rnd_N.
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0

generic_format beta fexp2 y
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0

y <> 0 -> (fexp2 (mag y) <= fexp1 (mag y))%Z
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0
generic_format beta fexp1 y
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0

y <> 0 -> (fexp2 (mag y) <= fexp1 (mag y))%Z
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0
generic_format beta fexp1 y
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0

generic_format beta fexp1 y
exact Fy.
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Nzy:y <> 0
round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0

round beta fexp1 (Znearest choice1) x = round beta fexp1 (Znearest choice1) x
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0
Valid_rnd (Znearest choice2)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0
generic_format beta fexp2 x
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0

round beta fexp1 (Znearest choice1) x = round beta fexp1 (Znearest choice1) x
reflexivity.
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0

Valid_rnd (Znearest choice2)
now apply valid_rnd_N.
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0

generic_format beta fexp2 x
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0

x <> 0 -> (fexp2 (mag x) <= fexp1 (mag x))%Z
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0
generic_format beta fexp1 x
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp4:forall ex ey : Z, (ex - 1 <= ey)%Z -> (fexp2 ex <= fexp1 ey)%Z
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0

x <> 0 -> (fexp2 (mag x) <= fexp1 (mag x))%Z
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0
generic_format beta fexp1 x
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Zy:y = 0

generic_format beta fexp1 x
exact Fx.
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Nzy:y <> 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Nzy:y <> 0
Px:0 < x

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Nzy:y <> 0
Px:0 < x
Py:0 < y

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Nzy:y <> 0
Px:0 < x
Py:0 < y
H:x < y

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Nzy:y <> 0
Px:0 < x
Py:0 < y
H:y <= x
round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Nzy:y <> 0
Px:0 < x
Py:0 < y
H:x < y

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Nzy:y <> 0
Px:0 < x
Py:0 < y
H:x <= y

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Nzy:y <> 0
Px:0 < x
Py:0 < y
H:x <= y

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (- (y - x))) = round beta fexp1 (Znearest choice1) (- (y - x))
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Nzy:y <> 0
Px:0 < x
Py:0 < y
H:x <= y

- round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) (round beta fexp2 (Znearest (fun t : Z => negb (choice2 (- (t + 1))%Z))) (y - x)) = - round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) (y - x)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Nzy:y <> 0
Px:0 < x
Py:0 < y
H:x <= y

round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) (round beta fexp2 (Znearest (fun t : Z => negb (choice2 (- (t + 1))%Z))) (y - x)) = round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) (y - x)
now apply round_round_minus_radix_ge_3_aux3.
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Nnx:0 <= x
Nny:0 <= y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nzx:x <> 0
Nzy:y <> 0
Px:0 < x
Py:0 < y
H:y <= x

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - y)) = round beta fexp1 (Znearest choice1) (x - y)
now apply round_round_minus_radix_ge_3_aux3. Qed.
beta:radix

(3 <= beta)%Z -> forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall choice1 choice2 : Z -> bool, round_round_plus_radix_ge_3_hyp fexp1 fexp2 -> forall x y : R, generic_format beta fexp1 x -> generic_format beta fexp1 y -> round_round_eq fexp1 fexp2 choice1 choice2 (x + y)
beta:radix

(3 <= beta)%Z -> forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall choice1 choice2 : Z -> bool, round_round_plus_radix_ge_3_hyp fexp1 fexp2 -> forall x y : R, generic_format beta fexp1 x -> generic_format beta fexp1 y -> round_round_eq fexp1 fexp2 choice1 choice2 (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y

round_round_eq fexp1 fexp2 choice1 choice2 (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Sx:x < 0
Sy:y < 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Sx:x < 0
Sy:0 <= y
round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Sx:0 <= x
Sy:y < 0
round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Sx:0 <= x
Sy:0 <= y
round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Sx:x < 0
Sy:y < 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Sx:x < 0
Sy:y < 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (- (- x - y))) = round beta fexp1 (Znearest choice1) (- (- x - y))
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Sx:x < 0
Sy:y < 0

- round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) (round beta fexp2 (Znearest (fun t : Z => negb (choice2 (- (t + 1))%Z))) (- x - y)) = - round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) (- x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Sx:x < 0
Sy:y < 0

round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) (round beta fexp2 (Znearest (fun t : Z => negb (choice2 (- (t + 1))%Z))) (- x - y)) = round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) (- x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Sx:x < 0
Sy:y < 0
Px:0 <= - x

round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) (round beta fexp2 (Znearest (fun t : Z => negb (choice2 (- (t + 1))%Z))) (- x - y)) = round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) (- x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Sx:x < 0
Sy:y < 0
Px:0 <= - x
Py:0 <= - y

round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) (round beta fexp2 (Znearest (fun t : Z => negb (choice2 (- (t + 1))%Z))) (- x - y)) = round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) (- x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 (- x)
Fy:generic_format beta fexp1 y
Sx:x < 0
Sy:y < 0
Px:0 <= - x
Py:0 <= - y

round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) (round beta fexp2 (Znearest (fun t : Z => negb (choice2 (- (t + 1))%Z))) (- x - y)) = round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) (- x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 (- x)
Fy:generic_format beta fexp1 (- y)
Sx:x < 0
Sy:y < 0
Px:0 <= - x
Py:0 <= - y

round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) (round beta fexp2 (Znearest (fun t : Z => negb (choice2 (- (t + 1))%Z))) (- x - y)) = round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) (- x - y)
now apply round_round_plus_radix_ge_3_aux.
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Sx:x < 0
Sy:0 <= y

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Sx:x < 0
Sy:0 <= y

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (y - - x)) = round beta fexp1 (Znearest choice1) (y - - x)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Sx:x < 0
Sy:0 <= y
Px:0 <= - x

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (y - - x)) = round beta fexp1 (Znearest choice1) (y - - x)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 (- x)
Fy:generic_format beta fexp1 y
Sx:x < 0
Sy:0 <= y
Px:0 <= - x

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (y - - x)) = round beta fexp1 (Znearest choice1) (y - - x)
now apply round_round_minus_radix_ge_3_aux.
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Sx:0 <= x
Sy:y < 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Sx:0 <= x
Sy:y < 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - - y)) = round beta fexp1 (Znearest choice1) (x - - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Sx:0 <= x
Sy:y < 0
Py:0 <= - y

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - - y)) = round beta fexp1 (Znearest choice1) (x - - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 (- y)
Sx:0 <= x
Sy:y < 0
Py:0 <= - y

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x - - y)) = round beta fexp1 (Znearest choice1) (x - - y)
now apply round_round_minus_radix_ge_3_aux.
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Sx:0 <= x
Sy:0 <= y

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x + y)) = round beta fexp1 (Znearest choice1) (x + y)
now apply round_round_plus_radix_ge_3_aux. Qed.
beta:radix

(3 <= beta)%Z -> forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall choice1 choice2 : Z -> bool, round_round_plus_radix_ge_3_hyp fexp1 fexp2 -> forall x y : R, generic_format beta fexp1 x -> generic_format beta fexp1 y -> round_round_eq fexp1 fexp2 choice1 choice2 (x - y)
beta:radix

(3 <= beta)%Z -> forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall choice1 choice2 : Z -> bool, round_round_plus_radix_ge_3_hyp fexp1 fexp2 -> forall x y : R, generic_format beta fexp1 x -> generic_format beta fexp1 y -> round_round_eq fexp1 fexp2 choice1 choice2 (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y

round_round_eq fexp1 fexp2 choice1 choice2 (x - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y

round_round_eq fexp1 fexp2 choice1 choice2 (x + - y)
beta:radix
Hbeta:(3 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_plus_radix_ge_3_hyp fexp1 fexp2
x, y:R
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 (- y)

round_round_eq fexp1 fexp2 choice1 choice2 (x + - y)
now apply round_round_plus_radix_ge_3. Qed. Section Double_round_plus_radix_ge_3_FLX. Variable prec : Z. Variable prec' : Z. Context { prec_gt_0_ : Prec_gt_0 prec }. Context { prec_gt_0_' : Prec_gt_0 prec' }.
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(2 * prec <= prec')%Z -> round_round_plus_radix_ge_3_hyp (FLX_exp prec) (FLX_exp prec')
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(2 * prec <= prec')%Z -> round_round_plus_radix_ge_3_hyp (FLX_exp prec) (FLX_exp prec')
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hprec:(2 * prec <= prec')%Z

round_round_plus_radix_ge_3_hyp (FLX_exp prec) (FLX_exp prec')
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hprec:(2 * prec <= prec')%Z

round_round_plus_radix_ge_3_hyp (fun e : Z => (e - prec)%Z) (fun e : Z => (e - prec')%Z)
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hprec:(2 * prec <= prec')%Z
ex, ey:Z

(ex - 1 <= ey)%Z -> (ex - prec' <= ey - prec)%Z
beta:radix
prec, prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hprec:(2 * prec <= prec')%Z
ex, ey:Z

(ex - 1 <= ey)%Z -> (ex - prec' <= ey - prec)%Z
omega. Qed.
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(3 <= beta)%Z -> forall choice1 choice2 : Z -> bool, (2 * prec <= prec')%Z -> forall x y : R, FLX_format beta prec x -> FLX_format beta prec y -> round_round_eq (FLX_exp prec) (FLX_exp prec') choice1 choice2 (x + y)
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(3 <= beta)%Z -> forall choice1 choice2 : Z -> bool, (2 * prec <= prec')%Z -> forall x y : R, FLX_format beta prec x -> FLX_format beta prec y -> round_round_eq (FLX_exp prec) (FLX_exp prec') choice1 choice2 (x + y)
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y

round_round_eq (FLX_exp prec) (FLX_exp prec') choice1 choice2 (x + y)
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y

(3 <= beta)%Z
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y
Valid_exp (FLX_exp prec)
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y
Valid_exp (FLX_exp prec')
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y
round_round_plus_radix_ge_3_hyp (FLX_exp prec) (FLX_exp prec')
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y
generic_format beta (FLX_exp prec) x
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y
generic_format beta (FLX_exp prec) y
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y

(3 <= beta)%Z
exact Hbeta.
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y

Valid_exp (FLX_exp prec)
now apply FLX_exp_valid.
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y

Valid_exp (FLX_exp prec')
now apply FLX_exp_valid.
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y

round_round_plus_radix_ge_3_hyp (FLX_exp prec) (FLX_exp prec')
now apply FLX_round_round_plus_radix_ge_3_hyp.
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y

generic_format beta (FLX_exp prec) x
now apply generic_format_FLX.
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y

generic_format beta (FLX_exp prec) y
now apply generic_format_FLX. Qed.
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(3 <= beta)%Z -> forall choice1 choice2 : Z -> bool, (2 * prec <= prec')%Z -> forall x y : R, FLX_format beta prec x -> FLX_format beta prec y -> round_round_eq (FLX_exp prec) (FLX_exp prec') choice1 choice2 (x - y)
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(3 <= beta)%Z -> forall choice1 choice2 : Z -> bool, (2 * prec <= prec')%Z -> forall x y : R, FLX_format beta prec x -> FLX_format beta prec y -> round_round_eq (FLX_exp prec) (FLX_exp prec') choice1 choice2 (x - y)
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y

round_round_eq (FLX_exp prec) (FLX_exp prec') choice1 choice2 (x - y)
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y

(3 <= beta)%Z
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y
Valid_exp (FLX_exp prec)
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y
Valid_exp (FLX_exp prec')
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y
round_round_plus_radix_ge_3_hyp (FLX_exp prec) (FLX_exp prec')
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y
generic_format beta (FLX_exp prec) x
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y
generic_format beta (FLX_exp prec) y
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y

(3 <= beta)%Z
exact Hbeta.
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y

Valid_exp (FLX_exp prec)
now apply FLX_exp_valid.
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y

Valid_exp (FLX_exp prec')
now apply FLX_exp_valid.
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y

round_round_plus_radix_ge_3_hyp (FLX_exp prec) (FLX_exp prec')
now apply FLX_round_round_plus_radix_ge_3_hyp.
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y

generic_format beta (FLX_exp prec) x
now apply generic_format_FLX.
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y

generic_format beta (FLX_exp prec) y
now apply generic_format_FLX. Qed. End Double_round_plus_radix_ge_3_FLX. Section Double_round_plus_radix_ge_3_FLT. Variable emin prec : Z. Variable emin' prec' : Z. Context { prec_gt_0_ : Prec_gt_0 prec }. Context { prec_gt_0_' : Prec_gt_0 prec' }.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(emin' <= emin)%Z -> (2 * prec <= prec')%Z -> round_round_plus_radix_ge_3_hyp (FLT_exp emin prec) (FLT_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(emin' <= emin)%Z -> (2 * prec <= prec')%Z -> round_round_plus_radix_ge_3_hyp (FLT_exp emin prec) (FLT_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec <= prec')%Z

round_round_plus_radix_ge_3_hyp (FLT_exp emin prec) (FLT_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec <= prec')%Z

round_round_plus_radix_ge_3_hyp (fun e : Z => Z.max (e - prec) emin) (fun e : Z => Z.max (e - prec') emin')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z

(Z.max (ex + 1 - prec) emin <= ey)%Z -> (Z.max (ex - prec') emin' <= Z.max (ey - prec) emin)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z
(Z.max (ex - 1 - prec) emin + 1 <= ey)%Z -> (Z.max (ex - prec') emin' <= Z.max (ey - prec) emin)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z
(Z.max (ex - prec) emin <= ey)%Z -> (Z.max (ex - prec') emin' <= Z.max (ey - prec) emin)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z
(ex - 1 <= ey)%Z -> (Z.max (ex - prec') emin' <= Z.max (ey - prec) emin)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z

(Z.max (ex + 1 - prec) emin <= ey)%Z -> (Z.max (ex - prec') emin' <= Z.max (ey - prec) emin)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z

(ex + 1 - prec >= emin)%Z /\ Z.max (ex + 1 - prec) emin = (ex + 1 - prec)%Z \/ (ex + 1 - prec < emin)%Z /\ Z.max (ex + 1 - prec) emin = emin -> (Z.max (ex + 1 - prec) emin <= ey)%Z -> (Z.max (ex - prec') emin' <= Z.max (ey - prec) emin)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z

(ex - prec' >= emin')%Z /\ Z.max (ex - prec') emin' = (ex - prec')%Z \/ (ex - prec' < emin')%Z /\ Z.max (ex - prec') emin' = emin' -> (ex + 1 - prec >= emin)%Z /\ Z.max (ex + 1 - prec) emin = (ex + 1 - prec)%Z \/ (ex + 1 - prec < emin)%Z /\ Z.max (ex + 1 - prec) emin = emin -> (Z.max (ex + 1 - prec) emin <= ey)%Z -> (Z.max (ex - prec') emin' <= Z.max (ey - prec) emin)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z

(ey - prec >= emin)%Z /\ Z.max (ey - prec) emin = (ey - prec)%Z \/ (ey - prec < emin)%Z /\ Z.max (ey - prec) emin = emin -> (ex - prec' >= emin')%Z /\ Z.max (ex - prec') emin' = (ex - prec')%Z \/ (ex - prec' < emin')%Z /\ Z.max (ex - prec') emin' = emin' -> (ex + 1 - prec >= emin)%Z /\ Z.max (ex + 1 - prec) emin = (ex + 1 - prec)%Z \/ (ex + 1 - prec < emin)%Z /\ Z.max (ex + 1 - prec) emin = emin -> (Z.max (ex + 1 - prec) emin <= ey)%Z -> (Z.max (ex - prec') emin' <= Z.max (ey - prec) emin)%Z
omega.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z

(Z.max (ex - 1 - prec) emin + 1 <= ey)%Z -> (Z.max (ex - prec') emin' <= Z.max (ey - prec) emin)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z

(ex - 1 - prec >= emin)%Z /\ Z.max (ex - 1 - prec) emin = (ex - 1 - prec)%Z \/ (ex - 1 - prec < emin)%Z /\ Z.max (ex - 1 - prec) emin = emin -> (Z.max (ex - 1 - prec) emin + 1 <= ey)%Z -> (Z.max (ex - prec') emin' <= Z.max (ey - prec) emin)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z

(ex - prec' >= emin')%Z /\ Z.max (ex - prec') emin' = (ex - prec')%Z \/ (ex - prec' < emin')%Z /\ Z.max (ex - prec') emin' = emin' -> (ex - 1 - prec >= emin)%Z /\ Z.max (ex - 1 - prec) emin = (ex - 1 - prec)%Z \/ (ex - 1 - prec < emin)%Z /\ Z.max (ex - 1 - prec) emin = emin -> (Z.max (ex - 1 - prec) emin + 1 <= ey)%Z -> (Z.max (ex - prec') emin' <= Z.max (ey - prec) emin)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z

(ey - prec >= emin)%Z /\ Z.max (ey - prec) emin = (ey - prec)%Z \/ (ey - prec < emin)%Z /\ Z.max (ey - prec) emin = emin -> (ex - prec' >= emin')%Z /\ Z.max (ex - prec') emin' = (ex - prec')%Z \/ (ex - prec' < emin')%Z /\ Z.max (ex - prec') emin' = emin' -> (ex - 1 - prec >= emin)%Z /\ Z.max (ex - 1 - prec) emin = (ex - 1 - prec)%Z \/ (ex - 1 - prec < emin)%Z /\ Z.max (ex - 1 - prec) emin = emin -> (Z.max (ex - 1 - prec) emin + 1 <= ey)%Z -> (Z.max (ex - prec') emin' <= Z.max (ey - prec) emin)%Z
omega.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z

(Z.max (ex - prec) emin <= ey)%Z -> (Z.max (ex - prec') emin' <= Z.max (ey - prec) emin)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z

(ex - prec >= emin)%Z /\ Z.max (ex - prec) emin = (ex - prec)%Z \/ (ex - prec < emin)%Z /\ Z.max (ex - prec) emin = emin -> (Z.max (ex - prec) emin <= ey)%Z -> (Z.max (ex - prec') emin' <= Z.max (ey - prec) emin)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z

(ex - prec' >= emin')%Z /\ Z.max (ex - prec') emin' = (ex - prec')%Z \/ (ex - prec' < emin')%Z /\ Z.max (ex - prec') emin' = emin' -> (ex - prec >= emin)%Z /\ Z.max (ex - prec) emin = (ex - prec)%Z \/ (ex - prec < emin)%Z /\ Z.max (ex - prec) emin = emin -> (Z.max (ex - prec) emin <= ey)%Z -> (Z.max (ex - prec') emin' <= Z.max (ey - prec) emin)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z

(ey - prec >= emin)%Z /\ Z.max (ey - prec) emin = (ey - prec)%Z \/ (ey - prec < emin)%Z /\ Z.max (ey - prec) emin = emin -> (ex - prec' >= emin')%Z /\ Z.max (ex - prec') emin' = (ex - prec')%Z \/ (ex - prec' < emin')%Z /\ Z.max (ex - prec') emin' = emin' -> (ex - prec >= emin)%Z /\ Z.max (ex - prec) emin = (ex - prec)%Z \/ (ex - prec < emin)%Z /\ Z.max (ex - prec) emin = emin -> (Z.max (ex - prec) emin <= ey)%Z -> (Z.max (ex - prec') emin' <= Z.max (ey - prec) emin)%Z
omega.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z

(ex - 1 <= ey)%Z -> (Z.max (ex - prec') emin' <= Z.max (ey - prec) emin)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z

(ex - 1 <= ey)%Z -> (Z.max (ex - prec') emin' <= Z.max (ey - prec) emin)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z

(ex - prec' >= emin')%Z /\ Z.max (ex - prec') emin' = (ex - prec')%Z \/ (ex - prec' < emin')%Z /\ Z.max (ex - prec') emin' = emin' -> (ex - 1 <= ey)%Z -> (Z.max (ex - prec') emin' <= Z.max (ey - prec) emin)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z

(ey - prec >= emin)%Z /\ Z.max (ey - prec) emin = (ey - prec)%Z \/ (ey - prec < emin)%Z /\ Z.max (ey - prec) emin = emin -> (ex - prec' >= emin')%Z /\ Z.max (ex - prec') emin' = (ex - prec')%Z \/ (ex - prec' < emin')%Z /\ Z.max (ex - prec') emin' = emin' -> (ex - 1 <= ey)%Z -> (Z.max (ex - prec') emin' <= Z.max (ey - prec) emin)%Z
omega. Qed.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(3 <= beta)%Z -> forall choice1 choice2 : Z -> bool, (emin' <= emin)%Z -> (2 * prec <= prec')%Z -> forall x y : R, FLT_format beta emin prec x -> FLT_format beta emin prec y -> round_round_eq (FLT_exp emin prec) (FLT_exp emin' prec') choice1 choice2 (x + y)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(3 <= beta)%Z -> forall choice1 choice2 : Z -> bool, (emin' <= emin)%Z -> (2 * prec <= prec')%Z -> forall x y : R, FLT_format beta emin prec x -> FLT_format beta emin prec y -> round_round_eq (FLT_exp emin prec) (FLT_exp emin' prec') choice1 choice2 (x + y)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y

round_round_eq (FLT_exp emin prec) (FLT_exp emin' prec') choice1 choice2 (x + y)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y

(3 <= beta)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y
Valid_exp (FLT_exp emin prec)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y
Valid_exp (FLT_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y
round_round_plus_radix_ge_3_hyp (FLT_exp emin prec) (FLT_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y
generic_format beta (FLT_exp emin prec) x
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y
generic_format beta (FLT_exp emin prec) y
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y

(3 <= beta)%Z
exact Hbeta.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y

Valid_exp (FLT_exp emin prec)
now apply FLT_exp_valid.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y

Valid_exp (FLT_exp emin' prec')
now apply FLT_exp_valid.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y

round_round_plus_radix_ge_3_hyp (FLT_exp emin prec) (FLT_exp emin' prec')
now apply FLT_round_round_plus_radix_ge_3_hyp.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y

generic_format beta (FLT_exp emin prec) x
now apply generic_format_FLT.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y

generic_format beta (FLT_exp emin prec) y
now apply generic_format_FLT. Qed.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(3 <= beta)%Z -> forall choice1 choice2 : Z -> bool, (emin' <= emin)%Z -> (2 * prec <= prec')%Z -> forall x y : R, FLT_format beta emin prec x -> FLT_format beta emin prec y -> round_round_eq (FLT_exp emin prec) (FLT_exp emin' prec') choice1 choice2 (x - y)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(3 <= beta)%Z -> forall choice1 choice2 : Z -> bool, (emin' <= emin)%Z -> (2 * prec <= prec')%Z -> forall x y : R, FLT_format beta emin prec x -> FLT_format beta emin prec y -> round_round_eq (FLT_exp emin prec) (FLT_exp emin' prec') choice1 choice2 (x - y)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y

round_round_eq (FLT_exp emin prec) (FLT_exp emin' prec') choice1 choice2 (x - y)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y

(3 <= beta)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y
Valid_exp (FLT_exp emin prec)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y
Valid_exp (FLT_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y
round_round_plus_radix_ge_3_hyp (FLT_exp emin prec) (FLT_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y
generic_format beta (FLT_exp emin prec) x
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y
generic_format beta (FLT_exp emin prec) y
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y

(3 <= beta)%Z
exact Hbeta.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y

Valid_exp (FLT_exp emin prec)
now apply FLT_exp_valid.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y

Valid_exp (FLT_exp emin' prec')
now apply FLT_exp_valid.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y

round_round_plus_radix_ge_3_hyp (FLT_exp emin prec) (FLT_exp emin' prec')
now apply FLT_round_round_plus_radix_ge_3_hyp.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y

generic_format beta (FLT_exp emin prec) x
now apply generic_format_FLT.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin' <= emin)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y

generic_format beta (FLT_exp emin prec) y
now apply generic_format_FLT. Qed. End Double_round_plus_radix_ge_3_FLT. Section Double_round_plus_radix_ge_3_FTZ. Variable emin prec : Z. Variable emin' prec' : Z. Context { prec_gt_0_ : Prec_gt_0 prec }. Context { prec_gt_0_' : Prec_gt_0 prec' }.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(emin' + prec' <= emin + 1)%Z -> (2 * prec <= prec')%Z -> round_round_plus_radix_ge_3_hyp (FTZ_exp emin prec) (FTZ_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(emin' + prec' <= emin + 1)%Z -> (2 * prec <= prec')%Z -> round_round_plus_radix_ge_3_hyp (FTZ_exp emin prec) (FTZ_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec <= prec')%Z

round_round_plus_radix_ge_3_hyp (FTZ_exp emin prec) (FTZ_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec <= prec')%Z

round_round_plus_radix_ge_3_hyp (fun e : Z => if (e - prec <? emin)%Z then (emin + prec - 1)%Z else (e - prec)%Z) (fun e : Z => if (e - prec' <? emin')%Z then (emin' + prec' - 1)%Z else (e - prec')%Z)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':(0 < prec')%Z
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec <= prec')%Z

round_round_plus_radix_ge_3_hyp (fun e : Z => if (e - prec <? emin)%Z then (emin + prec - 1)%Z else (e - prec)%Z) (fun e : Z => if (e - prec' <? emin')%Z then (emin' + prec' - 1)%Z else (e - prec')%Z)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':(0 < prec')%Z
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z

((if ex + 1 - prec <? emin then emin + prec - 1 else ex + 1 - prec) <= ey)%Z -> ((if ex - prec' <? emin' then emin' + prec' - 1 else ex - prec') <= (if ey - prec <? emin then emin + prec - 1 else ey - prec))%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':(0 < prec')%Z
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z
((if ex - 1 - prec <? emin then emin + prec - 1 else ex - 1 - prec) + 1 <= ey)%Z -> ((if ex - prec' <? emin' then emin' + prec' - 1 else ex - prec') <= (if ey - prec <? emin then emin + prec - 1 else ey - prec))%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':(0 < prec')%Z
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z
((if ex - prec <? emin then emin + prec - 1 else ex - prec) <= ey)%Z -> ((if ex - prec' <? emin' then emin' + prec' - 1 else ex - prec') <= (if ey - prec <? emin then emin + prec - 1 else ey - prec))%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':(0 < prec')%Z
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z
(ex - 1 <= ey)%Z -> ((if ex - prec' <? emin' then emin' + prec' - 1 else ex - prec') <= (if ey - prec <? emin then emin + prec - 1 else ey - prec))%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':(0 < prec')%Z
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z

((if ex + 1 - prec <? emin then emin + prec - 1 else ex + 1 - prec) <= ey)%Z -> ((if ex - prec' <? emin' then emin' + prec' - 1 else ex - prec') <= (if ey - prec <? emin then emin + prec - 1 else ey - prec))%Z
destruct (Z.ltb_spec (ex + 1 - prec) emin); destruct (Z.ltb_spec (ex - prec') emin'); destruct (Z.ltb_spec (ey - prec) emin); omega.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':(0 < prec')%Z
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z

((if ex - 1 - prec <? emin then emin + prec - 1 else ex - 1 - prec) + 1 <= ey)%Z -> ((if ex - prec' <? emin' then emin' + prec' - 1 else ex - prec') <= (if ey - prec <? emin then emin + prec - 1 else ey - prec))%Z
destruct (Z.ltb_spec (ex - 1 - prec) emin); destruct (Z.ltb_spec (ex - prec') emin'); destruct (Z.ltb_spec (ey - prec) emin); omega.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':(0 < prec')%Z
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z

((if ex - prec <? emin then emin + prec - 1 else ex - prec) <= ey)%Z -> ((if ex - prec' <? emin' then emin' + prec' - 1 else ex - prec') <= (if ey - prec <? emin then emin + prec - 1 else ey - prec))%Z
destruct (Z.ltb_spec (ex - prec) emin); destruct (Z.ltb_spec (ex - prec') emin'); destruct (Z.ltb_spec (ey - prec) emin); omega.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':(0 < prec')%Z
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z

(ex - 1 <= ey)%Z -> ((if ex - prec' <? emin' then emin' + prec' - 1 else ex - prec') <= (if ey - prec <? emin then emin + prec - 1 else ey - prec))%Z
destruct (Z.ltb_spec (ex - prec') emin'); destruct (Z.ltb_spec (ey - prec) emin); omega. Qed.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(3 <= beta)%Z -> forall choice1 choice2 : Z -> bool, (emin' + prec' <= emin + 1)%Z -> (2 * prec <= prec')%Z -> forall x y : R, FTZ_format beta emin prec x -> FTZ_format beta emin prec y -> round_round_eq (FTZ_exp emin prec) (FTZ_exp emin' prec') choice1 choice2 (x + y)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(3 <= beta)%Z -> forall choice1 choice2 : Z -> bool, (emin' + prec' <= emin + 1)%Z -> (2 * prec <= prec')%Z -> forall x y : R, FTZ_format beta emin prec x -> FTZ_format beta emin prec y -> round_round_eq (FTZ_exp emin prec) (FTZ_exp emin' prec') choice1 choice2 (x + y)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y

round_round_eq (FTZ_exp emin prec) (FTZ_exp emin' prec') choice1 choice2 (x + y)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y

(3 <= beta)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y
Valid_exp (FTZ_exp emin prec)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y
Valid_exp (FTZ_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y
round_round_plus_radix_ge_3_hyp (FTZ_exp emin prec) (FTZ_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y
generic_format beta (FTZ_exp emin prec) x
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y
generic_format beta (FTZ_exp emin prec) y
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y

(3 <= beta)%Z
exact Hbeta.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y

Valid_exp (FTZ_exp emin prec)
now apply FTZ_exp_valid.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y

Valid_exp (FTZ_exp emin' prec')
now apply FTZ_exp_valid.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y

round_round_plus_radix_ge_3_hyp (FTZ_exp emin prec) (FTZ_exp emin' prec')
now apply FTZ_round_round_plus_radix_ge_3_hyp.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y

generic_format beta (FTZ_exp emin prec) x
now apply generic_format_FTZ.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y

generic_format beta (FTZ_exp emin prec) y
now apply generic_format_FTZ. Qed.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(3 <= beta)%Z -> forall choice1 choice2 : Z -> bool, (emin' + prec' <= emin + 1)%Z -> (2 * prec <= prec')%Z -> forall x y : R, FTZ_format beta emin prec x -> FTZ_format beta emin prec y -> round_round_eq (FTZ_exp emin prec) (FTZ_exp emin' prec') choice1 choice2 (x - y)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(3 <= beta)%Z -> forall choice1 choice2 : Z -> bool, (emin' + prec' <= emin + 1)%Z -> (2 * prec <= prec')%Z -> forall x y : R, FTZ_format beta emin prec x -> FTZ_format beta emin prec y -> round_round_eq (FTZ_exp emin prec) (FTZ_exp emin' prec') choice1 choice2 (x - y)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y

round_round_eq (FTZ_exp emin prec) (FTZ_exp emin' prec') choice1 choice2 (x - y)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y

(3 <= beta)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y
Valid_exp (FTZ_exp emin prec)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y
Valid_exp (FTZ_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y
round_round_plus_radix_ge_3_hyp (FTZ_exp emin prec) (FTZ_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y
generic_format beta (FTZ_exp emin prec) x
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y
generic_format beta (FTZ_exp emin prec) y
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y

(3 <= beta)%Z
exact Hbeta.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y

Valid_exp (FTZ_exp emin prec)
now apply FTZ_exp_valid.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y

Valid_exp (FTZ_exp emin' prec')
now apply FTZ_exp_valid.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y

round_round_plus_radix_ge_3_hyp (FTZ_exp emin prec) (FTZ_exp emin' prec')
now apply FTZ_round_round_plus_radix_ge_3_hyp.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y

generic_format beta (FTZ_exp emin prec) x
now apply generic_format_FTZ.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(3 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin' + prec' <= emin + 1)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y

generic_format beta (FTZ_exp emin prec) y
now apply generic_format_FTZ. Qed. End Double_round_plus_radix_ge_3_FTZ. End Double_round_plus_radix_ge_3. End Double_round_plus.
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall (choice1 choice2 : Z -> bool) (x : R), 0 < x -> (fexp2 (mag x) <= fexp1 (mag x) - 1)%Z -> (fexp1 (mag x) <= mag x)%Z -> (Rabs (x - midp fexp1 x) <= / 2 * ulp beta fexp2 x -> round_round_eq fexp1 fexp2 choice1 choice2 x) -> round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall (choice1 choice2 : Z -> bool) (x : R), 0 < x -> (fexp2 (mag x) <= fexp1 (mag x) - 1)%Z -> (fexp1 (mag x) <= mag x)%Z -> (Rabs (x - midp fexp1 x) <= / 2 * ulp beta fexp2 x -> round_round_eq fexp1 fexp2 choice1 choice2 x) -> round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z

(Rabs (x - midp fexp1 x) <= / 2 * ulp beta fexp2 x -> round_round_eq fexp1 fexp2 choice1 choice2 x) -> round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z

(Rabs (x - (round beta fexp1 Zfloor x + / 2 * ulp beta fexp1 x)) <= / 2 * ulp beta fexp2 x -> round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x) -> round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R

(Rabs (x - (rd + / 2 * ulp beta fexp1 x)) <= / 2 * ulp beta fexp2 x -> round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x) -> round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R

(Rabs (x - (rd + / 2 * u1)) <= / 2 * ulp beta fexp2 x -> round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x) -> round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R

(Rabs (x - (rd + / 2 * u1)) <= / 2 * u2 -> round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x) -> round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cmid:Rabs (x - (rd + / 2 * u1)) <= / 2 * u2 -> round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cmid:Rabs (x - (rd + / 2 * u1)) <= / 2 * u2 -> round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
Fx:generic_format beta fexp1 x

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cmid:Rabs (x - (rd + / 2 * u1)) <= / 2 * u2 -> round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
Nfx:~ generic_format beta fexp1 x
round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cmid:Rabs (x - (rd + / 2 * u1)) <= / 2 * u2 -> round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
Fx:generic_format beta fexp1 x

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cmid:Rabs (x - (rd + / 2 * u1)) <= / 2 * u2 -> round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
Fx:generic_format beta fexp1 x

generic_format beta fexp2 x
now apply (generic_inclusion_mag beta fexp1); [omega|].
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cmid:Rabs (x - (rd + / 2 * u1)) <= / 2 * u2 -> round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
Nfx:~ generic_format beta fexp1 x

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cmid:Rabs (x - (rd + / 2 * u1)) <= / 2 * u2 -> round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
Nfx:~ generic_format beta fexp1 x
Hceil:round beta fexp1 Zceil x = rd + u1

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cmid:Rabs (x - (rd + / 2 * u1)) <= / 2 * u2 -> round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
Nfx:~ generic_format beta fexp1 x
Hceil:round beta fexp1 Zceil x = rd + u1
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cmid:Rabs (x - (rd + / 2 * u1)) <= / 2 * u2 -> round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
Nfx:~ generic_format beta fexp1 x
Hceil:round beta fexp1 Zceil x = rd + u1
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
H:x - rd < / 2 * (u1 - u2)

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cmid:Rabs (x - (rd + / 2 * u1)) <= / 2 * u2 -> round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
Nfx:~ generic_format beta fexp1 x
Hceil:round beta fexp1 Zceil x = rd + u1
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
H:/ 2 * (u1 - u2) <= x - rd
round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cmid:Rabs (x - (rd + / 2 * u1)) <= / 2 * u2 -> round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
Nfx:~ generic_format beta fexp1 x
Hceil:round beta fexp1 Zceil x = rd + u1
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
H:x - rd < / 2 * (u1 - u2)

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cmid:Rabs (x - (rd + / 2 * u1)) <= / 2 * u2 -> round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
Nfx:~ generic_format beta fexp1 x
Hceil:round beta fexp1 Zceil x = rd + u1
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
H:x - rd < / 2 * (u1 - u2)

x < midp fexp1 x - / 2 * ulp beta fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cmid:Rabs (x - (rd + / 2 * u1)) <= / 2 * u2 -> round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
Nfx:~ generic_format beta fexp1 x
Hceil:round beta fexp1 Zceil x = rd + u1
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
H:x - rd < / 2 * (u1 - u2)

x < round beta fexp1 Zfloor x + / 2 * ulp beta fexp1 x - / 2 * ulp beta fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cmid:Rabs (x - (rd + / 2 * u1)) <= / 2 * u2 -> round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
Nfx:~ generic_format beta fexp1 x
Hceil:round beta fexp1 Zceil x = rd + u1
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
H:x - rd < / 2 * (u1 - u2)

x < rd + / 2 * u1 - / 2 * u2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cmid:Rabs (x - (rd + / 2 * u1)) <= / 2 * u2 -> round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
Nfx:~ generic_format beta fexp1 x
Hceil:round beta fexp1 Zceil x = rd + u1
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
H:x - rd < / 2 * (u1 - u2)

x - rd < / 2 * u1 - / 2 * u2
now rewrite <- Rmult_minus_distr_l.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cmid:Rabs (x - (rd + / 2 * u1)) <= / 2 * u2 -> round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
Nfx:~ generic_format beta fexp1 x
Hceil:round beta fexp1 Zceil x = rd + u1
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
H:/ 2 * (u1 - u2) <= x - rd

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cmid:Rabs (x - (rd + / 2 * u1)) <= / 2 * u2 -> round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
Nfx:~ generic_format beta fexp1 x
Hceil:round beta fexp1 Zceil x = rd + u1
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
H:/ 2 * (u1 - u2) <= x - rd

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cmid:Rabs (x - (rd + / 2 * u1)) <= / 2 * u2 -> round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
Nfx:~ generic_format beta fexp1 x
Hceil:round beta fexp1 Zceil x = rd + u1
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
H:/ 2 * (u1 - u2) <= x - rd
H0:/ 2 * (u1 + u2) < x - rd

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cmid:Rabs (x - (rd + / 2 * u1)) <= / 2 * u2 -> round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
Nfx:~ generic_format beta fexp1 x
Hceil:round beta fexp1 Zceil x = rd + u1
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
H:/ 2 * (u1 - u2) <= x - rd
H0:x - rd <= / 2 * (u1 + u2)
round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cmid:Rabs (x - (rd + / 2 * u1)) <= / 2 * u2 -> round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
Nfx:~ generic_format beta fexp1 x
Hceil:round beta fexp1 Zceil x = rd + u1
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
H:/ 2 * (u1 - u2) <= x - rd
H0:/ 2 * (u1 + u2) < x - rd

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cmid:Rabs (x - (rd + / 2 * u1)) <= / 2 * u2 -> round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
Nfx:~ generic_format beta fexp1 x
Hceil:round beta fexp1 Zceil x = rd + u1
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
H:/ 2 * (u1 - u2) <= x - rd
H0:/ 2 * (u1 + u2) < x - rd

round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cmid:Rabs (x - (rd + / 2 * u1)) <= / 2 * u2 -> round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
Nfx:~ generic_format beta fexp1 x
Hceil:round beta fexp1 Zceil x = rd + u1
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
H:/ 2 * (u1 - u2) <= x - rd
H0:/ 2 * (u1 + u2) < x - rd
H1:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cmid:Rabs (x - (rd + / 2 * u1)) <= / 2 * u2 -> round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
Nfx:~ generic_format beta fexp1 x
Hceil:round beta fexp1 Zceil x = rd + u1
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
H:/ 2 * (u1 - u2) <= x - rd
H0:/ 2 * (u1 + u2) < x - rd

round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cmid:Rabs (x - (rd + / 2 * u1)) <= / 2 * u2 -> round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
Nfx:~ generic_format beta fexp1 x
Hceil:round beta fexp1 Zceil x = rd + u1
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
H:/ 2 * (u1 - u2) <= x - rd
H0:/ 2 * (u1 + u2) < x - rd

rd + u1 - x < / 2 * (u1 - u2)
lra.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cmid:Rabs (x - (rd + / 2 * u1)) <= / 2 * u2 -> round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
Nfx:~ generic_format beta fexp1 x
Hceil:round beta fexp1 Zceil x = rd + u1
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
H:/ 2 * (u1 - u2) <= x - rd
H0:/ 2 * (u1 + u2) < x - rd
H1:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cmid:Rabs (x - (rd + / 2 * u1)) <= / 2 * u2 -> round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
Nfx:~ generic_format beta fexp1 x
Hceil:round beta fexp1 Zceil x = rd + u1
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
H:/ 2 * (u1 - u2) <= x - rd
H0:/ 2 * (u1 + u2) < x - rd
H1:round beta fexp1 Zceil x - x < / 2 * (ulp beta fexp1 x - ulp beta fexp2 x)

midp' fexp1 x + / 2 * ulp beta fexp2 x < x
unfold midp'; lra.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2f1:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cmid:Rabs (x - (rd + / 2 * u1)) <= / 2 * u2 -> round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
Nfx:~ generic_format beta fexp1 x
Hceil:round beta fexp1 Zceil x = rd + u1
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
H:/ 2 * (u1 - u2) <= x - rd
H0:x - rd <= / 2 * (u1 + u2)

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
apply Cmid, Rabs_le; split; lra. } Qed. Section Double_round_sqrt. Definition round_round_sqrt_hyp fexp1 fexp2 := (forall ex, (2 * fexp1 ex <= fexp1 (2 * ex))%Z) /\ (forall ex, (2 * fexp1 ex <= fexp1 (2 * ex - 1))%Z) /\ (forall ex, (fexp1 (2 * ex) < 2 * ex)%Z -> (fexp2 ex + ex <= 2 * fexp1 ex - 2)%Z).
beta:radix

forall x : R, 0 < x -> mag x = (2 * mag (sqrt x) - 1)%Z \/ mag x = (2 * mag (sqrt x))%Z
beta:radix

forall x : R, 0 < x -> mag x = (2 * mag (sqrt x) - 1)%Z \/ mag x = (2 * mag (sqrt x))%Z
beta:radix
x:R
Px:0 < x

mag x = (2 * mag (sqrt x) - 1)%Z \/ mag x = (2 * mag (sqrt x))%Z
beta:radix
x:R
Px:0 < x

mag x = (2 * Z.div2 (mag x + 1) - 1)%Z \/ mag x = (2 * Z.div2 (mag x + 1))%Z
beta:radix
x:R
Px:0 < x

(mag x + 1)%Z = (2 * Z.div2 (mag x + 1) + (if Z.odd (mag x + 1) then 1 else 0))%Z -> mag x = (2 * Z.div2 (mag x + 1) - 1)%Z \/ mag x = (2 * Z.div2 (mag x + 1))%Z
destruct Z.odd ; intros ; omega. Qed.
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> round_round_sqrt_hyp fexp1 fexp2 -> forall x : R, 0 < x -> (fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z -> generic_format beta fexp1 x -> / 2 * ulp beta fexp2 (sqrt x) < Rabs (sqrt x - midp fexp1 (sqrt x))
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> round_round_sqrt_hyp fexp1 fexp2 -> forall x : R, 0 < x -> (fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z -> generic_format beta fexp1 x -> / 2 * ulp beta fexp2 (sqrt x) < Rabs (sqrt x - midp fexp1 (sqrt x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Fx:generic_format beta fexp1 x

/ 2 * ulp beta fexp2 (sqrt x) < Rabs (sqrt x - midp fexp1 (sqrt x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Fx:generic_format beta fexp1 x

(2 <= beta)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Fx:generic_format beta fexp1 x
Hbeta:(2 <= beta)%Z
/ 2 * ulp beta fexp2 (sqrt x) < Rabs (sqrt x - midp fexp1 (sqrt x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Fx:generic_format beta fexp1 x

(2 <= beta)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
beta_val:Z
beta_prop:(2 <=? beta_val)%Z = true
Hf2:(fexp2 (Raux.mag {| radix_val := beta_val; radix_prop := beta_prop |} (sqrt x)) <= fexp1 (Raux.mag {| radix_val := beta_val; radix_prop := beta_prop |} (sqrt x)) - 1)%Z
Fx:generic_format {| radix_val := beta_val; radix_prop := beta_prop |} fexp1 x

(2 <= {| radix_val := beta_val; radix_prop := beta_prop |})%Z
now apply Zle_bool_imp_le.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Fx:generic_format beta fexp1 x
Hbeta:(2 <= beta)%Z

/ 2 * ulp beta fexp2 (sqrt x) < Rabs (sqrt x - midp fexp1 (sqrt x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Fx:generic_format beta fexp1 x
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R

/ 2 * ulp beta fexp2 (sqrt x) < Rabs (sqrt x - midp fexp1 (sqrt x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Fx:generic_format beta fexp1 x
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R

/ 2 * ulp beta fexp2 (sqrt x) < Rabs (sqrt x - midp fexp1 (sqrt x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Fx:generic_format beta fexp1 x
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R

/ 2 * ulp beta fexp2 (sqrt x) < Rabs (sqrt x - midp fexp1 (sqrt x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Fx:generic_format beta fexp1 x
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R

/ 2 * ulp beta fexp2 (sqrt x) < Rabs (sqrt x - midp fexp1 (sqrt x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Fx:generic_format beta fexp1 x
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R

/ 2 * ulp beta fexp2 (sqrt x) < Rabs (sqrt x - midp fexp1 (sqrt x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Fx:generic_format beta fexp1 x
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R

/ 2 * bpow (cexp beta fexp2 (sqrt x)) < Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x))))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Fx:generic_format beta fexp1 x
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))

False
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Fx:generic_format beta fexp1 x
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))

generic_format beta fexp1 a
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Fx:generic_format beta fexp1 x
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
Fa:generic_format beta fexp1 a
False
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Fx:generic_format beta fexp1 x
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))

generic_format beta fexp1 a
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Fx:generic_format beta fexp1 x
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))

generic_format beta fexp1 (round beta fexp1 Zfloor (sqrt x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Fx:generic_format beta fexp1 x
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))

Valid_exp fexp1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Fx:generic_format beta fexp1 x
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
Valid_rnd Zfloor
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Fx:generic_format beta fexp1 x
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))

Valid_exp fexp1
exact Vfexp1.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Fx:generic_format beta fexp1 x
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))

Valid_rnd Zfloor
now apply valid_rnd_DN.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Fx:generic_format beta fexp1 x
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
Fa:generic_format beta fexp1 a

False
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))

generic_format beta fexp1 x -> generic_format beta fexp1 a -> False
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))

x = IZR (Ztrunc (x * bpow (- fexp1 (mag x)))) * bpow (fexp1 (mag x)) -> a = IZR (Ztrunc (a * bpow (- fexp1 (mag a)))) * bpow (fexp1 (mag a)) -> False
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z

x = IZR mx * bpow (fexp1 (mag x)) -> a = IZR (Ztrunc (a * bpow (- fexp1 (mag a)))) * bpow (fexp1 (mag a)) -> False
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z

x = IZR mx * bpow (fexp1 (mag x)) -> a = IZR ma * bpow (fexp1 (mag a)) -> False
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))

False
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))

0 <= a
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
False
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))

0 <= a
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))

round beta fexp1 Zfloor 0 <= a
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))

Valid_exp fexp1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Valid_rnd Zfloor
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
0 <= sqrt x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))

Valid_exp fexp1
exact Vfexp1.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))

Valid_rnd Zfloor
now apply valid_rnd_DN.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))

0 <= sqrt x
apply sqrt_pos.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a

False
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a

0 < / 2 * u1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
False
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a

0 < / 2 * u1
apply Rmult_lt_0_compat; [lra|apply bpow_gt_0].
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1

False
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1

0 < / 2 * u2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
False
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1

0 < / 2 * u2
apply Rmult_lt_0_compat; [lra|apply bpow_gt_0].
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2

False
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2

0 < b
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
False
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2

0 < b
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2

0 < / 2 * (u1 - u2)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2

/ 2 * 0 < / 2 * (u1 - u2)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2

0 < u1 - u2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2

u2 < u1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2

bpow (fexp2 (mag (sqrt x))) < bpow (fexp1 (mag (sqrt x)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2

(fexp2 (mag (sqrt x)) < fexp1 (mag (sqrt x)))%Z
omega.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b

False
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b

0 < b'
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
False
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b

0 < b'
now unfold b'; rewrite Rmult_plus_distr_l; apply Rplus_lt_0_compat.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'

False
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'

sqrt x <= a + b'
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
False
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'

sqrt x <= a + b'
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'

sqrt x - / 2 * u1 - a <= / 2 * u2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'

sqrt x - (a + / 2 * u1) <= / 2 * u2
now apply Rabs_le_inv.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'

False
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'

a + b <= sqrt x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
False
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'

a + b <= sqrt x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'

- / 2 * u2 <= - a - / 2 * u1 + sqrt x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'

- / 2 * u2 <= sqrt x - (a + / 2 * u1)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'

- (/ 2 * u2) <= sqrt x - (a + / 2 * u1)
now apply Rabs_le_inv in H; destruct H.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x

False
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z

False
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z

(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
False
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z

(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:mag x = (2 * mag (sqrt x) - 1)%Z

(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:mag x = (2 * mag (sqrt x))%Z
(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:mag x = (2 * mag (sqrt x) - 1)%Z

(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:mag x = (2 * mag (sqrt x) - 1)%Z

(fexp1 (mag x) < mag x)%Z
now apply mag_generic_gt; [|apply Rgt_not_eq|].
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:mag x = (2 * mag (sqrt x))%Z

(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:mag x = (2 * mag (sqrt x))%Z

(fexp1 (mag x) < mag x)%Z
now apply mag_generic_gt; [|apply Rgt_not_eq|].
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z

False
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z

a * a + u1 * a - u2 * a + b * b <= x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
False
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z

a * a + u1 * a - u2 * a + b * b <= x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z

(a + b) * (a + b) <= x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z

(a + b) * (a + b) <= sqrt x * sqrt x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
H':0 <= a + b

(a + b) * (a + b) <= sqrt x * sqrt x
now apply Rmult_le_compat.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x

False
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x

x <= a * a + u1 * a + u2 * a + b' * b'
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
False
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x

x <= a * a + u1 * a + u2 * a + b' * b'
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x

x <= (a + b') * (a + b')
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x

sqrt x * sqrt x <= (a + b') * (a + b')
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
H':0 <= sqrt x

sqrt x * sqrt x <= (a + b') * (a + b')
now apply Rmult_le_compat.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'

False
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0

False
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
False
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0

False
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0

0 < 0
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0

b * b <= 0
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0

b * b <= x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0
x <= 0
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0

b * b <= x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0

0 + 0 + - 0 + b * b <= x -> b * b <= x
now rewrite Ropp_0; do 3 rewrite Rplus_0_l.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0

x <= 0
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0

IZR mx * bpow (fexp1 (mag x)) <= 0
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0

IZR mx * bpow (fexp1 (mag x)) * bpow (- fexp1 (mag x)) <= 0 * bpow (- fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0

IZR mx <= 0
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0

IZR (Ztrunc (x * bpow (- fexp1 (mag x)))) <= 0
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0

IZR (Zfloor (x * bpow (- fexp1 (mag x)))) <= 0
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0

Zfloor (x * bpow (- fexp1 (mag x))) = 0%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0

0 <= x * bpow (- fexp1 (mag x)) < IZR (0 + 1)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0

x * bpow (- fexp1 (mag x)) < 1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0

x * bpow (- fexp1 (mag x)) * bpow (fexp1 (mag x)) < 1 * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0

x < bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0

x < bpow (2 * fexp1 (mag (sqrt x)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0

x < bpow (fexp1 (mag (sqrt x)) + fexp1 (mag (sqrt x)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0

x < bpow (fexp1 (mag (sqrt x))) * bpow (fexp1 (mag (sqrt x)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0

sqrt x * sqrt x < bpow (fexp1 (mag (sqrt x))) * bpow (fexp1 (mag (sqrt x)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0

sqrt x < bpow (fexp1 (mag (sqrt x)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0

b' < bpow (fexp1 (mag (sqrt x)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0

/ 2 * (u1 + u2) < u1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0

/ 2 * (u1 + u2) < / 2 * (u1 + u1)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0

u2 < u1
unfold u2, u1, ulp, cexp; apply bpow_lt; omega.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0

False
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a

False
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a

mag a = mag (sqrt x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
False
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a

mag a = mag (sqrt x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a

Valid_exp fexp1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
0 < round beta fexp1 Zfloor (sqrt x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a

Valid_exp fexp1
exact Vfexp1.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a

0 < round beta fexp1 Zfloor (sqrt x)
now fold a.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

False
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

0 < - (u2 * a) + b * b
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b
False
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

0 < - (u2 * a) + b * b
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

u2 * a < b ^ 2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

u2 * a < u2 ^ 2 * (/ 2) ^ 2 - 2 * u2 * (/ 2) ^ 2 * u1 + (/ 2) ^ 2 * u1 ^ 2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

(2 * u2 * a + u2 * u1) / 2 < (2 * u2 ^ 2 + 2 * u1 ^ 2) / 8
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

u2 * (a + / 2 * u1) < (2 * u2 ^ 2 + 2 * u1 ^ 2) / 8
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

u2 * (a + / 2 * u1) < / 4 * (u2 ^ 2 + u1 ^ 2)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

u2 * (a + / 2 * u1) < u2 * bpow (mag (sqrt x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
u2 * bpow (mag (sqrt x)) <= / 4 * (u2 ^ 2 + u1 ^ 2)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

u2 * (a + / 2 * u1) < u2 * bpow (mag (sqrt x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

a + / 2 * u1 < bpow (mag (sqrt x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

a + / 2 * bpow (fexp1 (mag a)) < bpow (mag a)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

a + / 2 * bpow (fexp1 (mag a)) < a + bpow (fexp1 (mag a))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
a + bpow (fexp1 (mag a)) <= bpow (mag a)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

a + / 2 * bpow (fexp1 (mag a)) < a + bpow (fexp1 (mag a))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

/ 2 * bpow (fexp1 (mag a)) < bpow (fexp1 (mag a))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

/ 2 * bpow (fexp1 (mag a)) < 1 * bpow (fexp1 (mag a))
apply Rmult_lt_compat_r; [apply bpow_gt_0|lra].
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

a + bpow (fexp1 (mag a)) <= bpow (mag a)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

a + bpow (fexp1 (mag a)) <= a + ulp beta fexp1 a
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
a + ulp beta fexp1 a <= bpow (mag a)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

a + ulp beta fexp1 a <= bpow (mag a)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

a < bpow (mag a)
apply Rabs_lt_inv, bpow_mag_gt.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

u2 * bpow (mag (sqrt x)) <= / 4 * (u2 ^ 2 + u1 ^ 2)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

u2 * bpow (mag (sqrt x)) <= bpow (-2) * u1 ^ 2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
bpow (-2) * u1 ^ 2 <= / 4 * (u2 ^ 2 + u1 ^ 2)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

u2 * bpow (mag (sqrt x)) <= bpow (-2) * u1 ^ 2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

u2 * bpow (mag (sqrt x)) <= bpow (-2) * (u1 * u1)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

(fexp2 (mag (sqrt x)) + mag (sqrt x) <= 2 * fexp1 (mag (sqrt x)) - 2)%Z
now apply Hexp.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

bpow (-2) * u1 ^ 2 <= / 4 * (u2 ^ 2 + u1 ^ 2)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

0 <= bpow (-2)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
0 <= u1 ^ 2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
bpow (-2) <= / 4
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
u1 ^ 2 <= u2 ^ 2 + u1 ^ 2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

0 <= bpow (-2)
apply bpow_ge_0.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

0 <= u1 ^ 2
apply pow2_ge_0.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

bpow (-2) <= / 4
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

/ IZR (beta * beta) <= / 4
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

4 <= IZR (beta * beta)
change 4%Z with (2 * 2)%Z; apply IZR_le, Zmult_le_compat; omega.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

u1 ^ 2 <= u2 ^ 2 + u1 ^ 2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

0 <= u2 ^ 2
apply pow2_ge_0.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

False
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

x <= a * a + u1 * a
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b
Hr':x <= a * a + u1 * a
False
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

x <= a * a + u1 * a
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

x <= a * a + u1 * a
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

x <= (a + u1) * a
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

x <= (a + bpow (fexp1 (mag (sqrt x)))) * a
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

x <= (IZR ma + 1) * bpow (fexp1 (mag (sqrt x))) * (IZR ma * bpow (fexp1 (mag (sqrt x))))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

x <= IZR ma * ((IZR ma + 1) * bpow (fexp1 (mag (sqrt x)))) * bpow (fexp1 (mag (sqrt x)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

x <= IZR ma * (IZR ma + 1) * bpow (2 * fexp1 (mag (sqrt x)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

x * bpow (-2 * fexp1 (mag (sqrt x))) <= IZR ma * (IZR ma + 1)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

IZR mx * bpow (fexp1 (mag x) - 2 * fexp1 (mag (sqrt x))) <= IZR ma * (IZR ma + 1)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

IZR mx * IZR (beta ^ (fexp1 (mag x) - 2 * fexp1 (mag (sqrt x)))) <= IZR ma * (IZR ma + 1)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

IZR (mx * beta ^ (fexp1 (mag x) - 2 * fexp1 (mag (sqrt x)))) <= IZR (ma * (ma + 1))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

IZR (mx * beta ^ (fexp1 (mag x) - 2 * fexp1 (mag (sqrt x)))) < IZR (Z.succ (ma * (ma + 1)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

IZR mx * IZR (beta ^ (fexp1 (mag x) - 2 * fexp1 (mag (sqrt x)))) < IZR ma * (IZR ma + 1) + 1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

IZR mx * bpow (fexp1 (mag x) - 2 * fexp1 (mag (sqrt x))) < IZR ma * (IZR ma + 1) + 1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

IZR mx * bpow (fexp1 (mag x)) < (IZR ma * (IZR ma + 1) + 1) * bpow (2 * fexp1 (mag (sqrt x)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

x < (IZR ma * (IZR ma + 1) + 1) * bpow (2 * fexp1 (mag (sqrt x)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

x < (IZR ma * (IZR ma + 1) + 1) * bpow (fexp1 (mag (sqrt x)) + fexp1 (mag (sqrt x)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

x < (IZR ma * (IZR ma + 1) + 1) * (bpow (fexp1 (mag (sqrt x))) * bpow (fexp1 (mag (sqrt x))))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

x < a * a + u1 * a + u1 * u1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

a * a + u1 * a + u2 * a + b' * b' < a * a + u1 * a + u1 * u1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

u2 * a + b' * b' < u1 * u1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

u2 * a + u2 * / 2 * u1 < u2 * / 2 * u1 - b' ^ 2 + u1 ^ 2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

(a + / 2 * u1) * u2 < u2 * / 2 * u1 - b' ^ 2 + u1 ^ 2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

(a + / 2 * u1) * u2 < bpow (mag (sqrt x)) * u2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b
bpow (mag (sqrt x)) * u2 <= u2 * / 2 * u1 - b' ^ 2 + u1 ^ 2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

(a + / 2 * u1) * u2 < bpow (mag (sqrt x)) * u2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

a + / 2 * u1 < bpow (mag (sqrt x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

a + u1 <= bpow (mag (sqrt x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

a + bpow (cexp beta fexp1 (sqrt x)) <= bpow (mag (sqrt x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

a + bpow (cexp beta fexp1 a) <= bpow (mag (sqrt x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

a + ulp beta fexp1 a <= bpow (mag (sqrt x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

0 < a
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b
generic_format beta fexp1 a
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b
a < bpow (mag (sqrt x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

0 < a
exact Pa.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

generic_format beta fexp1 a
now apply round_DN_pt.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

a < bpow (mag (sqrt x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

a <= sqrt x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b
sqrt x < bpow (mag (sqrt x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

a <= sqrt x
now apply round_DN_pt.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

sqrt x < bpow (mag (sqrt x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

Rabs (sqrt x) < bpow (mag (sqrt x))
apply bpow_mag_gt.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

bpow (mag (sqrt x)) * u2 <= u2 * / 2 * u1 - b' ^ 2 + u1 ^ 2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

bpow (mag (sqrt x)) * u2 <= / 2 * u1 ^ 2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b
/ 2 * u1 ^ 2 <= u2 * / 2 * u1 - b' ^ 2 + u1 ^ 2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

bpow (mag (sqrt x)) * u2 <= / 2 * u1 ^ 2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

bpow (mag (sqrt x)) * u2 <= bpow (-2) * u1 ^ 2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b
bpow (-2) * u1 ^ 2 <= / 2 * u1 ^ 2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

bpow (mag (sqrt x)) * u2 <= bpow (-2) * u1 ^ 2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

bpow (mag (sqrt x)) * u2 <= bpow (-2) * (u1 * u1)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

bpow (mag (sqrt x)) * bpow (fexp2 (mag (sqrt x))) <= bpow (-2) * (bpow (fexp1 (mag (sqrt x))) * bpow (fexp1 (mag (sqrt x))))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

bpow (mag (sqrt x) + fexp2 (mag (sqrt x))) <= bpow (2 * fexp1 (mag (sqrt x)) - 2)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

(mag (sqrt x) + fexp2 (mag (sqrt x)) <= 2 * fexp1 (mag (sqrt x)) - 2)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

(fexp2 (mag (sqrt x)) + mag (sqrt x) <= 2 * fexp1 (mag (sqrt x)) - 2)%Z
now apply Hexp.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

bpow (-2) * u1 ^ 2 <= / 2 * u1 ^ 2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

bpow (-2) <= / 2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

/ IZR (beta * (beta * 1)) <= / 2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

/ IZR (beta * beta) <= / 2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

2 <= IZR (beta * beta)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

(2 <= beta * beta)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

(1 * 2 <= beta * beta)%Z
apply Zmult_le_compat; omega.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

/ 2 * u1 ^ 2 <= u2 * / 2 * u1 - b' ^ 2 + u1 ^ 2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

u2 ^ 2 < u1 ^ 2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

u2 * u2 < u1 * u1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b
H':0 <= u2

u2 * u2 < u1 * u1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b
H':0 <= u2

u2 < u1
unfold u1, u2, ulp, cexp; apply bpow_lt; omega.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b
Hr':x <= a * a + u1 * a

False
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b
Hr':x <= a * a + u1 * a

a * a + u1 * a < a * a + u1 * a
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b
Hr':x <= a * a + u1 * a

a * a + u1 * a < a * a + u1 * a - u2 * a + b * b
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b
Hr':x <= a * a + u1 * a
a * a + u1 * a - u2 * a + b * b <= a * a + u1 * a
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b
Hr':x <= a * a + u1 * a

a * a + u1 * a < a * a + u1 * a - u2 * a + b * b
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b
Hr':x <= a * a + u1 * a

a * a + u1 * a + 0 < a * a + u1 * a - u2 * a + b * b
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b
Hr':x <= a * a + u1 * a

a * a + u1 * a + 0 < a * a + u1 * a + (- (u2 * a) + b * b)
now apply Rplus_lt_compat_l.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hbeta:(2 <= beta)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b
Hr':x <= a * a + u1 * a

a * a + u1 * a - u2 * a + b * b <= a * a + u1 * a
now apply Rle_trans with x. Qed.
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall choice1 choice2 : Z -> bool, round_round_sqrt_hyp fexp1 fexp2 -> forall x : R, generic_format beta fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 (sqrt x)
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall choice1 choice2 : Z -> bool, round_round_sqrt_hyp fexp1 fexp2 -> forall x : R, generic_format beta fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 (sqrt x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x

round_round_eq fexp1 fexp2 choice1 choice2 (sqrt x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (sqrt x)) = round beta fexp1 (Znearest choice1) (sqrt x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Npx:x <= 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (sqrt x)) = round beta fexp1 (Znearest choice1) (sqrt x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (sqrt x)) = round beta fexp1 (Znearest choice1) (sqrt x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Npx:x <= 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (sqrt x)) = round beta fexp1 (Znearest choice1) (sqrt x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Npx:x <= 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) 0) = round beta fexp1 (Znearest choice1) 0
now rewrite round_0; [|apply valid_rnd_N].
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (sqrt x)) = round beta fexp1 (Znearest choice1) (sqrt x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (sqrt x)) = round beta fexp1 (Znearest choice1) (sqrt x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z

(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (sqrt x)) = round beta fexp1 (Znearest choice1) (sqrt x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z

(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hx:x <= 1

(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hx:1 < x
(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hx:x <= 1

(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hx:x <= 1

(mag x <= mag (sqrt x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hx:x <= 1

x <= sqrt x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hx:x <= 1

sqrt x * sqrt x <= sqrt x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hx:x <= 1

sqrt x * sqrt x <= sqrt x * 1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hx:x <= 1

0 <= sqrt x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hx:x <= 1
sqrt x <= 1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hx:x <= 1

0 <= sqrt x
apply sqrt_pos.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hx:x <= 1

sqrt x <= 1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hx:x <= 1

sqrt x <= sqrt 1
now apply sqrt_le_1_alt.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hx:1 < x

(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hx:1 < x

(2 * fexp1 1 <= fexp1 (2 * 1 - 1))%Z -> (fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hx:1 < x

(2 * fexp1 1 <= fexp1 1)%Z -> (fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hx:1 < x
Hexp10:(2 * fexp1 1 <= fexp1 1)%Z

(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hx:1 < x
Hf0:(fexp1 1 < 1)%Z

(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hx:1 < x
Hf0:(fexp1 1 < 1)%Z

(1 <= mag (sqrt x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hx:1 < x
Hf0:(fexp1 1 < 1)%Z

bpow (1 - 1) <= Rabs (sqrt x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hx:1 < x
Hf0:(fexp1 1 < 1)%Z

bpow 0 <= Rabs (sqrt x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hx:1 < x
Hf0:(fexp1 1 < 1)%Z

1 <= Rabs (sqrt x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hx:1 < x
Hf0:(fexp1 1 < 1)%Z

1 <= sqrt x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hx:1 < x
Hf0:(fexp1 1 < 1)%Z

sqrt 1 <= sqrt x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hx:1 < x
Hf0:(fexp1 1 < 1)%Z

1 <= x
now apply Rlt_le.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (sqrt x)) = round beta fexp1 (Znearest choice1) (sqrt x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z

(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (sqrt x)) = round beta fexp1 (Znearest choice1) (sqrt x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z

(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z

(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
H:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z

(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
Hlx:mag x = (2 * mag (sqrt x) - 1)%Z

(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
Hlx:mag x = (2 * mag (sqrt x))%Z
(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
Hlx:mag x = (2 * mag (sqrt x) - 1)%Z

(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
Hlx:mag x = (2 * mag (sqrt x) - 1)%Z

(fexp1 (mag x) < mag x)%Z
now apply mag_generic_gt; [|apply Rgt_not_eq|].
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
Hlx:mag x = (2 * mag (sqrt x))%Z

(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
Hlx:mag x = (2 * mag (sqrt x))%Z

(fexp1 (mag x) < mag x)%Z
now apply mag_generic_gt; [|apply Rgt_not_eq|].
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
H:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z

(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
H:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z

(fexp2 (mag (sqrt x)) + mag (sqrt x) <= 2 * fexp1 (mag (sqrt x)) - 2)%Z -> (fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
omega.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (sqrt x)) = round beta fexp1 (Znearest choice1) (sqrt x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z

Valid_exp fexp1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Valid_exp fexp2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
0 < sqrt x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
(fexp1 (mag (sqrt x)) <= mag (sqrt x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Rabs (sqrt x - midp fexp1 (sqrt x)) <= / 2 * ulp beta fexp2 (sqrt x) -> round_round_eq fexp1 fexp2 choice1 choice2 (sqrt x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z

Valid_exp fexp1
exact Vfexp1.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z

Valid_exp fexp2
exact Vfexp2.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z

0 < sqrt x
now apply sqrt_lt_R0.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z

(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
omega.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z

(fexp1 (mag (sqrt x)) <= mag (sqrt x))%Z
omega.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z

Rabs (sqrt x - midp fexp1 (sqrt x)) <= / 2 * ulp beta fexp2 (sqrt x) -> round_round_eq fexp1 fexp2 choice1 choice2 (sqrt x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hmid:Rabs (sqrt x - midp fexp1 (sqrt x)) <= / 2 * ulp beta fexp2 (sqrt x)

/ 2 * ulp beta fexp2 (sqrt x) < Rabs (sqrt x - midp fexp1 (sqrt x))
apply (round_round_sqrt_aux fexp1 fexp2 Vfexp1 Vfexp2 Hexp x Px Hf2 Fx). Qed. Section Double_round_sqrt_FLX. Variable prec : Z. Variable prec' : Z. Context { prec_gt_0_ : Prec_gt_0 prec }. Context { prec_gt_0_' : Prec_gt_0 prec' }.
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(2 * prec + 2 <= prec')%Z -> round_round_sqrt_hyp (FLX_exp prec) (FLX_exp prec')
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(2 * prec + 2 <= prec')%Z -> round_round_sqrt_hyp (FLX_exp prec) (FLX_exp prec')
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hprec:(2 * prec + 2 <= prec')%Z

round_round_sqrt_hyp (FLX_exp prec) (FLX_exp prec')
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hprec:(2 * prec + 2 <= prec')%Z

round_round_sqrt_hyp (fun e : Z => (e - prec)%Z) (fun e : Z => (e - prec')%Z)
beta:radix
prec, prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hprec:(2 * prec + 2 <= prec')%Z

round_round_sqrt_hyp (fun e : Z => (e - prec)%Z) (fun e : Z => (e - prec')%Z)
unfold round_round_sqrt_hyp; split; [|split]; intro ex; omega. Qed.
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

forall choice1 choice2 : Z -> bool, (2 * prec + 2 <= prec')%Z -> forall x : R, FLX_format beta prec x -> round_round_eq (FLX_exp prec) (FLX_exp prec') choice1 choice2 (sqrt x)
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

forall choice1 choice2 : Z -> bool, (2 * prec + 2 <= prec')%Z -> forall x : R, FLX_format beta prec x -> round_round_eq (FLX_exp prec) (FLX_exp prec') choice1 choice2 (sqrt x)
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hprec:(2 * prec + 2 <= prec')%Z
x:R
Fx:FLX_format beta prec x

round_round_eq (FLX_exp prec) (FLX_exp prec') choice1 choice2 (sqrt x)
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hprec:(2 * prec + 2 <= prec')%Z
x:R
Fx:FLX_format beta prec x

Valid_exp (FLX_exp prec)
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hprec:(2 * prec + 2 <= prec')%Z
x:R
Fx:FLX_format beta prec x
Valid_exp (FLX_exp prec')
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hprec:(2 * prec + 2 <= prec')%Z
x:R
Fx:FLX_format beta prec x
round_round_sqrt_hyp (FLX_exp prec) (FLX_exp prec')
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hprec:(2 * prec + 2 <= prec')%Z
x:R
Fx:FLX_format beta prec x
generic_format beta (FLX_exp prec) x
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hprec:(2 * prec + 2 <= prec')%Z
x:R
Fx:FLX_format beta prec x

Valid_exp (FLX_exp prec)
now apply FLX_exp_valid.
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hprec:(2 * prec + 2 <= prec')%Z
x:R
Fx:FLX_format beta prec x

Valid_exp (FLX_exp prec')
now apply FLX_exp_valid.
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hprec:(2 * prec + 2 <= prec')%Z
x:R
Fx:FLX_format beta prec x

round_round_sqrt_hyp (FLX_exp prec) (FLX_exp prec')
now apply FLX_round_round_sqrt_hyp.
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hprec:(2 * prec + 2 <= prec')%Z
x:R
Fx:FLX_format beta prec x

generic_format beta (FLX_exp prec) x
now apply generic_format_FLX. Qed. End Double_round_sqrt_FLX. Section Double_round_sqrt_FLT. Variable emin prec : Z. Variable emin' prec' : Z. Context { prec_gt_0_ : Prec_gt_0 prec }. Context { prec_gt_0_' : Prec_gt_0 prec' }.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(emin <= 0)%Z -> (emin' <= emin - prec - 2)%Z \/ (2 * emin' <= emin - 4 * prec - 2)%Z -> (2 * prec + 2 <= prec')%Z -> round_round_sqrt_hyp (FLT_exp emin prec) (FLT_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(emin <= 0)%Z -> (emin' <= emin - prec - 2)%Z \/ (2 * emin' <= emin - 4 * prec - 2)%Z -> (2 * prec + 2 <= prec')%Z -> round_round_sqrt_hyp (FLT_exp emin prec) (FLT_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin <= 0)%Z
Heminprec:(emin' <= emin - prec - 2)%Z \/ (2 * emin' <= emin - 4 * prec - 2)%Z
Hprec:(2 * prec + 2 <= prec')%Z

round_round_sqrt_hyp (FLT_exp emin prec) (FLT_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin <= 0)%Z
Heminprec:(emin' <= emin - prec - 2)%Z \/ (2 * emin' <= emin - 4 * prec - 2)%Z
Hprec:(2 * prec + 2 <= prec')%Z

round_round_sqrt_hyp (fun e : Z => Z.max (e - prec) emin) (fun e : Z => Z.max (e - prec') emin')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin <= 0)%Z
Heminprec:(emin' <= emin - prec - 2)%Z \/ (2 * emin' <= emin - 4 * prec - 2)%Z
Hprec:(2 * prec + 2 <= prec')%Z

round_round_sqrt_hyp (fun e : Z => Z.max (e - prec) emin) (fun e : Z => Z.max (e - prec') emin')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin <= 0)%Z
Heminprec:(emin' <= emin - prec - 2)%Z \/ (2 * emin' <= emin - 4 * prec - 2)%Z
Hprec:(2 * prec + 2 <= prec')%Z
ex:Z

(2 * Z.max (ex - prec) emin <= Z.max (2 * ex - prec) emin)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin <= 0)%Z
Heminprec:(emin' <= emin - prec - 2)%Z \/ (2 * emin' <= emin - 4 * prec - 2)%Z
Hprec:(2 * prec + 2 <= prec')%Z
ex:Z
(2 * Z.max (ex - prec) emin <= Z.max (2 * ex - 1 - prec) emin)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin <= 0)%Z
Heminprec:(emin' <= emin - prec - 2)%Z \/ (2 * emin' <= emin - 4 * prec - 2)%Z
Hprec:(2 * prec + 2 <= prec')%Z
ex:Z
(Z.max (2 * ex - prec) emin < 2 * ex)%Z -> (Z.max (ex - prec') emin' + ex <= 2 * Z.max (ex - prec) emin - 2)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin <= 0)%Z
Heminprec:(emin' <= emin - prec - 2)%Z \/ (2 * emin' <= emin - 4 * prec - 2)%Z
Hprec:(2 * prec + 2 <= prec')%Z
ex:Z

(2 * Z.max (ex - prec) emin <= Z.max (2 * ex - prec) emin)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin <= 0)%Z
Heminprec:(emin' <= emin - prec - 2)%Z \/ (2 * emin' <= emin - 4 * prec - 2)%Z
Hprec:(2 * prec + 2 <= prec')%Z
ex:Z

(ex - prec >= emin)%Z /\ Z.max (ex - prec) emin = (ex - prec)%Z \/ (ex - prec < emin)%Z /\ Z.max (ex - prec) emin = emin -> (2 * Z.max (ex - prec) emin <= Z.max (2 * ex - prec) emin)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin <= 0)%Z
Heminprec:(emin' <= emin - prec - 2)%Z \/ (2 * emin' <= emin - 4 * prec - 2)%Z
Hprec:(2 * prec + 2 <= prec')%Z
ex:Z

(2 * ex - prec >= emin)%Z /\ Z.max (2 * ex - prec) emin = (2 * ex - prec)%Z \/ (2 * ex - prec < emin)%Z /\ Z.max (2 * ex - prec) emin = emin -> (ex - prec >= emin)%Z /\ Z.max (ex - prec) emin = (ex - prec)%Z \/ (ex - prec < emin)%Z /\ Z.max (ex - prec) emin = emin -> (2 * Z.max (ex - prec) emin <= Z.max (2 * ex - prec) emin)%Z
omega.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin <= 0)%Z
Heminprec:(emin' <= emin - prec - 2)%Z \/ (2 * emin' <= emin - 4 * prec - 2)%Z
Hprec:(2 * prec + 2 <= prec')%Z
ex:Z

(2 * Z.max (ex - prec) emin <= Z.max (2 * ex - 1 - prec) emin)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin <= 0)%Z
Heminprec:(emin' <= emin - prec - 2)%Z \/ (2 * emin' <= emin - 4 * prec - 2)%Z
Hprec:(2 * prec + 2 <= prec')%Z
ex:Z

(ex - prec >= emin)%Z /\ Z.max (ex - prec) emin = (ex - prec)%Z \/ (ex - prec < emin)%Z /\ Z.max (ex - prec) emin = emin -> (2 * Z.max (ex - prec) emin <= Z.max (2 * ex - 1 - prec) emin)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin <= 0)%Z
Heminprec:(emin' <= emin - prec - 2)%Z \/ (2 * emin' <= emin - 4 * prec - 2)%Z
Hprec:(2 * prec + 2 <= prec')%Z
ex:Z

(2 * ex - 1 - prec >= emin)%Z /\ Z.max (2 * ex - 1 - prec) emin = (2 * ex - 1 - prec)%Z \/ (2 * ex - 1 - prec < emin)%Z /\ Z.max (2 * ex - 1 - prec) emin = emin -> (ex - prec >= emin)%Z /\ Z.max (ex - prec) emin = (ex - prec)%Z \/ (ex - prec < emin)%Z /\ Z.max (ex - prec) emin = emin -> (2 * Z.max (ex - prec) emin <= Z.max (2 * ex - 1 - prec) emin)%Z
omega.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin <= 0)%Z
Heminprec:(emin' <= emin - prec - 2)%Z \/ (2 * emin' <= emin - 4 * prec - 2)%Z
Hprec:(2 * prec + 2 <= prec')%Z
ex:Z

(Z.max (2 * ex - prec) emin < 2 * ex)%Z -> (Z.max (ex - prec') emin' + ex <= 2 * Z.max (ex - prec) emin - 2)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin <= 0)%Z
Heminprec:(emin' <= emin - prec - 2)%Z \/ (2 * emin' <= emin - 4 * prec - 2)%Z
Hprec:(2 * prec + 2 <= prec')%Z
ex:Z

(2 * ex - prec >= emin)%Z /\ Z.max (2 * ex - prec) emin = (2 * ex - prec)%Z \/ (2 * ex - prec < emin)%Z /\ Z.max (2 * ex - prec) emin = emin -> (Z.max (2 * ex - prec) emin < 2 * ex)%Z -> (Z.max (ex - prec') emin' + ex <= 2 * Z.max (ex - prec) emin - 2)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin <= 0)%Z
Heminprec:(emin' <= emin - prec - 2)%Z \/ (2 * emin' <= emin - 4 * prec - 2)%Z
Hprec:(2 * prec + 2 <= prec')%Z
ex:Z

(ex - prec' >= emin')%Z /\ Z.max (ex - prec') emin' = (ex - prec')%Z \/ (ex - prec' < emin')%Z /\ Z.max (ex - prec') emin' = emin' -> (2 * ex - prec >= emin)%Z /\ Z.max (2 * ex - prec) emin = (2 * ex - prec)%Z \/ (2 * ex - prec < emin)%Z /\ Z.max (2 * ex - prec) emin = emin -> (Z.max (2 * ex - prec) emin < 2 * ex)%Z -> (Z.max (ex - prec') emin' + ex <= 2 * Z.max (ex - prec) emin - 2)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin <= 0)%Z
Heminprec:(emin' <= emin - prec - 2)%Z \/ (2 * emin' <= emin - 4 * prec - 2)%Z
Hprec:(2 * prec + 2 <= prec')%Z
ex:Z

(ex - prec >= emin)%Z /\ Z.max (ex - prec) emin = (ex - prec)%Z \/ (ex - prec < emin)%Z /\ Z.max (ex - prec) emin = emin -> (ex - prec' >= emin')%Z /\ Z.max (ex - prec') emin' = (ex - prec')%Z \/ (ex - prec' < emin')%Z /\ Z.max (ex - prec') emin' = emin' -> (2 * ex - prec >= emin)%Z /\ Z.max (2 * ex - prec) emin = (2 * ex - prec)%Z \/ (2 * ex - prec < emin)%Z /\ Z.max (2 * ex - prec) emin = emin -> (Z.max (2 * ex - prec) emin < 2 * ex)%Z -> (Z.max (ex - prec') emin' + ex <= 2 * Z.max (ex - prec) emin - 2)%Z
omega. Qed.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

forall choice1 choice2 : Z -> bool, (emin <= 0)%Z -> (emin' <= emin - prec - 2)%Z \/ (2 * emin' <= emin - 4 * prec - 2)%Z -> (2 * prec + 2 <= prec')%Z -> forall x : R, FLT_format beta emin prec x -> round_round_eq (FLT_exp emin prec) (FLT_exp emin' prec') choice1 choice2 (sqrt x)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

forall choice1 choice2 : Z -> bool, (emin <= 0)%Z -> (emin' <= emin - prec - 2)%Z \/ (2 * emin' <= emin - 4 * prec - 2)%Z -> (2 * prec + 2 <= prec')%Z -> forall x : R, FLT_format beta emin prec x -> round_round_eq (FLT_exp emin prec) (FLT_exp emin' prec') choice1 choice2 (sqrt x)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hemin:(emin <= 0)%Z
Heminprec:(emin' <= emin - prec - 2)%Z \/ (2 * emin' <= emin - 4 * prec - 2)%Z
Hprec:(2 * prec + 2 <= prec')%Z
x:R
Fx:FLT_format beta emin prec x

round_round_eq (FLT_exp emin prec) (FLT_exp emin' prec') choice1 choice2 (sqrt x)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hemin:(emin <= 0)%Z
Heminprec:(emin' <= emin - prec - 2)%Z \/ (2 * emin' <= emin - 4 * prec - 2)%Z
Hprec:(2 * prec + 2 <= prec')%Z
x:R
Fx:FLT_format beta emin prec x

Valid_exp (FLT_exp emin prec)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hemin:(emin <= 0)%Z
Heminprec:(emin' <= emin - prec - 2)%Z \/ (2 * emin' <= emin - 4 * prec - 2)%Z
Hprec:(2 * prec + 2 <= prec')%Z
x:R
Fx:FLT_format beta emin prec x
Valid_exp (FLT_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hemin:(emin <= 0)%Z
Heminprec:(emin' <= emin - prec - 2)%Z \/ (2 * emin' <= emin - 4 * prec - 2)%Z
Hprec:(2 * prec + 2 <= prec')%Z
x:R
Fx:FLT_format beta emin prec x
round_round_sqrt_hyp (FLT_exp emin prec) (FLT_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hemin:(emin <= 0)%Z
Heminprec:(emin' <= emin - prec - 2)%Z \/ (2 * emin' <= emin - 4 * prec - 2)%Z
Hprec:(2 * prec + 2 <= prec')%Z
x:R
Fx:FLT_format beta emin prec x
generic_format beta (FLT_exp emin prec) x
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hemin:(emin <= 0)%Z
Heminprec:(emin' <= emin - prec - 2)%Z \/ (2 * emin' <= emin - 4 * prec - 2)%Z
Hprec:(2 * prec + 2 <= prec')%Z
x:R
Fx:FLT_format beta emin prec x

Valid_exp (FLT_exp emin prec)
now apply FLT_exp_valid.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hemin:(emin <= 0)%Z
Heminprec:(emin' <= emin - prec - 2)%Z \/ (2 * emin' <= emin - 4 * prec - 2)%Z
Hprec:(2 * prec + 2 <= prec')%Z
x:R
Fx:FLT_format beta emin prec x

Valid_exp (FLT_exp emin' prec')
now apply FLT_exp_valid.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hemin:(emin <= 0)%Z
Heminprec:(emin' <= emin - prec - 2)%Z \/ (2 * emin' <= emin - 4 * prec - 2)%Z
Hprec:(2 * prec + 2 <= prec')%Z
x:R
Fx:FLT_format beta emin prec x

round_round_sqrt_hyp (FLT_exp emin prec) (FLT_exp emin' prec')
now apply FLT_round_round_sqrt_hyp.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Hemin:(emin <= 0)%Z
Heminprec:(emin' <= emin - prec - 2)%Z \/ (2 * emin' <= emin - 4 * prec - 2)%Z
Hprec:(2 * prec + 2 <= prec')%Z
x:R
Fx:FLT_format beta emin prec x

generic_format beta (FLT_exp emin prec) x
now apply generic_format_FLT. Qed. End Double_round_sqrt_FLT. Section Double_round_sqrt_FTZ. Variable emin prec : Z. Variable emin' prec' : Z. Context { prec_gt_0_ : Prec_gt_0 prec }. Context { prec_gt_0_' : Prec_gt_0 prec' }.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(2 * (emin' + prec') <= emin + prec <= 1)%Z -> (2 * prec + 2 <= prec')%Z -> round_round_sqrt_hyp (FTZ_exp emin prec) (FTZ_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(2 * (emin' + prec') <= emin + prec <= 1)%Z -> (2 * prec + 2 <= prec')%Z -> round_round_sqrt_hyp (FTZ_exp emin prec) (FTZ_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(2 * (emin' + prec') <= emin + prec <= 1)%Z
Hprec:(2 * prec + 2 <= prec')%Z

round_round_sqrt_hyp (FTZ_exp emin prec) (FTZ_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(2 * (emin' + prec') <= emin + prec <= 1)%Z
Hprec:(2 * prec + 2 <= prec')%Z

round_round_sqrt_hyp (fun e : Z => if (e - prec <? emin)%Z then (emin + prec - 1)%Z else (e - prec)%Z) (fun e : Z => if (e - prec' <? emin')%Z then (emin' + prec' - 1)%Z else (e - prec')%Z)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':(0 < prec')%Z
Hemin:(2 * (emin' + prec') <= emin + prec <= 1)%Z
Hprec:(2 * prec + 2 <= prec')%Z

round_round_sqrt_hyp (fun e : Z => if (e - prec <? emin)%Z then (emin + prec - 1)%Z else (e - prec)%Z) (fun e : Z => if (e - prec' <? emin')%Z then (emin' + prec' - 1)%Z else (e - prec')%Z)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':(0 < prec')%Z
Hemin:(2 * (emin' + prec') <= emin + prec <= 1)%Z
Hprec:(2 * prec + 2 <= prec')%Z
ex:Z

(2 * (if ex - prec <? emin then emin + prec - 1 else ex - prec) <= (if 2 * ex - prec <? emin then emin + prec - 1 else 2 * ex - prec))%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':(0 < prec')%Z
Hemin:(2 * (emin' + prec') <= emin + prec <= 1)%Z
Hprec:(2 * prec + 2 <= prec')%Z
ex:Z
(2 * (if ex - prec <? emin then emin + prec - 1 else ex - prec) <= (if 2 * ex - 1 - prec <? emin then emin + prec - 1 else 2 * ex - 1 - prec))%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':(0 < prec')%Z
Hemin:(2 * (emin' + prec') <= emin + prec <= 1)%Z
Hprec:(2 * prec + 2 <= prec')%Z
ex:Z
((if 2 * ex - prec <? emin then emin + prec - 1 else 2 * ex - prec) < 2 * ex)%Z -> ((if ex - prec' <? emin' then emin' + prec' - 1 else ex - prec') + ex <= 2 * (if ex - prec <? emin then emin + prec - 1 else ex - prec) - 2)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':(0 < prec')%Z
Hemin:(2 * (emin' + prec') <= emin + prec <= 1)%Z
Hprec:(2 * prec + 2 <= prec')%Z
ex:Z

(2 * (if ex - prec <? emin then emin + prec - 1 else ex - prec) <= (if 2 * ex - prec <? emin then emin + prec - 1 else 2 * ex - prec))%Z
destruct (Z.ltb_spec (ex - prec) emin); destruct (Z.ltb_spec (2 * ex - prec) emin); omega.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':(0 < prec')%Z
Hemin:(2 * (emin' + prec') <= emin + prec <= 1)%Z
Hprec:(2 * prec + 2 <= prec')%Z
ex:Z

(2 * (if ex - prec <? emin then emin + prec - 1 else ex - prec) <= (if 2 * ex - 1 - prec <? emin then emin + prec - 1 else 2 * ex - 1 - prec))%Z
destruct (Z.ltb_spec (ex - prec) emin); destruct (Z.ltb_spec (2 * ex - 1 - prec) emin); omega.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':(0 < prec')%Z
Hemin:(2 * (emin' + prec') <= emin + prec <= 1)%Z
Hprec:(2 * prec + 2 <= prec')%Z
ex:Z

((if 2 * ex - prec <? emin then emin + prec - 1 else 2 * ex - prec) < 2 * ex)%Z -> ((if ex - prec' <? emin' then emin' + prec' - 1 else ex - prec') + ex <= 2 * (if ex - prec <? emin then emin + prec - 1 else ex - prec) - 2)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':(0 < prec')%Z
Hemin:(2 * (emin' + prec') <= emin + prec <= 1)%Z
Hprec:(2 * prec + 2 <= prec')%Z
ex:Z
H:((if 2 * ex - prec <? emin then emin + prec - 1 else 2 * ex - prec) < 2 * ex)%Z

((if ex - prec' <? emin' then emin' + prec' - 1 else ex - prec') + ex <= 2 * (if ex - prec <? emin then emin + prec - 1 else ex - prec) - 2)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':(0 < prec')%Z
Hemin:(2 * (emin' + prec') <= emin + prec <= 1)%Z
Hprec:(2 * prec + 2 <= prec')%Z
ex:Z
H:((if 2 * ex - prec <? emin then emin + prec - 1 else 2 * ex - prec) < 2 * ex)%Z
H':(emin <= 2 * ex - prec)%Z

((if ex - prec' <? emin' then emin' + prec' - 1 else ex - prec') + ex <= 2 * (if ex - prec <? emin then emin + prec - 1 else ex - prec) - 2)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':(0 < prec')%Z
Hemin:(2 * (emin' + prec') <= emin + prec <= 1)%Z
Hprec:(2 * prec + 2 <= prec')%Z
ex:Z
H:((if 2 * ex - prec <? emin then emin + prec - 1 else 2 * ex - prec) < 2 * ex)%Z
H':(2 * ex - prec < emin)%Z
((if ex - prec' <? emin' then emin' + prec' - 1 else ex - prec') + ex <= 2 * (if ex - prec <? emin then emin + prec - 1 else ex - prec) - 2)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':(0 < prec')%Z
Hemin:(2 * (emin' + prec') <= emin + prec <= 1)%Z
Hprec:(2 * prec + 2 <= prec')%Z
ex:Z
H:((if 2 * ex - prec <? emin then emin + prec - 1 else 2 * ex - prec) < 2 * ex)%Z
H':(emin <= 2 * ex - prec)%Z

((if ex - prec' <? emin' then emin' + prec' - 1 else ex - prec') + ex <= 2 * (if ex - prec <? emin then emin + prec - 1 else ex - prec) - 2)%Z
destruct (Z.ltb_spec (ex - prec') emin'); destruct (Z.ltb_spec (ex - prec) emin); omega.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':(0 < prec')%Z
Hemin:(2 * (emin' + prec') <= emin + prec <= 1)%Z
Hprec:(2 * prec + 2 <= prec')%Z
ex:Z
H:((if 2 * ex - prec <? emin then emin + prec - 1 else 2 * ex - prec) < 2 * ex)%Z
H':(2 * ex - prec < emin)%Z

((if ex - prec' <? emin' then emin' + prec' - 1 else ex - prec') + ex <= 2 * (if ex - prec <? emin then emin + prec - 1 else ex - prec) - 2)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':(0 < prec')%Z
Hemin:(2 * (emin' + prec') <= emin + prec <= 1)%Z
Hprec:(2 * prec + 2 <= prec')%Z
ex:Z
H:((if 2 * ex - prec <? emin then emin + prec - 1 else 2 * ex - prec) < 2 * ex)%Z
H':(2 * ex - prec < emin)%Z

False
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':(0 < prec')%Z
Hemin:(2 * (emin' + prec') <= emin + prec <= 1)%Z
Hprec:(2 * prec + 2 <= prec')%Z
ex:Z
H:(emin + prec - 1 < 2 * ex)%Z
H':(2 * ex - prec < emin)%Z

False
omega. Qed.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(4 <= beta)%Z -> forall choice1 choice2 : Z -> bool, (2 * (emin' + prec') <= emin + prec <= 1)%Z -> (2 * prec + 2 <= prec')%Z -> forall x : R, FTZ_format beta emin prec x -> round_round_eq (FTZ_exp emin prec) (FTZ_exp emin' prec') choice1 choice2 (sqrt x)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(4 <= beta)%Z -> forall choice1 choice2 : Z -> bool, (2 * (emin' + prec') <= emin + prec <= 1)%Z -> (2 * prec + 2 <= prec')%Z -> forall x : R, FTZ_format beta emin prec x -> round_round_eq (FTZ_exp emin prec) (FTZ_exp emin' prec') choice1 choice2 (sqrt x)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(4 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(2 * (emin' + prec') <= emin + prec <= 1)%Z
Hprec:(2 * prec + 2 <= prec')%Z
x:R
Fx:FTZ_format beta emin prec x

round_round_eq (FTZ_exp emin prec) (FTZ_exp emin' prec') choice1 choice2 (sqrt x)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(4 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(2 * (emin' + prec') <= emin + prec <= 1)%Z
Hprec:(2 * prec + 2 <= prec')%Z
x:R
Fx:FTZ_format beta emin prec x

Valid_exp (FTZ_exp emin prec)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(4 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(2 * (emin' + prec') <= emin + prec <= 1)%Z
Hprec:(2 * prec + 2 <= prec')%Z
x:R
Fx:FTZ_format beta emin prec x
Valid_exp (FTZ_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(4 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(2 * (emin' + prec') <= emin + prec <= 1)%Z
Hprec:(2 * prec + 2 <= prec')%Z
x:R
Fx:FTZ_format beta emin prec x
round_round_sqrt_hyp (FTZ_exp emin prec) (FTZ_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(4 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(2 * (emin' + prec') <= emin + prec <= 1)%Z
Hprec:(2 * prec + 2 <= prec')%Z
x:R
Fx:FTZ_format beta emin prec x
generic_format beta (FTZ_exp emin prec) x
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(4 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(2 * (emin' + prec') <= emin + prec <= 1)%Z
Hprec:(2 * prec + 2 <= prec')%Z
x:R
Fx:FTZ_format beta emin prec x

Valid_exp (FTZ_exp emin prec)
now apply FTZ_exp_valid.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(4 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(2 * (emin' + prec') <= emin + prec <= 1)%Z
Hprec:(2 * prec + 2 <= prec')%Z
x:R
Fx:FTZ_format beta emin prec x

Valid_exp (FTZ_exp emin' prec')
now apply FTZ_exp_valid.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(4 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(2 * (emin' + prec') <= emin + prec <= 1)%Z
Hprec:(2 * prec + 2 <= prec')%Z
x:R
Fx:FTZ_format beta emin prec x

round_round_sqrt_hyp (FTZ_exp emin prec) (FTZ_exp emin' prec')
now apply FTZ_round_round_sqrt_hyp.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(4 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(2 * (emin' + prec') <= emin + prec <= 1)%Z
Hprec:(2 * prec + 2 <= prec')%Z
x:R
Fx:FTZ_format beta emin prec x

generic_format beta (FTZ_exp emin prec) x
now apply generic_format_FTZ. Qed. End Double_round_sqrt_FTZ. Section Double_round_sqrt_radix_ge_4. Definition round_round_sqrt_radix_ge_4_hyp fexp1 fexp2 := (forall ex, (2 * fexp1 ex <= fexp1 (2 * ex))%Z) /\ (forall ex, (2 * fexp1 ex <= fexp1 (2 * ex - 1))%Z) /\ (forall ex, (fexp1 (2 * ex) < 2 * ex)%Z -> (fexp2 ex + ex <= 2 * fexp1 ex - 1)%Z).
beta:radix

(4 <= beta)%Z -> forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> round_round_sqrt_radix_ge_4_hyp fexp1 fexp2 -> forall x : R, 0 < x -> (fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z -> generic_format beta fexp1 x -> / 2 * ulp beta fexp2 (sqrt x) < Rabs (sqrt x - midp fexp1 (sqrt x))
beta:radix

(4 <= beta)%Z -> forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> round_round_sqrt_radix_ge_4_hyp fexp1 fexp2 -> forall x : R, 0 < x -> (fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z -> generic_format beta fexp1 x -> / 2 * ulp beta fexp2 (sqrt x) < Rabs (sqrt x - midp fexp1 (sqrt x))
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Fx:generic_format beta fexp1 x

/ 2 * ulp beta fexp2 (sqrt x) < Rabs (sqrt x - midp fexp1 (sqrt x))
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Fx:generic_format beta fexp1 x
a:=round beta fexp1 Zfloor (sqrt x):R

/ 2 * ulp beta fexp2 (sqrt x) < Rabs (sqrt x - midp fexp1 (sqrt x))
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Fx:generic_format beta fexp1 x
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R

/ 2 * ulp beta fexp2 (sqrt x) < Rabs (sqrt x - midp fexp1 (sqrt x))
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Fx:generic_format beta fexp1 x
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R

/ 2 * ulp beta fexp2 (sqrt x) < Rabs (sqrt x - midp fexp1 (sqrt x))
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Fx:generic_format beta fexp1 x
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R

/ 2 * ulp beta fexp2 (sqrt x) < Rabs (sqrt x - midp fexp1 (sqrt x))
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Fx:generic_format beta fexp1 x
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R

/ 2 * ulp beta fexp2 (sqrt x) < Rabs (sqrt x - midp fexp1 (sqrt x))
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Fx:generic_format beta fexp1 x
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R

/ 2 * bpow (cexp beta fexp2 (sqrt x)) < Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x))))
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Fx:generic_format beta fexp1 x
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))

False
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Fx:generic_format beta fexp1 x
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))

generic_format beta fexp1 a
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Fx:generic_format beta fexp1 x
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
Fa:generic_format beta fexp1 a
False
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Fx:generic_format beta fexp1 x
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))

generic_format beta fexp1 a
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Fx:generic_format beta fexp1 x
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))

generic_format beta fexp1 (round beta fexp1 Zfloor (sqrt x))
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Fx:generic_format beta fexp1 x
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))

Valid_exp fexp1
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Fx:generic_format beta fexp1 x
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
Valid_rnd Zfloor
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Fx:generic_format beta fexp1 x
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))

Valid_exp fexp1
exact Vfexp1.
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Fx:generic_format beta fexp1 x
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))

Valid_rnd Zfloor
now apply valid_rnd_DN.
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Fx:generic_format beta fexp1 x
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
Fa:generic_format beta fexp1 a

False
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))

generic_format beta fexp1 x -> generic_format beta fexp1 a -> False
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))

x = IZR (Ztrunc (x * bpow (- fexp1 (mag x)))) * bpow (fexp1 (mag x)) -> a = IZR (Ztrunc (a * bpow (- fexp1 (mag a)))) * bpow (fexp1 (mag a)) -> False
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z

x = IZR mx * bpow (fexp1 (mag x)) -> a = IZR (Ztrunc (a * bpow (- fexp1 (mag a)))) * bpow (fexp1 (mag a)) -> False
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z

x = IZR mx * bpow (fexp1 (mag x)) -> a = IZR ma * bpow (fexp1 (mag a)) -> False
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))

False
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))

0 <= a
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
False
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))

0 <= a
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))

round beta fexp1 Zfloor 0 <= a
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))

Valid_exp fexp1
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Valid_rnd Zfloor
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
0 <= sqrt x
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))

Valid_exp fexp1
exact Vfexp1.
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))

Valid_rnd Zfloor
now apply valid_rnd_DN.
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))

0 <= sqrt x
apply sqrt_pos.
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a

False
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a

0 < / 2 * u1
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
False
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a

0 < / 2 * u1
apply Rmult_lt_0_compat; [lra|apply bpow_gt_0].
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1

False
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1

0 < / 2 * u2
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
False
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1

0 < / 2 * u2
apply Rmult_lt_0_compat; [lra|apply bpow_gt_0].
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2

False
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2

0 < b
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
False
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2

0 < b
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2

0 < / 2 * (u1 - u2)
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2

/ 2 * 0 < / 2 * (u1 - u2)
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2

0 < u1 - u2
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2

u2 < u1
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2

bpow (fexp2 (mag (sqrt x))) < bpow (fexp1 (mag (sqrt x)))
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2

(fexp2 (mag (sqrt x)) < fexp1 (mag (sqrt x)))%Z
omega.
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b

False
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b

0 < b'
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
False
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b

0 < b'
now unfold b'; rewrite Rmult_plus_distr_l; apply Rplus_lt_0_compat.
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'

False
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'

sqrt x <= a + b'
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
False
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'

sqrt x <= a + b'
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'

sqrt x - / 2 * u1 - a <= / 2 * u2
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'

sqrt x - (a + / 2 * u1) <= / 2 * u2
now apply Rabs_le_inv.
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'

False
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'

a + b <= sqrt x
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
False
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'

a + b <= sqrt x
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'

- / 2 * u2 <= - a - / 2 * u1 + sqrt x
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'

- / 2 * u2 <= sqrt x - (a + / 2 * u1)
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'

- (/ 2 * u2) <= sqrt x - (a + / 2 * u1)
now apply Rabs_le_inv in H; destruct H.
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x

False
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z

False
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z

(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
False
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z

(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:mag x = (2 * mag (sqrt x) - 1)%Z

(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:mag x = (2 * mag (sqrt x))%Z
(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:mag x = (2 * mag (sqrt x) - 1)%Z

(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:mag x = (2 * mag (sqrt x) - 1)%Z

(fexp1 (mag x) < mag x)%Z
now apply mag_generic_gt; [|apply Rgt_not_eq|].
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:mag x = (2 * mag (sqrt x))%Z

(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:mag x = (2 * mag (sqrt x))%Z

(fexp1 (mag x) < mag x)%Z
now apply mag_generic_gt; [|apply Rgt_not_eq|].
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z

False
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z

a * a + u1 * a - u2 * a + b * b <= x
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
False
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z

a * a + u1 * a - u2 * a + b * b <= x
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z

(a + b) * (a + b) <= x
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z

(a + b) * (a + b) <= sqrt x * sqrt x
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
H':0 <= a + b

(a + b) * (a + b) <= sqrt x * sqrt x
now apply Rmult_le_compat.
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x

False
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x

x <= a * a + u1 * a + u2 * a + b' * b'
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
False
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x

x <= a * a + u1 * a + u2 * a + b' * b'
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x

x <= (a + b') * (a + b')
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x

sqrt x * sqrt x <= (a + b') * (a + b')
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
H':0 <= sqrt x

sqrt x * sqrt x <= (a + b') * (a + b')
now apply Rmult_le_compat.
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'

False
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0

False
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
False
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0

False
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0

0 < 0
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0

b * b <= 0
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0

b * b <= x
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0
x <= 0
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0

b * b <= x
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0

0 + 0 + - 0 + b * b <= x -> b * b <= x
now rewrite Ropp_0; do 3 rewrite Rplus_0_l.
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0

x <= 0
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0

IZR mx * bpow (fexp1 (mag x)) <= 0
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0

IZR mx * bpow (fexp1 (mag x)) * bpow (- fexp1 (mag x)) <= 0 * bpow (- fexp1 (mag x))
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0

IZR mx <= 0
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0

IZR (Ztrunc (x * bpow (- fexp1 (mag x)))) <= 0
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0

IZR (Zfloor (x * bpow (- fexp1 (mag x)))) <= 0
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0

Zfloor (x * bpow (- fexp1 (mag x))) = 0%Z
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0

0 <= x * bpow (- fexp1 (mag x)) < IZR (0 + 1)
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0

x * bpow (- fexp1 (mag x)) < 1
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0

x * bpow (- fexp1 (mag x)) * bpow (fexp1 (mag x)) < 1 * bpow (fexp1 (mag x))
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0

x < bpow (fexp1 (mag x))
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0

x < bpow (2 * fexp1 (mag (sqrt x)))
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0

x < bpow (fexp1 (mag (sqrt x)) + fexp1 (mag (sqrt x)))
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0

x < bpow (fexp1 (mag (sqrt x))) * bpow (fexp1 (mag (sqrt x)))
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0

sqrt x * sqrt x < bpow (fexp1 (mag (sqrt x))) * bpow (fexp1 (mag (sqrt x)))
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0

sqrt x < bpow (fexp1 (mag (sqrt x)))
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0

b' < bpow (fexp1 (mag (sqrt x)))
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0

/ 2 * (u1 + u2) < u1
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0

/ 2 * (u1 + u2) < / 2 * (u1 + u1)
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Za:a = 0

u2 < u1
unfold u2, u1, ulp, cexp; apply bpow_lt; omega.
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0

False
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a

False
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a

mag a = mag (sqrt x)
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
False
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a

mag a = mag (sqrt x)
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a

Valid_exp fexp1
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
0 < round beta fexp1 Zfloor (sqrt x)
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a

Valid_exp fexp1
exact Vfexp1.
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a

0 < round beta fexp1 Zfloor (sqrt x)
now fold a.
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

False
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

0 < - (u2 * a) + b * b
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b
False
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

0 < - (u2 * a) + b * b
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

u2 * a < b ^ 2
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

u2 * a < u2 ^ 2 * (/ 2) ^ 2 - 2 * u2 * (/ 2) ^ 2 * u1 + (/ 2) ^ 2 * u1 ^ 2
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

(2 * u2 * a + u2 * u1) / 2 < (2 * u2 ^ 2 + 2 * u1 ^ 2) / 8
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

u2 * (a + / 2 * u1) < (2 * u2 ^ 2 + 2 * u1 ^ 2) / 8
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

u2 * (a + / 2 * u1) < / 4 * (u2 ^ 2 + u1 ^ 2)
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

u2 * (a + / 2 * u1) < u2 * bpow (mag (sqrt x))
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
u2 * bpow (mag (sqrt x)) <= / 4 * (u2 ^ 2 + u1 ^ 2)
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

u2 * (a + / 2 * u1) < u2 * bpow (mag (sqrt x))
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

a + / 2 * u1 < bpow (mag (sqrt x))
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

a + / 2 * bpow (fexp1 (mag a)) < bpow (mag a)
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

a + / 2 * bpow (fexp1 (mag a)) < a + ulp beta fexp1 a
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
a + ulp beta fexp1 a <= bpow (mag a)
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

a + / 2 * bpow (fexp1 (mag a)) < a + ulp beta fexp1 a
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

/ 2 * bpow (fexp1 (mag a)) < ulp beta fexp1 a
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

/ 2 * bpow (fexp1 (mag a)) < 1 * ulp beta fexp1 a
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

/ 2 * bpow (fexp1 (mag a)) < 1 * bpow (cexp beta fexp1 a)
apply Rmult_lt_compat_r; [apply bpow_gt_0|lra].
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

a + ulp beta fexp1 a <= bpow (mag a)
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

a < bpow (mag a)
apply Rabs_lt_inv, bpow_mag_gt.
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

u2 * bpow (mag (sqrt x)) <= / 4 * (u2 ^ 2 + u1 ^ 2)
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

u2 * bpow (mag (sqrt x)) <= bpow (-1) * u1 ^ 2
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
bpow (-1) * u1 ^ 2 <= / 4 * (u2 ^ 2 + u1 ^ 2)
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

u2 * bpow (mag (sqrt x)) <= bpow (-1) * u1 ^ 2
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

u2 * bpow (mag (sqrt x)) <= bpow (-1) * (u1 * u1)
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

(fexp2 (mag (sqrt x)) + mag (sqrt x) <= 2 * fexp1 (mag (sqrt x)) - 1)%Z
now apply Hexp.
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

bpow (-1) * u1 ^ 2 <= / 4 * (u2 ^ 2 + u1 ^ 2)
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

0 <= bpow (-1)
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
0 <= u1 ^ 2
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
bpow (-1) <= / 4
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
u1 ^ 2 <= u2 ^ 2 + u1 ^ 2
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

0 <= bpow (-1)
apply bpow_ge_0.
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

0 <= u1 ^ 2
apply pow2_ge_0.
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

bpow (-1) <= / 4
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

/ IZR beta <= / 4
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

4 <= IZR beta
now apply IZR_le.
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

u1 ^ 2 <= u2 ^ 2 + u1 ^ 2
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)

0 <= u2 ^ 2
apply pow2_ge_0.
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

False
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

x <= a * a + u1 * a
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b
Hr':x <= a * a + u1 * a
False
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

x <= a * a + u1 * a
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

x <= a * a + u1 * a
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

x <= (a + u1) * a
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

x <= (a + bpow (fexp1 (mag (sqrt x)))) * a
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

x <= (IZR ma + 1) * bpow (fexp1 (mag (sqrt x))) * (IZR ma * bpow (fexp1 (mag (sqrt x))))
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

x <= IZR ma * ((IZR ma + 1) * bpow (fexp1 (mag (sqrt x)))) * bpow (fexp1 (mag (sqrt x)))
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

x <= IZR ma * (IZR ma + 1) * bpow (2 * fexp1 (mag (sqrt x)))
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

x * bpow (-2 * fexp1 (mag (sqrt x))) <= IZR ma * (IZR ma + 1)
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

IZR mx * bpow (fexp1 (mag x) - 2 * fexp1 (mag (sqrt x))) <= IZR ma * (IZR ma + 1)
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

IZR mx * IZR (beta ^ (fexp1 (mag x) - 2 * fexp1 (mag (sqrt x)))) <= IZR ma * (IZR ma + 1)
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

IZR (mx * beta ^ (fexp1 (mag x) - 2 * fexp1 (mag (sqrt x)))) <= IZR (ma * (ma + 1))
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

IZR (mx * beta ^ (fexp1 (mag x) - 2 * fexp1 (mag (sqrt x)))) < IZR (Z.succ (ma * (ma + 1)))
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

IZR mx * IZR (beta ^ (fexp1 (mag x) - 2 * fexp1 (mag (sqrt x)))) < IZR ma * (IZR ma + 1) + 1
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

IZR mx * bpow (fexp1 (mag x) - 2 * fexp1 (mag (sqrt x))) < IZR ma * (IZR ma + 1) + 1
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

IZR mx * bpow (fexp1 (mag x)) < (IZR ma * (IZR ma + 1) + 1) * bpow (2 * fexp1 (mag (sqrt x)))
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

x < (IZR ma * (IZR ma + 1) + 1) * bpow (2 * fexp1 (mag (sqrt x)))
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

x < (IZR ma * (IZR ma + 1) + 1) * bpow (fexp1 (mag (sqrt x)) + fexp1 (mag (sqrt x)))
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

x < (IZR ma * (IZR ma + 1) + 1) * (bpow (fexp1 (mag (sqrt x))) * bpow (fexp1 (mag (sqrt x))))
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

x < a * a + u1 * a + u1 * u1
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

a * a + u1 * a + u2 * a + b' * b' < a * a + u1 * a + u1 * u1
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

u2 * a + b' * b' < u1 * u1
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

u2 * a + u2 * / 2 * u1 < u2 * / 2 * u1 - b' ^ 2 + u1 ^ 2
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

(a + / 2 * u1) * u2 < u2 * / 2 * u1 - b' ^ 2 + u1 ^ 2
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

(a + / 2 * u1) * u2 < bpow (mag (sqrt x)) * u2
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b
bpow (mag (sqrt x)) * u2 <= u2 * / 2 * u1 - b' ^ 2 + u1 ^ 2
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

(a + / 2 * u1) * u2 < bpow (mag (sqrt x)) * u2
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

a + / 2 * u1 < bpow (mag (sqrt x))
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

a + u1 <= bpow (mag (sqrt x))
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

a + bpow (cexp beta fexp1 (sqrt x)) <= bpow (mag (sqrt x))
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

a + bpow (cexp beta fexp1 a) <= bpow (mag (sqrt x))
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

a + ulp beta fexp1 a <= bpow (mag (sqrt x))
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

0 < a
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b
generic_format beta fexp1 a
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b
a < bpow (mag (sqrt x))
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

0 < a
exact Pa.
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

generic_format beta fexp1 a
now apply round_DN_pt.
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

a < bpow (mag (sqrt x))
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

a <= sqrt x
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b
sqrt x < bpow (mag (sqrt x))
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

a <= sqrt x
now apply round_DN_pt.
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

sqrt x < bpow (mag (sqrt x))
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

Rabs (sqrt x) < bpow (mag (sqrt x))
apply bpow_mag_gt.
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

bpow (mag (sqrt x)) * u2 <= u2 * / 2 * u1 - b' ^ 2 + u1 ^ 2
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

bpow (mag (sqrt x)) * u2 <= / 2 * u1 ^ 2
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b
/ 2 * u1 ^ 2 <= u2 * / 2 * u1 - b' ^ 2 + u1 ^ 2
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

bpow (mag (sqrt x)) * u2 <= / 2 * u1 ^ 2
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

bpow (mag (sqrt x)) * u2 <= bpow (-1) * u1 ^ 2
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b
bpow (-1) * u1 ^ 2 <= / 2 * u1 ^ 2
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

bpow (mag (sqrt x)) * u2 <= bpow (-1) * u1 ^ 2
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

bpow (mag (sqrt x)) * u2 <= bpow (-1) * (u1 * u1)
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

bpow (mag (sqrt x)) * bpow (fexp2 (mag (sqrt x))) <= bpow (-1) * (bpow (fexp1 (mag (sqrt x))) * bpow (fexp1 (mag (sqrt x))))
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

bpow (mag (sqrt x) + fexp2 (mag (sqrt x))) <= bpow (2 * fexp1 (mag (sqrt x)) - 1)
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

(mag (sqrt x) + fexp2 (mag (sqrt x)) <= 2 * fexp1 (mag (sqrt x)) - 1)%Z
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

(fexp2 (mag (sqrt x)) + mag (sqrt x) <= 2 * fexp1 (mag (sqrt x)) - 1)%Z
now apply Hexp.
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

bpow (-1) * u1 ^ 2 <= / 2 * u1 ^ 2
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

bpow (-1) <= / 2
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

/ IZR (beta * 1) <= / 2
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

/ IZR beta <= / 2
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

2 <= IZR beta
apply IZR_le; omega.
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

/ 2 * u1 ^ 2 <= u2 * / 2 * u1 - b' ^ 2 + u1 ^ 2
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

u2 ^ 2 < u1 ^ 2
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b

u2 * u2 < u1 * u1
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b
H':0 <= u2

u2 * u2 < u1 * u1
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag (sqrt x)))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b
H':0 <= u2

u2 < u1
unfold u1, u2, ulp, cexp; apply bpow_lt; omega.
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b
Hr':x <= a * a + u1 * a

False
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b
Hr':x <= a * a + u1 * a

a * a + u1 * a < a * a + u1 * a
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b
Hr':x <= a * a + u1 * a

a * a + u1 * a < a * a + u1 * a - u2 * a + b * b
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b
Hr':x <= a * a + u1 * a
a * a + u1 * a - u2 * a + b * b <= a * a + u1 * a
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b
Hr':x <= a * a + u1 * a

a * a + u1 * a < a * a + u1 * a - u2 * a + b * b
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b
Hr':x <= a * a + u1 * a

a * a + u1 * a + 0 < a * a + u1 * a - u2 * a + b * b
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b
Hr':x <= a * a + u1 * a

a * a + u1 * a + 0 < a * a + u1 * a + (- (u2 * a) + b * b)
now apply Rplus_lt_compat_l.
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Px:0 < x
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
a:=round beta fexp1 Zfloor (sqrt x):R
u1:=bpow (fexp1 (mag (sqrt x))):R
u2:=bpow (fexp2 (mag (sqrt x))):R
b:=/ 2 * (u1 - u2):R
b':=/ 2 * (u1 + u2):R
H:Rabs (sqrt x - (round beta fexp1 Zfloor (sqrt x) + / 2 * bpow (cexp beta fexp1 (sqrt x)))) <= / 2 * bpow (cexp beta fexp2 (sqrt x))
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
ma:=Ztrunc (a * bpow (- fexp1 (mag a))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fa:a = IZR ma * bpow (fexp1 (mag a))
Nna:0 <= a
Phu1:0 < / 2 * u1
Phu2:0 < / 2 * u2
Pb:0 < b
Pb':0 < b'
Hr:sqrt x <= a + b'
Hl:a + b <= sqrt x
Hf1:(2 * fexp1 (mag (sqrt x)) <= fexp1 (mag x))%Z
Hlx:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
Hsl:a * a + u1 * a - u2 * a + b * b <= x
Hsr:x <= a * a + u1 * a + u2 * a + b' * b'
Nza:a <> 0
Pa:0 < a
Hla:mag a = mag (sqrt x)
Hl':0 < - (u2 * a) + b * b
Hr':x <= a * a + u1 * a

a * a + u1 * a - u2 * a + b * b <= a * a + u1 * a
now apply Rle_trans with x. Qed.
beta:radix

(4 <= beta)%Z -> forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall choice1 choice2 : Z -> bool, round_round_sqrt_radix_ge_4_hyp fexp1 fexp2 -> forall x : R, generic_format beta fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 (sqrt x)
beta:radix

(4 <= beta)%Z -> forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall choice1 choice2 : Z -> bool, round_round_sqrt_radix_ge_4_hyp fexp1 fexp2 -> forall x : R, generic_format beta fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 (sqrt x)
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x

round_round_eq fexp1 fexp2 choice1 choice2 (sqrt x)
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (sqrt x)) = round beta fexp1 (Znearest choice1) (sqrt x)
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Npx:x <= 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (sqrt x)) = round beta fexp1 (Znearest choice1) (sqrt x)
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (sqrt x)) = round beta fexp1 (Znearest choice1) (sqrt x)
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Npx:x <= 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (sqrt x)) = round beta fexp1 (Znearest choice1) (sqrt x)
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Npx:x <= 0

sqrt x = 0
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Npx:x <= 0
Hs:sqrt x = 0
round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (sqrt x)) = round beta fexp1 (Znearest choice1) (sqrt x)
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Npx:x <= 0

sqrt x = 0
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Npx:x <= 0
Zx:x = 0

sqrt x = 0
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Npx:x <= 0
Nzx:x <> 0
sqrt x = 0
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Npx:x <= 0
Zx:x = 0

sqrt x = 0
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Npx:x <= 0
Zx:x = 0

sqrt 0 = 0
exact sqrt_0.
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Npx:x <= 0
Nzx:x <> 0

sqrt x = 0
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Npx:x <= 0
Nzx:x <> 0

match Rcase_abs x with | left _ => 0 | right a => Rsqrt {| nonneg := x; cond_nonneg := Rge_le x 0 a |} end = 0
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Npx:x <= 0
Nzx:x <> 0
r:x < 0

0 = 0
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Npx:x <= 0
Nzx:x <> 0
r:x >= 0
Rsqrt {| nonneg := x; cond_nonneg := Rge_le x 0 r |} = 0
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Npx:x <= 0
Nzx:x <> 0
r:x < 0

0 = 0
reflexivity.
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Npx:x <= 0
Nzx:x <> 0
r:x >= 0

Rsqrt {| nonneg := x; cond_nonneg := Rge_le x 0 r |} = 0
casetype False; lra.
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Npx:x <= 0
Hs:sqrt x = 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (sqrt x)) = round beta fexp1 (Znearest choice1) (sqrt x)
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Npx:x <= 0
Hs:sqrt x = 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) 0) = round beta fexp1 (Znearest choice1) 0
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Npx:x <= 0
Hs:sqrt x = 0

round beta fexp1 (Znearest choice1) 0 = round beta fexp1 (Znearest choice1) 0
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Npx:x <= 0
Hs:sqrt x = 0
Valid_rnd (Znearest choice2)
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Npx:x <= 0
Hs:sqrt x = 0

round beta fexp1 (Znearest choice1) 0 = round beta fexp1 (Znearest choice1) 0
reflexivity.
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Npx:x <= 0
Hs:sqrt x = 0

Valid_rnd (Znearest choice2)
now apply valid_rnd_N.
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (sqrt x)) = round beta fexp1 (Znearest choice1) (sqrt x)
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (sqrt x)) = round beta fexp1 (Znearest choice1) (sqrt x)
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z

(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (sqrt x)) = round beta fexp1 (Znearest choice1) (sqrt x)
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z

(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hx:x <= 1

(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hx:1 < x
(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hx:x <= 1

(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hx:x <= 1

(mag x <= mag (sqrt x))%Z
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hx:x <= 1

x <= sqrt x
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hx:x <= 1

sqrt x * sqrt x <= sqrt x
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hx:x <= 1

sqrt x * sqrt x <= sqrt x * 1
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hx:x <= 1

0 <= sqrt x
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hx:x <= 1
sqrt x <= 1
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hx:x <= 1

0 <= sqrt x
apply sqrt_pos.
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hx:x <= 1

sqrt x <= 1
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hx:x <= 1

sqrt x <= sqrt 1
now apply sqrt_le_1_alt.
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hx:1 < x

(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hx:1 < x

(2 * fexp1 1 <= fexp1 (2 * 1 - 1))%Z -> (fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hx:1 < x

(2 * fexp1 1 <= fexp1 1)%Z -> (fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hx:1 < x
Hexp10:(2 * fexp1 1 <= fexp1 1)%Z

(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hx:1 < x
Hf0:(fexp1 1 < 1)%Z

(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hx:1 < x
Hf0:(fexp1 1 < 1)%Z

(1 <= mag (sqrt x))%Z
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hx:1 < x
Hf0:(fexp1 1 < 1)%Z

bpow (1 - 1) <= Rabs (sqrt x)
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hx:1 < x
Hf0:(fexp1 1 < 1)%Z

bpow 0 <= Rabs (sqrt x)
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hx:1 < x
Hf0:(fexp1 1 < 1)%Z

1 <= Rabs (sqrt x)
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hx:1 < x
Hf0:(fexp1 1 < 1)%Z

1 <= sqrt x
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hx:1 < x
Hf0:(fexp1 1 < 1)%Z

sqrt 1 <= sqrt x
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hx:1 < x
Hf0:(fexp1 1 < 1)%Z

1 <= x
now apply Rlt_le.
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (sqrt x)) = round beta fexp1 (Znearest choice1) (sqrt x)
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z

(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (sqrt x)) = round beta fexp1 (Znearest choice1) (sqrt x)
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z

(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z

(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
H:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z

(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
Hlx:mag x = (2 * mag (sqrt x) - 1)%Z

(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
Hlx:mag x = (2 * mag (sqrt x))%Z
(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
Hlx:mag x = (2 * mag (sqrt x) - 1)%Z

(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
Hlx:mag x = (2 * mag (sqrt x) - 1)%Z

(fexp1 (mag x) < mag x)%Z
now apply mag_generic_gt; [|apply Rgt_not_eq|].
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
Hlx:mag x = (2 * mag (sqrt x))%Z

(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
Hlx:mag x = (2 * mag (sqrt x))%Z

(fexp1 (mag x) < mag x)%Z
now apply mag_generic_gt; [|apply Rgt_not_eq|].
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
H:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z

(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
H:(fexp1 (2 * mag (sqrt x)) < 2 * mag (sqrt x))%Z

(fexp2 (mag (sqrt x)) + mag (sqrt x) <= 2 * fexp1 (mag (sqrt x)) - 1)%Z -> (fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
omega.
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (sqrt x)) = round beta fexp1 (Znearest choice1) (sqrt x)
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z

Valid_exp fexp1
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Valid_exp fexp2
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
0 < sqrt x
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
(fexp1 (mag (sqrt x)) <= mag (sqrt x))%Z
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Rabs (sqrt x - midp fexp1 (sqrt x)) <= / 2 * ulp beta fexp2 (sqrt x) -> round_round_eq fexp1 fexp2 choice1 choice2 (sqrt x)
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z

Valid_exp fexp1
exact Vfexp1.
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z

Valid_exp fexp2
exact Vfexp2.
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z

0 < sqrt x
now apply sqrt_lt_R0.
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z

(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
omega.
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z

(fexp1 (mag (sqrt x)) <= mag (sqrt x))%Z
omega.
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z

Rabs (sqrt x - midp fexp1 (sqrt x)) <= / 2 * ulp beta fexp2 (sqrt x) -> round_round_eq fexp1 fexp2 choice1 choice2 (sqrt x)
beta:radix
Hbeta:(4 <= beta)%Z
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_sqrt_radix_ge_4_hyp fexp1 fexp2
x:R
Fx:generic_format beta fexp1 x
Px:0 < x
Hfx:(fexp1 (mag x) < mag x)%Z
Hfsx:(fexp1 (mag (sqrt x)) < mag (sqrt x))%Z
Hf2:(fexp2 (mag (sqrt x)) <= fexp1 (mag (sqrt x)) - 1)%Z
Hmid:Rabs (sqrt x - midp fexp1 (sqrt x)) <= / 2 * ulp beta fexp2 (sqrt x)

/ 2 * ulp beta fexp2 (sqrt x) < Rabs (sqrt x - midp fexp1 (sqrt x))
apply (round_round_sqrt_radix_ge_4_aux Hbeta fexp1 fexp2 Vfexp1 Vfexp2 Hexp x Px Hf2 Fx). Qed. Section Double_round_sqrt_radix_ge_4_FLX. Variable prec : Z. Variable prec' : Z. Context { prec_gt_0_ : Prec_gt_0 prec }. Context { prec_gt_0_' : Prec_gt_0 prec' }.
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(2 * prec + 1 <= prec')%Z -> round_round_sqrt_radix_ge_4_hyp (FLX_exp prec) (FLX_exp prec')
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(2 * prec + 1 <= prec')%Z -> round_round_sqrt_radix_ge_4_hyp (FLX_exp prec) (FLX_exp prec')
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hprec:(2 * prec + 1 <= prec')%Z

round_round_sqrt_radix_ge_4_hyp (FLX_exp prec) (FLX_exp prec')
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hprec:(2 * prec + 1 <= prec')%Z

round_round_sqrt_radix_ge_4_hyp (fun e : Z => (e - prec)%Z) (fun e : Z => (e - prec')%Z)
beta:radix
prec, prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hprec:(2 * prec + 1 <= prec')%Z

round_round_sqrt_radix_ge_4_hyp (fun e : Z => (e - prec)%Z) (fun e : Z => (e - prec')%Z)
unfold round_round_sqrt_radix_ge_4_hyp; split; [|split]; intro ex; omega. Qed.
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(4 <= beta)%Z -> forall choice1 choice2 : Z -> bool, (2 * prec + 1 <= prec')%Z -> forall x : R, FLX_format beta prec x -> round_round_eq (FLX_exp prec) (FLX_exp prec') choice1 choice2 (sqrt x)
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(4 <= beta)%Z -> forall choice1 choice2 : Z -> bool, (2 * prec + 1 <= prec')%Z -> forall x : R, FLX_format beta prec x -> round_round_eq (FLX_exp prec) (FLX_exp prec') choice1 choice2 (sqrt x)
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(4 <= beta)%Z
choice1, choice2:Z -> bool
Hprec:(2 * prec + 1 <= prec')%Z
x:R
Fx:FLX_format beta prec x

round_round_eq (FLX_exp prec) (FLX_exp prec') choice1 choice2 (sqrt x)
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(4 <= beta)%Z
choice1, choice2:Z -> bool
Hprec:(2 * prec + 1 <= prec')%Z
x:R
Fx:FLX_format beta prec x

(4 <= beta)%Z
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(4 <= beta)%Z
choice1, choice2:Z -> bool
Hprec:(2 * prec + 1 <= prec')%Z
x:R
Fx:FLX_format beta prec x
Valid_exp (FLX_exp prec)
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(4 <= beta)%Z
choice1, choice2:Z -> bool
Hprec:(2 * prec + 1 <= prec')%Z
x:R
Fx:FLX_format beta prec x
Valid_exp (FLX_exp prec')
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(4 <= beta)%Z
choice1, choice2:Z -> bool
Hprec:(2 * prec + 1 <= prec')%Z
x:R
Fx:FLX_format beta prec x
round_round_sqrt_radix_ge_4_hyp (FLX_exp prec) (FLX_exp prec')
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(4 <= beta)%Z
choice1, choice2:Z -> bool
Hprec:(2 * prec + 1 <= prec')%Z
x:R
Fx:FLX_format beta prec x
generic_format beta (FLX_exp prec) x
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(4 <= beta)%Z
choice1, choice2:Z -> bool
Hprec:(2 * prec + 1 <= prec')%Z
x:R
Fx:FLX_format beta prec x

(4 <= beta)%Z
exact Hbeta.
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(4 <= beta)%Z
choice1, choice2:Z -> bool
Hprec:(2 * prec + 1 <= prec')%Z
x:R
Fx:FLX_format beta prec x

Valid_exp (FLX_exp prec)
now apply FLX_exp_valid.
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(4 <= beta)%Z
choice1, choice2:Z -> bool
Hprec:(2 * prec + 1 <= prec')%Z
x:R
Fx:FLX_format beta prec x

Valid_exp (FLX_exp prec')
now apply FLX_exp_valid.
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(4 <= beta)%Z
choice1, choice2:Z -> bool
Hprec:(2 * prec + 1 <= prec')%Z
x:R
Fx:FLX_format beta prec x

round_round_sqrt_radix_ge_4_hyp (FLX_exp prec) (FLX_exp prec')
now apply FLX_round_round_sqrt_radix_ge_4_hyp.
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(4 <= beta)%Z
choice1, choice2:Z -> bool
Hprec:(2 * prec + 1 <= prec')%Z
x:R
Fx:FLX_format beta prec x

generic_format beta (FLX_exp prec) x
now apply generic_format_FLX. Qed. End Double_round_sqrt_radix_ge_4_FLX. Section Double_round_sqrt_radix_ge_4_FLT. Variable emin prec : Z. Variable emin' prec' : Z. Context { prec_gt_0_ : Prec_gt_0 prec }. Context { prec_gt_0_' : Prec_gt_0 prec' }.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(emin <= 0)%Z -> (emin' <= emin - prec - 1)%Z \/ (2 * emin' <= emin - 4 * prec)%Z -> (2 * prec + 1 <= prec')%Z -> round_round_sqrt_radix_ge_4_hyp (FLT_exp emin prec) (FLT_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(emin <= 0)%Z -> (emin' <= emin - prec - 1)%Z \/ (2 * emin' <= emin - 4 * prec)%Z -> (2 * prec + 1 <= prec')%Z -> round_round_sqrt_radix_ge_4_hyp (FLT_exp emin prec) (FLT_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin <= 0)%Z
Heminprec:(emin' <= emin - prec - 1)%Z \/ (2 * emin' <= emin - 4 * prec)%Z
Hprec:(2 * prec + 1 <= prec')%Z

round_round_sqrt_radix_ge_4_hyp (FLT_exp emin prec) (FLT_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin <= 0)%Z
Heminprec:(emin' <= emin - prec - 1)%Z \/ (2 * emin' <= emin - 4 * prec)%Z
Hprec:(2 * prec + 1 <= prec')%Z

round_round_sqrt_radix_ge_4_hyp (fun e : Z => Z.max (e - prec) emin) (fun e : Z => Z.max (e - prec') emin')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin <= 0)%Z
Heminprec:(emin' <= emin - prec - 1)%Z \/ (2 * emin' <= emin - 4 * prec)%Z
Hprec:(2 * prec + 1 <= prec')%Z

round_round_sqrt_radix_ge_4_hyp (fun e : Z => Z.max (e - prec) emin) (fun e : Z => Z.max (e - prec') emin')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin <= 0)%Z
Heminprec:(emin' <= emin - prec - 1)%Z \/ (2 * emin' <= emin - 4 * prec)%Z
Hprec:(2 * prec + 1 <= prec')%Z
ex:Z

(2 * Z.max (ex - prec) emin <= Z.max (2 * ex - prec) emin)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin <= 0)%Z
Heminprec:(emin' <= emin - prec - 1)%Z \/ (2 * emin' <= emin - 4 * prec)%Z
Hprec:(2 * prec + 1 <= prec')%Z
ex:Z
(2 * Z.max (ex - prec) emin <= Z.max (2 * ex - 1 - prec) emin)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin <= 0)%Z
Heminprec:(emin' <= emin - prec - 1)%Z \/ (2 * emin' <= emin - 4 * prec)%Z
Hprec:(2 * prec + 1 <= prec')%Z
ex:Z
(Z.max (2 * ex - prec) emin < 2 * ex)%Z -> (Z.max (ex - prec') emin' + ex <= 2 * Z.max (ex - prec) emin - 1)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin <= 0)%Z
Heminprec:(emin' <= emin - prec - 1)%Z \/ (2 * emin' <= emin - 4 * prec)%Z
Hprec:(2 * prec + 1 <= prec')%Z
ex:Z

(2 * Z.max (ex - prec) emin <= Z.max (2 * ex - prec) emin)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin <= 0)%Z
Heminprec:(emin' <= emin - prec - 1)%Z \/ (2 * emin' <= emin - 4 * prec)%Z
Hprec:(2 * prec + 1 <= prec')%Z
ex:Z

(ex - prec >= emin)%Z /\ Z.max (ex - prec) emin = (ex - prec)%Z \/ (ex - prec < emin)%Z /\ Z.max (ex - prec) emin = emin -> (2 * Z.max (ex - prec) emin <= Z.max (2 * ex - prec) emin)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin <= 0)%Z
Heminprec:(emin' <= emin - prec - 1)%Z \/ (2 * emin' <= emin - 4 * prec)%Z
Hprec:(2 * prec + 1 <= prec')%Z
ex:Z

(2 * ex - prec >= emin)%Z /\ Z.max (2 * ex - prec) emin = (2 * ex - prec)%Z \/ (2 * ex - prec < emin)%Z /\ Z.max (2 * ex - prec) emin = emin -> (ex - prec >= emin)%Z /\ Z.max (ex - prec) emin = (ex - prec)%Z \/ (ex - prec < emin)%Z /\ Z.max (ex - prec) emin = emin -> (2 * Z.max (ex - prec) emin <= Z.max (2 * ex - prec) emin)%Z
omega.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin <= 0)%Z
Heminprec:(emin' <= emin - prec - 1)%Z \/ (2 * emin' <= emin - 4 * prec)%Z
Hprec:(2 * prec + 1 <= prec')%Z
ex:Z

(2 * Z.max (ex - prec) emin <= Z.max (2 * ex - 1 - prec) emin)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin <= 0)%Z
Heminprec:(emin' <= emin - prec - 1)%Z \/ (2 * emin' <= emin - 4 * prec)%Z
Hprec:(2 * prec + 1 <= prec')%Z
ex:Z

(ex - prec >= emin)%Z /\ Z.max (ex - prec) emin = (ex - prec)%Z \/ (ex - prec < emin)%Z /\ Z.max (ex - prec) emin = emin -> (2 * Z.max (ex - prec) emin <= Z.max (2 * ex - 1 - prec) emin)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin <= 0)%Z
Heminprec:(emin' <= emin - prec - 1)%Z \/ (2 * emin' <= emin - 4 * prec)%Z
Hprec:(2 * prec + 1 <= prec')%Z
ex:Z

(2 * ex - 1 - prec >= emin)%Z /\ Z.max (2 * ex - 1 - prec) emin = (2 * ex - 1 - prec)%Z \/ (2 * ex - 1 - prec < emin)%Z /\ Z.max (2 * ex - 1 - prec) emin = emin -> (ex - prec >= emin)%Z /\ Z.max (ex - prec) emin = (ex - prec)%Z \/ (ex - prec < emin)%Z /\ Z.max (ex - prec) emin = emin -> (2 * Z.max (ex - prec) emin <= Z.max (2 * ex - 1 - prec) emin)%Z
omega.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin <= 0)%Z
Heminprec:(emin' <= emin - prec - 1)%Z \/ (2 * emin' <= emin - 4 * prec)%Z
Hprec:(2 * prec + 1 <= prec')%Z
ex:Z

(Z.max (2 * ex - prec) emin < 2 * ex)%Z -> (Z.max (ex - prec') emin' + ex <= 2 * Z.max (ex - prec) emin - 1)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin <= 0)%Z
Heminprec:(emin' <= emin - prec - 1)%Z \/ (2 * emin' <= emin - 4 * prec)%Z
Hprec:(2 * prec + 1 <= prec')%Z
ex:Z

(2 * ex - prec >= emin)%Z /\ Z.max (2 * ex - prec) emin = (2 * ex - prec)%Z \/ (2 * ex - prec < emin)%Z /\ Z.max (2 * ex - prec) emin = emin -> (Z.max (2 * ex - prec) emin < 2 * ex)%Z -> (Z.max (ex - prec') emin' + ex <= 2 * Z.max (ex - prec) emin - 1)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin <= 0)%Z
Heminprec:(emin' <= emin - prec - 1)%Z \/ (2 * emin' <= emin - 4 * prec)%Z
Hprec:(2 * prec + 1 <= prec')%Z
ex:Z

(ex - prec' >= emin')%Z /\ Z.max (ex - prec') emin' = (ex - prec')%Z \/ (ex - prec' < emin')%Z /\ Z.max (ex - prec') emin' = emin' -> (2 * ex - prec >= emin)%Z /\ Z.max (2 * ex - prec) emin = (2 * ex - prec)%Z \/ (2 * ex - prec < emin)%Z /\ Z.max (2 * ex - prec) emin = emin -> (Z.max (2 * ex - prec) emin < 2 * ex)%Z -> (Z.max (ex - prec') emin' + ex <= 2 * Z.max (ex - prec) emin - 1)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin <= 0)%Z
Heminprec:(emin' <= emin - prec - 1)%Z \/ (2 * emin' <= emin - 4 * prec)%Z
Hprec:(2 * prec + 1 <= prec')%Z
ex:Z

(ex - prec >= emin)%Z /\ Z.max (ex - prec) emin = (ex - prec)%Z \/ (ex - prec < emin)%Z /\ Z.max (ex - prec) emin = emin -> (ex - prec' >= emin')%Z /\ Z.max (ex - prec') emin' = (ex - prec')%Z \/ (ex - prec' < emin')%Z /\ Z.max (ex - prec') emin' = emin' -> (2 * ex - prec >= emin)%Z /\ Z.max (2 * ex - prec) emin = (2 * ex - prec)%Z \/ (2 * ex - prec < emin)%Z /\ Z.max (2 * ex - prec) emin = emin -> (Z.max (2 * ex - prec) emin < 2 * ex)%Z -> (Z.max (ex - prec') emin' + ex <= 2 * Z.max (ex - prec) emin - 1)%Z
omega. Qed.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(4 <= beta)%Z -> forall choice1 choice2 : Z -> bool, (emin <= 0)%Z -> (emin' <= emin - prec - 1)%Z \/ (2 * emin' <= emin - 4 * prec)%Z -> (2 * prec + 1 <= prec')%Z -> forall x : R, FLT_format beta emin prec x -> round_round_eq (FLT_exp emin prec) (FLT_exp emin' prec') choice1 choice2 (sqrt x)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(4 <= beta)%Z -> forall choice1 choice2 : Z -> bool, (emin <= 0)%Z -> (emin' <= emin - prec - 1)%Z \/ (2 * emin' <= emin - 4 * prec)%Z -> (2 * prec + 1 <= prec')%Z -> forall x : R, FLT_format beta emin prec x -> round_round_eq (FLT_exp emin prec) (FLT_exp emin' prec') choice1 choice2 (sqrt x)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(4 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin <= 0)%Z
Heminprec:(emin' <= emin - prec - 1)%Z \/ (2 * emin' <= emin - 4 * prec)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x:R
Fx:FLT_format beta emin prec x

round_round_eq (FLT_exp emin prec) (FLT_exp emin' prec') choice1 choice2 (sqrt x)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(4 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin <= 0)%Z
Heminprec:(emin' <= emin - prec - 1)%Z \/ (2 * emin' <= emin - 4 * prec)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x:R
Fx:FLT_format beta emin prec x

(4 <= beta)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(4 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin <= 0)%Z
Heminprec:(emin' <= emin - prec - 1)%Z \/ (2 * emin' <= emin - 4 * prec)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x:R
Fx:FLT_format beta emin prec x
Valid_exp (FLT_exp emin prec)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(4 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin <= 0)%Z
Heminprec:(emin' <= emin - prec - 1)%Z \/ (2 * emin' <= emin - 4 * prec)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x:R
Fx:FLT_format beta emin prec x
Valid_exp (FLT_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(4 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin <= 0)%Z
Heminprec:(emin' <= emin - prec - 1)%Z \/ (2 * emin' <= emin - 4 * prec)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x:R
Fx:FLT_format beta emin prec x
round_round_sqrt_radix_ge_4_hyp (FLT_exp emin prec) (FLT_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(4 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin <= 0)%Z
Heminprec:(emin' <= emin - prec - 1)%Z \/ (2 * emin' <= emin - 4 * prec)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x:R
Fx:FLT_format beta emin prec x
generic_format beta (FLT_exp emin prec) x
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(4 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin <= 0)%Z
Heminprec:(emin' <= emin - prec - 1)%Z \/ (2 * emin' <= emin - 4 * prec)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x:R
Fx:FLT_format beta emin prec x

(4 <= beta)%Z
exact Hbeta.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(4 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin <= 0)%Z
Heminprec:(emin' <= emin - prec - 1)%Z \/ (2 * emin' <= emin - 4 * prec)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x:R
Fx:FLT_format beta emin prec x

Valid_exp (FLT_exp emin prec)
now apply FLT_exp_valid.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(4 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin <= 0)%Z
Heminprec:(emin' <= emin - prec - 1)%Z \/ (2 * emin' <= emin - 4 * prec)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x:R
Fx:FLT_format beta emin prec x

Valid_exp (FLT_exp emin' prec')
now apply FLT_exp_valid.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(4 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin <= 0)%Z
Heminprec:(emin' <= emin - prec - 1)%Z \/ (2 * emin' <= emin - 4 * prec)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x:R
Fx:FLT_format beta emin prec x

round_round_sqrt_radix_ge_4_hyp (FLT_exp emin prec) (FLT_exp emin' prec')
now apply FLT_round_round_sqrt_radix_ge_4_hyp.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(4 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(emin <= 0)%Z
Heminprec:(emin' <= emin - prec - 1)%Z \/ (2 * emin' <= emin - 4 * prec)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x:R
Fx:FLT_format beta emin prec x

generic_format beta (FLT_exp emin prec) x
now apply generic_format_FLT. Qed. End Double_round_sqrt_radix_ge_4_FLT. Section Double_round_sqrt_radix_ge_4_FTZ. Variable emin prec : Z. Variable emin' prec' : Z. Context { prec_gt_0_ : Prec_gt_0 prec }. Context { prec_gt_0_' : Prec_gt_0 prec' }.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(2 * (emin' + prec') <= emin + prec <= 1)%Z -> (2 * prec + 1 <= prec')%Z -> round_round_sqrt_radix_ge_4_hyp (FTZ_exp emin prec) (FTZ_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(2 * (emin' + prec') <= emin + prec <= 1)%Z -> (2 * prec + 1 <= prec')%Z -> round_round_sqrt_radix_ge_4_hyp (FTZ_exp emin prec) (FTZ_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(2 * (emin' + prec') <= emin + prec <= 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z

round_round_sqrt_radix_ge_4_hyp (FTZ_exp emin prec) (FTZ_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(2 * (emin' + prec') <= emin + prec <= 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z

round_round_sqrt_radix_ge_4_hyp (fun e : Z => if (e - prec <? emin)%Z then (emin + prec - 1)%Z else (e - prec)%Z) (fun e : Z => if (e - prec' <? emin')%Z then (emin' + prec' - 1)%Z else (e - prec')%Z)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':(0 < prec')%Z
Hemin:(2 * (emin' + prec') <= emin + prec <= 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z

round_round_sqrt_radix_ge_4_hyp (fun e : Z => if (e - prec <? emin)%Z then (emin + prec - 1)%Z else (e - prec)%Z) (fun e : Z => if (e - prec' <? emin')%Z then (emin' + prec' - 1)%Z else (e - prec')%Z)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':(0 < prec')%Z
Hemin:(2 * (emin' + prec') <= emin + prec <= 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z
ex:Z

(2 * (if ex - prec <? emin then emin + prec - 1 else ex - prec) <= (if 2 * ex - prec <? emin then emin + prec - 1 else 2 * ex - prec))%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':(0 < prec')%Z
Hemin:(2 * (emin' + prec') <= emin + prec <= 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z
ex:Z
(2 * (if ex - prec <? emin then emin + prec - 1 else ex - prec) <= (if 2 * ex - 1 - prec <? emin then emin + prec - 1 else 2 * ex - 1 - prec))%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':(0 < prec')%Z
Hemin:(2 * (emin' + prec') <= emin + prec <= 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z
ex:Z
((if 2 * ex - prec <? emin then emin + prec - 1 else 2 * ex - prec) < 2 * ex)%Z -> ((if ex - prec' <? emin' then emin' + prec' - 1 else ex - prec') + ex <= 2 * (if ex - prec <? emin then emin + prec - 1 else ex - prec) - 1)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':(0 < prec')%Z
Hemin:(2 * (emin' + prec') <= emin + prec <= 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z
ex:Z

(2 * (if ex - prec <? emin then emin + prec - 1 else ex - prec) <= (if 2 * ex - prec <? emin then emin + prec - 1 else 2 * ex - prec))%Z
destruct (Z.ltb_spec (ex - prec) emin); destruct (Z.ltb_spec (2 * ex - prec) emin); omega.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':(0 < prec')%Z
Hemin:(2 * (emin' + prec') <= emin + prec <= 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z
ex:Z

(2 * (if ex - prec <? emin then emin + prec - 1 else ex - prec) <= (if 2 * ex - 1 - prec <? emin then emin + prec - 1 else 2 * ex - 1 - prec))%Z
destruct (Z.ltb_spec (ex - prec) emin); destruct (Z.ltb_spec (2 * ex - 1 - prec) emin); omega.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':(0 < prec')%Z
Hemin:(2 * (emin' + prec') <= emin + prec <= 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z
ex:Z

((if 2 * ex - prec <? emin then emin + prec - 1 else 2 * ex - prec) < 2 * ex)%Z -> ((if ex - prec' <? emin' then emin' + prec' - 1 else ex - prec') + ex <= 2 * (if ex - prec <? emin then emin + prec - 1 else ex - prec) - 1)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':(0 < prec')%Z
Hemin:(2 * (emin' + prec') <= emin + prec <= 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z
ex:Z
H:((if 2 * ex - prec <? emin then emin + prec - 1 else 2 * ex - prec) < 2 * ex)%Z

((if ex - prec' <? emin' then emin' + prec' - 1 else ex - prec') + ex <= 2 * (if ex - prec <? emin then emin + prec - 1 else ex - prec) - 1)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':(0 < prec')%Z
Hemin:(2 * (emin' + prec') <= emin + prec <= 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z
ex:Z
H:((if 2 * ex - prec <? emin then emin + prec - 1 else 2 * ex - prec) < 2 * ex)%Z
H':(emin <= 2 * ex - prec)%Z

((if ex - prec' <? emin' then emin' + prec' - 1 else ex - prec') + ex <= 2 * (if ex - prec <? emin then emin + prec - 1 else ex - prec) - 1)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':(0 < prec')%Z
Hemin:(2 * (emin' + prec') <= emin + prec <= 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z
ex:Z
H:((if 2 * ex - prec <? emin then emin + prec - 1 else 2 * ex - prec) < 2 * ex)%Z
H':(2 * ex - prec < emin)%Z
((if ex - prec' <? emin' then emin' + prec' - 1 else ex - prec') + ex <= 2 * (if ex - prec <? emin then emin + prec - 1 else ex - prec) - 1)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':(0 < prec')%Z
Hemin:(2 * (emin' + prec') <= emin + prec <= 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z
ex:Z
H:((if 2 * ex - prec <? emin then emin + prec - 1 else 2 * ex - prec) < 2 * ex)%Z
H':(emin <= 2 * ex - prec)%Z

((if ex - prec' <? emin' then emin' + prec' - 1 else ex - prec') + ex <= 2 * (if ex - prec <? emin then emin + prec - 1 else ex - prec) - 1)%Z
destruct (Z.ltb_spec (ex - prec') emin'); destruct (Z.ltb_spec (ex - prec) emin); omega.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':(0 < prec')%Z
Hemin:(2 * (emin' + prec') <= emin + prec <= 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z
ex:Z
H:((if 2 * ex - prec <? emin then emin + prec - 1 else 2 * ex - prec) < 2 * ex)%Z
H':(2 * ex - prec < emin)%Z

((if ex - prec' <? emin' then emin' + prec' - 1 else ex - prec') + ex <= 2 * (if ex - prec <? emin then emin + prec - 1 else ex - prec) - 1)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':(0 < prec')%Z
Hemin:(2 * (emin' + prec') <= emin + prec <= 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z
ex:Z
H:((if 2 * ex - prec <? emin then emin + prec - 1 else 2 * ex - prec) < 2 * ex)%Z
H':(2 * ex - prec < emin)%Z

False
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':(0 < prec')%Z
Hemin:(2 * (emin' + prec') <= emin + prec <= 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z
ex:Z
H:(emin + prec - 1 < 2 * ex)%Z
H':(2 * ex - prec < emin)%Z

False
omega. Qed.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(4 <= beta)%Z -> forall choice1 choice2 : Z -> bool, (2 * (emin' + prec') <= emin + prec <= 1)%Z -> (2 * prec + 1 <= prec')%Z -> forall x : R, FTZ_format beta emin prec x -> round_round_eq (FTZ_exp emin prec) (FTZ_exp emin' prec') choice1 choice2 (sqrt x)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(4 <= beta)%Z -> forall choice1 choice2 : Z -> bool, (2 * (emin' + prec') <= emin + prec <= 1)%Z -> (2 * prec + 1 <= prec')%Z -> forall x : R, FTZ_format beta emin prec x -> round_round_eq (FTZ_exp emin prec) (FTZ_exp emin' prec') choice1 choice2 (sqrt x)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(4 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(2 * (emin' + prec') <= emin + prec <= 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x:R
Fx:FTZ_format beta emin prec x

round_round_eq (FTZ_exp emin prec) (FTZ_exp emin' prec') choice1 choice2 (sqrt x)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(4 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(2 * (emin' + prec') <= emin + prec <= 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x:R
Fx:FTZ_format beta emin prec x

(4 <= beta)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(4 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(2 * (emin' + prec') <= emin + prec <= 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x:R
Fx:FTZ_format beta emin prec x
Valid_exp (FTZ_exp emin prec)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(4 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(2 * (emin' + prec') <= emin + prec <= 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x:R
Fx:FTZ_format beta emin prec x
Valid_exp (FTZ_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(4 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(2 * (emin' + prec') <= emin + prec <= 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x:R
Fx:FTZ_format beta emin prec x
round_round_sqrt_radix_ge_4_hyp (FTZ_exp emin prec) (FTZ_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(4 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(2 * (emin' + prec') <= emin + prec <= 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x:R
Fx:FTZ_format beta emin prec x
generic_format beta (FTZ_exp emin prec) x
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(4 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(2 * (emin' + prec') <= emin + prec <= 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x:R
Fx:FTZ_format beta emin prec x

(4 <= beta)%Z
exact Hbeta.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(4 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(2 * (emin' + prec') <= emin + prec <= 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x:R
Fx:FTZ_format beta emin prec x

Valid_exp (FTZ_exp emin prec)
now apply FTZ_exp_valid.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(4 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(2 * (emin' + prec') <= emin + prec <= 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x:R
Fx:FTZ_format beta emin prec x

Valid_exp (FTZ_exp emin' prec')
now apply FTZ_exp_valid.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(4 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(2 * (emin' + prec') <= emin + prec <= 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x:R
Fx:FTZ_format beta emin prec x

round_round_sqrt_radix_ge_4_hyp (FTZ_exp emin prec) (FTZ_exp emin' prec')
now apply FTZ_round_round_sqrt_radix_ge_4_hyp.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hbeta:(4 <= beta)%Z
choice1, choice2:Z -> bool
Hemin:(2 * (emin' + prec') <= emin + prec <= 1)%Z
Hprec:(2 * prec + 1 <= prec')%Z
x:R
Fx:FTZ_format beta emin prec x

generic_format beta (FTZ_exp emin prec) x
now apply generic_format_FTZ. Qed. End Double_round_sqrt_radix_ge_4_FTZ. End Double_round_sqrt_radix_ge_4. End Double_round_sqrt. Section Double_round_div.
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall choice1 choice2 : Z -> bool, (exists n : Z, beta = (2 * n)%Z) -> forall x : R, 0 < x -> (fexp2 (mag x) <= fexp1 (mag x) - 1)%Z -> (fexp1 (mag x) <= mag x)%Z -> x = midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall choice1 choice2 : Z -> bool, (exists n : Z, beta = (2 * n)%Z) -> forall x : R, 0 < x -> (fexp2 (mag x) <= fexp1 (mag x) - 1)%Z -> (fexp1 (mag x) <= mag x)%Z -> x = midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z

x = midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z

x = midp fexp1 x -> round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z

x = round beta fexp1 Zfloor x + / 2 * ulp beta fexp1 x -> round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R

x = rd + / 2 * ulp beta fexp1 x -> round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R

x = rd + / 2 * u -> round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
H:- rd + x = - rd + (rd + / 2 * u)

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R

- rd + x = / 2 * u -> round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R

x - rd = / 2 * u -> round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x - rd = / 2 * u

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x - rd = / 2 * u

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x - rd = / 2 * u

(2 <= beta)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x - rd = / 2 * u
Hbeta:(2 <= beta)%Z
round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x - rd = / 2 * u

(2 <= beta)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n, beta_val:Z
beta_prop:(2 <=? beta_val)%Z = true
Ebeta:{| radix_val := beta_val; radix_prop := beta_prop |} = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (Raux.mag {| radix_val := beta_val; radix_prop := beta_prop |} x) <= fexp1 (Raux.mag {| radix_val := beta_val; radix_prop := beta_prop |} x) - 1)%Z
Hf1:(fexp1 (Raux.mag {| radix_val := beta_val; radix_prop := beta_prop |} x) <= Raux.mag {| radix_val := beta_val; radix_prop := beta_prop |} x)%Z
rd:=round {| radix_val := beta_val; radix_prop := beta_prop |} fexp1 Zfloor x:R
u:=ulp {| radix_val := beta_val; radix_prop := beta_prop |} fexp1 x:R
Xmid:x - rd = / 2 * u

(2 <= {| radix_val := beta_val; radix_prop := beta_prop |})%Z
now apply Zle_bool_imp_le.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x - rd = / 2 * u
Hbeta:(2 <= beta)%Z

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z

generic_format beta fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta

generic_format beta fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta

F2R f = x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta
Hf:F2R f = x
generic_format beta fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta

F2R f = x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta

IZR (Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x))) * bpow (cexp beta fexp2 x) = x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta

(IZR (Zfloor (scaled_mantissa beta fexp2 rd)) + IZR (n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)))) * bpow (cexp beta fexp2 x) = x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta

IZR (Zfloor (scaled_mantissa beta fexp2 rd)) * bpow (cexp beta fexp2 x) + IZR (n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x))) * bpow (cexp beta fexp2 x) = x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta

IZR (Zfloor (scaled_mantissa beta fexp2 rd)) * bpow (cexp beta fexp2 x) + IZR n * IZR (beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x))) * bpow (cexp beta fexp2 x) = x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta

IZR (Zfloor (scaled_mantissa beta fexp2 rd)) * bpow (cexp beta fexp2 x) + IZR n * bpow (fexp1 (mag x) - 1 - fexp2 (mag x)) * bpow (cexp beta fexp2 x) = x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta

IZR (Zfloor (scaled_mantissa beta fexp2 rd)) * bpow (cexp beta fexp2 x) + IZR n * bpow (fexp1 (mag x) - 1) = x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta

IZR (Zfloor (scaled_mantissa beta fexp2 rd)) * bpow (cexp beta fexp2 x) + IZR n * (bpow (fexp1 (mag x)) * bpow (- (1))) = x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta

IZR (Zfloor (scaled_mantissa beta fexp2 rd)) * bpow (cexp beta fexp2 x) + IZR n * (bpow (-1) * bpow (fexp1 (mag x))) = x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta

IZR (Zfloor (scaled_mantissa beta fexp2 rd)) * bpow (cexp beta fexp2 x) + IZR n * bpow (-1) * bpow (fexp1 (mag x)) = x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta

IZR (Zfloor (scaled_mantissa beta fexp2 rd)) * bpow (cexp beta fexp2 x) + IZR n * / IZR (beta * 1) * bpow (fexp1 (mag x)) = x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta

IZR (Zfloor (scaled_mantissa beta fexp2 rd)) * bpow (cexp beta fexp2 x) + IZR n * / IZR beta * bpow (fexp1 (mag x)) = x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta

IZR (Zfloor (scaled_mantissa beta fexp2 rd)) * bpow (cexp beta fexp2 x) + IZR n * / IZR (2 * n) * bpow (fexp1 (mag x)) = x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta

IZR (Zfloor (scaled_mantissa beta fexp2 rd)) * bpow (cexp beta fexp2 x) + IZR n * / (2 * IZR n) * bpow (fexp1 (mag x)) = x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta

IZR (Zfloor (scaled_mantissa beta fexp2 rd)) * bpow (cexp beta fexp2 x) + IZR n * (/ 2 * / IZR n) * bpow (fexp1 (mag x)) = x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta

IZR (Zfloor (scaled_mantissa beta fexp2 rd)) * bpow (cexp beta fexp2 x) + / 2 * (IZR n * / IZR n) * bpow (fexp1 (mag x)) = x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta

IZR (Zfloor (scaled_mantissa beta fexp2 rd)) * bpow (cexp beta fexp2 x) + / 2 * bpow (fexp1 (mag x)) = x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta

IZR (Zfloor (scaled_mantissa beta fexp2 rd)) * bpow (cexp beta fexp2 x) + / 2 * bpow (cexp beta fexp1 x) = x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta

IZR (Zfloor (scaled_mantissa beta fexp2 rd)) * ulp beta fexp2 x + / 2 * ulp beta fexp1 x = x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta

IZR (Zfloor (scaled_mantissa beta fexp2 rd)) * ulp beta fexp2 x + / 2 * u = rd + / 2 * u
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta

IZR (Zfloor (scaled_mantissa beta fexp2 rd)) * ulp beta fexp2 x = rd
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta

IZR (Zfloor (scaled_mantissa beta fexp2 rd)) * bpow (cexp beta fexp2 x) = rd
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta
Zrd:rd = 0

IZR (Zfloor (scaled_mantissa beta fexp2 rd)) * bpow (cexp beta fexp2 x) = rd
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta
Nzrd:rd <> 0
IZR (Zfloor (scaled_mantissa beta fexp2 rd)) * bpow (cexp beta fexp2 x) = rd
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta
Zrd:rd = 0

IZR (Zfloor (scaled_mantissa beta fexp2 rd)) * bpow (cexp beta fexp2 x) = rd
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta
Zrd:rd = 0

IZR (Zfloor (scaled_mantissa beta fexp2 0)) * bpow (cexp beta fexp2 x) = 0
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta
Zrd:rd = 0

IZR (Zfloor 0) * bpow (cexp beta fexp2 x) = 0
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta
Zrd:rd = 0

0 * bpow (cexp beta fexp2 x) = 0
now rewrite Rmult_0_l.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta
Nzrd:rd <> 0

IZR (Zfloor (scaled_mantissa beta fexp2 rd)) * bpow (cexp beta fexp2 x) = rd
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta
Nzrd:rd <> 0

0 <= rd
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta
Nzrd:rd <> 0
Nnrd:0 <= rd
IZR (Zfloor (scaled_mantissa beta fexp2 rd)) * bpow (cexp beta fexp2 x) = rd
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta
Nzrd:rd <> 0

0 <= rd
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta
Nzrd:rd <> 0

Valid_exp fexp1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta
Nzrd:rd <> 0
generic_format beta fexp1 0
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta
Nzrd:rd <> 0
0 <= x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta
Nzrd:rd <> 0

Valid_exp fexp1
exact Vfexp1.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta
Nzrd:rd <> 0

generic_format beta fexp1 0
apply generic_format_0.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta
Nzrd:rd <> 0

0 <= x
now apply Rlt_le.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta
Nzrd:rd <> 0
Nnrd:0 <= rd

IZR (Zfloor (scaled_mantissa beta fexp2 rd)) * bpow (cexp beta fexp2 x) = rd
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta
Nzrd:rd <> 0
Nnrd:0 <= rd
Prd:0 < rd

IZR (Zfloor (scaled_mantissa beta fexp2 rd)) * bpow (cexp beta fexp2 x) = rd
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta
Nzrd:rd <> 0
Nnrd:0 <= rd
Prd:0 < rd

mag rd = mag x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta
Nzrd:rd <> 0
Nnrd:0 <= rd
Prd:0 < rd
Lrd:mag rd = mag x
IZR (Zfloor (scaled_mantissa beta fexp2 rd)) * bpow (cexp beta fexp2 x) = rd
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta
Nzrd:rd <> 0
Nnrd:0 <= rd
Prd:0 < rd

mag rd = mag x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta
Nzrd:rd <> 0
Nnrd:0 <= rd
Prd:0 < rd

(mag rd <= mag x)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta
Nzrd:rd <> 0
Nnrd:0 <= rd
Prd:0 < rd
(mag x <= mag rd)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta
Nzrd:rd <> 0
Nnrd:0 <= rd
Prd:0 < rd

(mag rd <= mag x)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta
Nzrd:rd <> 0
Nnrd:0 <= rd
Prd:0 < rd

rd <= x
now apply round_DN_pt.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta
Nzrd:rd <> 0
Nnrd:0 <= rd
Prd:0 < rd

(mag x <= mag rd)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta
Nzrd:rd <> 0
Nnrd:0 <= rd
Prd:0 < rd

Valid_exp fexp1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta
Nzrd:rd <> 0
Nnrd:0 <= rd
Prd:0 < rd
Valid_rnd Zfloor
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta
Nzrd:rd <> 0
Nnrd:0 <= rd
Prd:0 < rd
round beta fexp1 Zfloor x <> 0
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta
Nzrd:rd <> 0
Nnrd:0 <= rd
Prd:0 < rd

Valid_exp fexp1
exact Vfexp1.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta
Nzrd:rd <> 0
Nnrd:0 <= rd
Prd:0 < rd

Valid_rnd Zfloor
now apply valid_rnd_DN.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta
Nzrd:rd <> 0
Nnrd:0 <= rd
Prd:0 < rd

round beta fexp1 Zfloor x <> 0
exact Nzrd.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta
Nzrd:rd <> 0
Nnrd:0 <= rd
Prd:0 < rd
Lrd:mag rd = mag x

IZR (Zfloor (scaled_mantissa beta fexp2 rd)) * bpow (cexp beta fexp2 x) = rd
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta
Nzrd:rd <> 0
Nnrd:0 <= rd
Prd:0 < rd
Lrd:mag rd = mag x

IZR (Zfloor (rd * bpow (- cexp beta fexp2 rd))) * bpow (cexp beta fexp2 x) = rd
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta
Nzrd:rd <> 0
Nnrd:0 <= rd
Prd:0 < rd
Lrd:mag rd = mag x

IZR (Zfloor (round beta fexp1 Zfloor x * bpow (- cexp beta fexp2 rd))) * bpow (cexp beta fexp2 x) = rd
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta
Nzrd:rd <> 0
Nnrd:0 <= rd
Prd:0 < rd
Lrd:mag rd = mag x

IZR (Zfloor (IZR (Zfloor (x * bpow (- fexp1 (mag x)))) * bpow (fexp1 (mag x)) * bpow (- fexp2 (mag rd)))) * bpow (fexp2 (mag x)) = rd
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta
Nzrd:rd <> 0
Nnrd:0 <= rd
Prd:0 < rd
Lrd:mag rd = mag x

IZR (Zfloor (IZR (Zfloor (x * bpow (- fexp1 (mag x)))) * bpow (fexp1 (mag x) - fexp2 (mag rd)))) * bpow (fexp2 (mag x)) = rd
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta
Nzrd:rd <> 0
Nnrd:0 <= rd
Prd:0 < rd
Lrd:mag rd = mag x

IZR (Zfloor (IZR (Zfloor (x * bpow (- fexp1 (mag x)))) * bpow (fexp1 (mag x) - fexp2 (mag x)))) * bpow (fexp2 (mag x)) = rd
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta
Nzrd:rd <> 0
Nnrd:0 <= rd
Prd:0 < rd
Lrd:mag rd = mag x

IZR (Zfloor (IZR (Zfloor (x * bpow (- fexp1 (mag x)))) * IZR (beta ^ (fexp1 (mag x) - fexp2 (mag x))))) * bpow (fexp2 (mag x)) = rd
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta
Nzrd:rd <> 0
Nnrd:0 <= rd
Prd:0 < rd
Lrd:mag rd = mag x

IZR (Zfloor (IZR (Zfloor (x * bpow (- fexp1 (mag x))) * beta ^ (fexp1 (mag x) - fexp2 (mag x))))) * bpow (fexp2 (mag x)) = rd
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta
Nzrd:rd <> 0
Nnrd:0 <= rd
Prd:0 < rd
Lrd:mag rd = mag x

IZR (Zfloor (x * bpow (- fexp1 (mag x))) * beta ^ (fexp1 (mag x) - fexp2 (mag x))) * bpow (fexp2 (mag x)) = rd
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta
Nzrd:rd <> 0
Nnrd:0 <= rd
Prd:0 < rd
Lrd:mag rd = mag x
IZR (Zfloor (x * bpow (- fexp1 (mag x))) * beta ^ (fexp1 (mag x) - fexp2 (mag x))) <= IZR (Zfloor (x * bpow (- fexp1 (mag x))) * beta ^ (fexp1 (mag x) - fexp2 (mag x))) < IZR (Zfloor (x * bpow (- fexp1 (mag x))) * beta ^ (fexp1 (mag x) - fexp2 (mag x)) + 1)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta
Nzrd:rd <> 0
Nnrd:0 <= rd
Prd:0 < rd
Lrd:mag rd = mag x

IZR (Zfloor (x * bpow (- fexp1 (mag x))) * beta ^ (fexp1 (mag x) - fexp2 (mag x))) * bpow (fexp2 (mag x)) = rd
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta
Nzrd:rd <> 0
Nnrd:0 <= rd
Prd:0 < rd
Lrd:mag rd = mag x

IZR (Zfloor (x * bpow (- fexp1 (mag x)))) * IZR (beta ^ (fexp1 (mag x) - fexp2 (mag x))) * bpow (fexp2 (mag x)) = rd
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta
Nzrd:rd <> 0
Nnrd:0 <= rd
Prd:0 < rd
Lrd:mag rd = mag x

IZR (Zfloor (x * bpow (- fexp1 (mag x)))) * bpow (fexp1 (mag x) - fexp2 (mag x)) * bpow (fexp2 (mag x)) = rd
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta
Nzrd:rd <> 0
Nnrd:0 <= rd
Prd:0 < rd
Lrd:mag rd = mag x

IZR (Zfloor (x * bpow (- fexp1 (mag x)))) * bpow (fexp1 (mag x)) = rd
now unfold rd.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta
Nzrd:rd <> 0
Nnrd:0 <= rd
Prd:0 < rd
Lrd:mag rd = mag x

IZR (Zfloor (x * bpow (- fexp1 (mag x))) * beta ^ (fexp1 (mag x) - fexp2 (mag x))) <= IZR (Zfloor (x * bpow (- fexp1 (mag x))) * beta ^ (fexp1 (mag x) - fexp2 (mag x))) < IZR (Zfloor (x * bpow (- fexp1 (mag x))) * beta ^ (fexp1 (mag x) - fexp2 (mag x)) + 1)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta
Nzrd:rd <> 0
Nnrd:0 <= rd
Prd:0 < rd
Lrd:mag rd = mag x

IZR (Zfloor (x * bpow (- fexp1 (mag x))) * beta ^ (fexp1 (mag x) - fexp2 (mag x))) < IZR (Zfloor (x * bpow (- fexp1 (mag x))) * beta ^ (fexp1 (mag x) - fexp2 (mag x)) + 1)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta
Nzrd:rd <> 0
Nnrd:0 <= rd
Prd:0 < rd
Lrd:mag rd = mag x

IZR (Zfloor (x * bpow (- fexp1 (mag x))) * beta ^ (fexp1 (mag x) - fexp2 (mag x))) < IZR (Zfloor (x * bpow (- fexp1 (mag x))) * beta ^ (fexp1 (mag x) - fexp2 (mag x))) + 1
simpl; lra.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta
Hf:F2R f = x

generic_format beta fexp2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta
Hf:F2R f = x

x <> 0 -> (cexp beta fexp2 x <= Fexp f)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
n:Z
Ebeta:beta = (2 * n)%Z
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
Hf1:(fexp1 (mag x) <= mag x)%Z
rd:=round beta fexp1 Zfloor x:R
u:=ulp beta fexp1 x:R
Xmid:x = rd + / 2 * u
Hbeta:(2 <= beta)%Z
f:={| Fnum := Zfloor (scaled_mantissa beta fexp2 rd) + n * beta ^ (fexp1 (mag x) - 1 - fexp2 (mag x)); Fexp := cexp beta fexp2 x |}:float beta
Hf:F2R f = x

(cexp beta fexp2 x <= Fexp f)%Z
apply Z.le_refl. Qed.
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall (choice1 choice2 : Z -> bool) (x : R), 0 < x -> (mag x <= fexp1 (mag x) - 2)%Z -> round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall (choice1 choice2 : Z -> bool) (x : R), 0 < x -> (mag x <= fexp1 (mag x) - 2)%Z -> round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z

round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z

bpow (mag x - 1) <= x < bpow (mag x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z

bpow (mag x - 1) <= x < bpow (mag x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
ex:Z
Hex:x <> 0 -> bpow (ex - 1) <= Rabs x < bpow ex
Hf1:(Build_mag_prop beta x ex Hex <= fexp1 (Build_mag_prop beta x ex Hex) - 2)%Z

bpow (ex - 1) <= x < bpow ex
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
ex:Z
Hex:x <> 0 -> bpow (ex - 1) <= Rabs x < bpow ex
Hf1:(Build_mag_prop beta x ex Hex <= fexp1 (Build_mag_prop beta x ex Hex) - 2)%Z

bpow (ex - 1) <= Rabs x < bpow ex
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
ex:Z
Hex:x <> 0 -> bpow (ex - 1) <= Rabs x < bpow ex
Hf1:(Build_mag_prop beta x ex Hex <= fexp1 (Build_mag_prop beta x ex Hex) - 2)%Z

x <> 0
now apply Rgt_not_eq.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z
Hlx:bpow (mag x - 1) <= x < bpow (mag x)

round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z
Hlx:bpow (mag x - 1) <= x < bpow (mag x)

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z
Hlx:bpow (mag x - 1) <= x < bpow (mag x)

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = 0
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
x'':=round beta fexp2 (Znearest choice2) x:R

round beta fexp1 (Znearest choice1) x'' = 0
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
x'':=round beta fexp2 (Znearest choice2) x:R
Nzx'':x'' <> 0

round beta fexp1 (Znearest choice1) x'' = 0
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
x'':=round beta fexp2 (Znearest choice2) x:R
Nzx'':x'' <> 0
H:(fexp2 (mag x) <= mag x)%Z

round beta fexp1 (Znearest choice1) x'' = 0
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
x'':=round beta fexp2 (Znearest choice2) x:R
Nzx'':x'' <> 0
H:(mag x < fexp2 (mag x))%Z
round beta fexp1 (Znearest choice1) x'' = 0
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
x'':=round beta fexp2 (Znearest choice2) x:R
Nzx'':x'' <> 0
H:(fexp2 (mag x) <= mag x)%Z

round beta fexp1 (Znearest choice1) x'' = 0
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
x'':=round beta fexp2 (Znearest choice2) x:R
Nzx'':x'' <> 0
H:(fexp2 (mag x) <= mag x)%Z
H0:x'' < bpow (mag x)

round beta fexp1 (Znearest choice1) x'' = 0
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
x'':=round beta fexp2 (Znearest choice2) x:R
Nzx'':x'' <> 0
H:(fexp2 (mag x) <= mag x)%Z
H0:bpow (mag x) <= x''
round beta fexp1 (Znearest choice1) x'' = 0
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
x'':=round beta fexp2 (Znearest choice2) x:R
Nzx'':x'' <> 0
H:(fexp2 (mag x) <= mag x)%Z
H0:x'' < bpow (mag x)

round beta fexp1 (Znearest choice1) x'' = 0
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
x'':=round beta fexp2 (Znearest choice2) x:R
Nzx'':x'' <> 0
H:(fexp2 (mag x) <= mag x)%Z
H0:x'' < bpow (mag x)

bpow (mag x - 1) <= x''
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
x'':=round beta fexp2 (Znearest choice2) x:R
Nzx'':x'' <> 0
H:(fexp2 (mag x) <= mag x)%Z
H0:x'' < bpow (mag x)

0 < round beta fexp2 (Znearest choice2) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
x'':=round beta fexp2 (Znearest choice2) x:R
Nzx'':x'' <> 0
H:(fexp2 (mag x) <= mag x)%Z
H0:x'' < bpow (mag x)

0 <= round beta fexp2 (Znearest choice2) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
x'':=round beta fexp2 (Znearest choice2) x:R
Nzx'':x'' <> 0
H:(fexp2 (mag x) <= mag x)%Z
H0:x'' < bpow (mag x)

round beta fexp2 (Znearest choice2) 0 <= round beta fexp2 (Znearest choice2) x
now apply round_le; [|apply valid_rnd_N|apply Rlt_le].
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
x'':=round beta fexp2 (Znearest choice2) x:R
Nzx'':x'' <> 0
H:(fexp2 (mag x) <= mag x)%Z
H0:bpow (mag x) <= x''

round beta fexp1 (Znearest choice1) x'' = 0
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
x'':=round beta fexp2 (Znearest choice2) x:R
Nzx'':x'' <> 0
H:(fexp2 (mag x) <= mag x)%Z
H0:bpow (mag x) <= x''

x'' = bpow (mag x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
x'':=round beta fexp2 (Znearest choice2) x:R
Nzx'':x'' <> 0
H:(fexp2 (mag x) <= mag x)%Z
H0:bpow (mag x) <= x''
Hx'':x'' = bpow (mag x)
round beta fexp1 (Znearest choice1) x'' = 0
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
x'':=round beta fexp2 (Znearest choice2) x:R
Nzx'':x'' <> 0
H:(fexp2 (mag x) <= mag x)%Z
H0:bpow (mag x) <= x''

x'' = bpow (mag x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
x'':=round beta fexp2 (Znearest choice2) x:R
Nzx'':x'' <> 0
H:(fexp2 (mag x) <= mag x)%Z
H0:bpow (mag x) <= x''

x'' <= bpow (mag x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
x'':=round beta fexp2 (Znearest choice2) x:R
Nzx'':x'' <> 0
H:(fexp2 (mag x) <= mag x)%Z
H0:bpow (mag x) <= x''

x'' <= round beta fexp2 (Znearest choice2) (bpow (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
x'':=round beta fexp2 (Znearest choice2) x:R
Nzx'':x'' <> 0
H:(fexp2 (mag x) <= mag x)%Z
H0:bpow (mag x) <= x''
generic_format beta fexp2 (bpow (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
x'':=round beta fexp2 (Znearest choice2) x:R
Nzx'':x'' <> 0
H:(fexp2 (mag x) <= mag x)%Z
H0:bpow (mag x) <= x''

x'' <= round beta fexp2 (Znearest choice2) (bpow (mag x))
now apply round_le; [|apply valid_rnd_N|apply Rlt_le].
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
x'':=round beta fexp2 (Znearest choice2) x:R
Nzx'':x'' <> 0
H:(fexp2 (mag x) <= mag x)%Z
H0:bpow (mag x) <= x''

generic_format beta fexp2 (bpow (mag x))
now apply generic_format_bpow'.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
x'':=round beta fexp2 (Znearest choice2) x:R
Nzx'':x'' <> 0
H:(fexp2 (mag x) <= mag x)%Z
H0:bpow (mag x) <= x''
Hx'':x'' = bpow (mag x)

round beta fexp1 (Znearest choice1) x'' = 0
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
x'':=round beta fexp2 (Znearest choice2) x:R
Nzx'':x'' <> 0
H:(fexp2 (mag x) <= mag x)%Z
H0:bpow (mag x) <= x''
Hx'':x'' = bpow (mag x)

round beta fexp1 (Znearest choice1) (bpow (mag x)) = 0
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
x'':=round beta fexp2 (Znearest choice2) x:R
Nzx'':x'' <> 0
H:(fexp2 (mag x) <= mag x)%Z
H0:bpow (mag x) <= x''
Hx'':x'' = bpow (mag x)

IZR (Znearest choice1 (bpow (mag x) * bpow (- fexp1 (mag (bpow (mag x)))))) * bpow (fexp1 (mag (bpow (mag x)))) = 0
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
x'':=round beta fexp2 (Znearest choice2) x:R
Nzx'':x'' <> 0
H:(fexp2 (mag x) <= mag x)%Z
H0:bpow (mag x) <= x''
Hx'':x'' = bpow (mag x)

IZR (Znearest choice1 (bpow (mag x) * bpow (- fexp1 (mag x + 1)%Z))) * bpow (fexp1 (mag x + 1)%Z) = 0
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
x'':=round beta fexp2 (Znearest choice2) x:R
Nzx'':x'' <> 0
H:(fexp2 (mag x) <= mag x)%Z
H0:bpow (mag x) <= x''
Hx'':x'' = bpow (mag x)
Hf11:fexp1 (mag x + 1)%Z = fexp1 (mag x)

IZR (Znearest choice1 (bpow (mag x) * bpow (- fexp1 (mag x + 1)%Z))) * bpow (fexp1 (mag x + 1)%Z) = 0
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
x'':=round beta fexp2 (Znearest choice2) x:R
Nzx'':x'' <> 0
H:(fexp2 (mag x) <= mag x)%Z
H0:bpow (mag x) <= x''
Hx'':x'' = bpow (mag x)
Hf11:fexp1 (mag x + 1)%Z = fexp1 (mag x)

IZR (Znearest choice1 (bpow (mag x) * bpow (- fexp1 (mag x)))) * bpow (fexp1 (mag x)) = 0
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
x'':=round beta fexp2 (Znearest choice2) x:R
Nzx'':x'' <> 0
H:(fexp2 (mag x) <= mag x)%Z
H0:bpow (mag x) <= x''
Hx'':x'' = bpow (mag x)
Hf11:fexp1 (mag x + 1)%Z = fexp1 (mag x)

IZR (Znearest choice1 (bpow (mag x) * bpow (- fexp1 (mag x)))) * bpow (fexp1 (mag x)) * bpow (- fexp1 (mag x)) = 0 * bpow (- fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
x'':=round beta fexp2 (Znearest choice2) x:R
Nzx'':x'' <> 0
H:(fexp2 (mag x) <= mag x)%Z
H0:bpow (mag x) <= x''
Hx'':x'' = bpow (mag x)
Hf11:fexp1 (mag x + 1)%Z = fexp1 (mag x)

IZR (Znearest choice1 (bpow (mag x - fexp1 (mag x)))) = 0
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
x'':=round beta fexp2 (Znearest choice2) x:R
Nzx'':x'' <> 0
H:(fexp2 (mag x) <= mag x)%Z
H0:bpow (mag x) <= x''
Hx'':x'' = bpow (mag x)
Hf11:fexp1 (mag x + 1)%Z = fexp1 (mag x)

Znearest choice1 (bpow (mag x - fexp1 (mag x))) = 0%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
x'':=round beta fexp2 (Znearest choice2) x:R
Nzx'':x'' <> 0
H:(fexp2 (mag x) <= mag x)%Z
H0:bpow (mag x) <= x''
Hx'':x'' = bpow (mag x)
Hf11:fexp1 (mag x + 1)%Z = fexp1 (mag x)

Rabs (bpow (mag x - fexp1 (mag x)) - 0) < / 2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
x'':=round beta fexp2 (Znearest choice2) x:R
Nzx'':x'' <> 0
H:(fexp2 (mag x) <= mag x)%Z
H0:bpow (mag x) <= x''
Hx'':x'' = bpow (mag x)
Hf11:fexp1 (mag x + 1)%Z = fexp1 (mag x)

Rabs (bpow (mag x - fexp1 (mag x))) < / 2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
x'':=round beta fexp2 (Znearest choice2) x:R
Nzx'':x'' <> 0
H:(fexp2 (mag x) <= mag x)%Z
H0:bpow (mag x) <= x''
Hx'':x'' = bpow (mag x)
Hf11:fexp1 (mag x + 1)%Z = fexp1 (mag x)

bpow (mag x - fexp1 (mag x)) < / 2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
x'':=round beta fexp2 (Znearest choice2) x:R
Nzx'':x'' <> 0
H:(fexp2 (mag x) <= mag x)%Z
H0:bpow (mag x) <= x''
Hx'':x'' = bpow (mag x)
Hf11:fexp1 (mag x + 1)%Z = fexp1 (mag x)

bpow (-2) < / 2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
x'':=round beta fexp2 (Znearest choice2) x:R
Nzx'':x'' <> 0
H:(fexp2 (mag x) <= mag x)%Z
H0:bpow (mag x) <= x''
Hx'':x'' = bpow (mag x)
Hf11:fexp1 (mag x + 1)%Z = fexp1 (mag x)

/ IZR (beta * beta) < / 2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
x'':=round beta fexp2 (Znearest choice2) x:R
Nzx'':x'' <> 0
H:(fexp2 (mag x) <= mag x)%Z
H0:bpow (mag x) <= x''
Hx'':x'' = bpow (mag x)
Hf11:fexp1 (mag x + 1)%Z = fexp1 (mag x)

(2 <= beta)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
x'':=round beta fexp2 (Znearest choice2) x:R
Nzx'':x'' <> 0
H:(fexp2 (mag x) <= mag x)%Z
H0:bpow (mag x) <= x''
Hx'':x'' = bpow (mag x)
Hf11:fexp1 (mag x + 1)%Z = fexp1 (mag x)
Hbeta:(2 <= beta)%Z
/ IZR (beta * beta) < / 2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
x'':=round beta fexp2 (Znearest choice2) x:R
Nzx'':x'' <> 0
H:(fexp2 (mag x) <= mag x)%Z
H0:bpow (mag x) <= x''
Hx'':x'' = bpow (mag x)
Hf11:fexp1 (mag x + 1)%Z = fexp1 (mag x)

(2 <= beta)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
beta_val:Z
beta_prop:(2 <=? beta_val)%Z = true
Hf1:(Raux.mag {| radix_val := beta_val; radix_prop := beta_prop |} x <= fexp1 (Raux.mag {| radix_val := beta_val; radix_prop := beta_prop |} x) - 2)%Z
Hlx:Raux.bpow {| radix_val := beta_val; radix_prop := beta_prop |} (Raux.mag {| radix_val := beta_val; radix_prop := beta_prop |} x - 1) <= x < Raux.bpow {| radix_val := beta_val; radix_prop := beta_prop |} (Raux.mag {| radix_val := beta_val; radix_prop := beta_prop |} x)
x'':=round {| radix_val := beta_val; radix_prop := beta_prop |} fexp2 (Znearest choice2) x:R
Nzx'':x'' <> 0
H:(fexp2 (Raux.mag {| radix_val := beta_val; radix_prop := beta_prop |} x) <= Raux.mag {| radix_val := beta_val; radix_prop := beta_prop |} x)%Z
H0:Raux.bpow {| radix_val := beta_val; radix_prop := beta_prop |} (Raux.mag {| radix_val := beta_val; radix_prop := beta_prop |} x) <= x''
Hx'':x'' = Raux.bpow {| radix_val := beta_val; radix_prop := beta_prop |} (Raux.mag {| radix_val := beta_val; radix_prop := beta_prop |} x)
Hf11:fexp1 (Raux.mag {| radix_val := beta_val; radix_prop := beta_prop |} x + 1)%Z = fexp1 (Raux.mag {| radix_val := beta_val; radix_prop := beta_prop |} x)

(2 <= beta_val)%Z
now apply Zle_bool_imp_le.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
x'':=round beta fexp2 (Znearest choice2) x:R
Nzx'':x'' <> 0
H:(fexp2 (mag x) <= mag x)%Z
H0:bpow (mag x) <= x''
Hx'':x'' = bpow (mag x)
Hf11:fexp1 (mag x + 1)%Z = fexp1 (mag x)
Hbeta:(2 <= beta)%Z

/ IZR (beta * beta) < / 2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
x'':=round beta fexp2 (Znearest choice2) x:R
Nzx'':x'' <> 0
H:(fexp2 (mag x) <= mag x)%Z
H0:bpow (mag x) <= x''
Hx'':x'' = bpow (mag x)
Hf11:fexp1 (mag x + 1)%Z = fexp1 (mag x)
Hbeta:(2 <= beta)%Z

0 < 2 * IZR (beta * beta)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
x'':=round beta fexp2 (Znearest choice2) x:R
Nzx'':x'' <> 0
H:(fexp2 (mag x) <= mag x)%Z
H0:bpow (mag x) <= x''
Hx'':x'' = bpow (mag x)
Hf11:fexp1 (mag x + 1)%Z = fexp1 (mag x)
Hbeta:(2 <= beta)%Z
2 < IZR (beta * beta)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
x'':=round beta fexp2 (Znearest choice2) x:R
Nzx'':x'' <> 0
H:(fexp2 (mag x) <= mag x)%Z
H0:bpow (mag x) <= x''
Hx'':x'' = bpow (mag x)
Hf11:fexp1 (mag x + 1)%Z = fexp1 (mag x)
Hbeta:(2 <= beta)%Z

0 < 2 * IZR (beta * beta)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
x'':=round beta fexp2 (Znearest choice2) x:R
Nzx'':x'' <> 0
H:(fexp2 (mag x) <= mag x)%Z
H0:bpow (mag x) <= x''
Hx'':x'' = bpow (mag x)
Hf11:fexp1 (mag x + 1)%Z = fexp1 (mag x)
Hbeta:(2 <= beta)%Z

0 < IZR (beta * beta)
rewrite mult_IZR; apply Rmult_lt_0_compat; apply IZR_lt; omega.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
x'':=round beta fexp2 (Znearest choice2) x:R
Nzx'':x'' <> 0
H:(fexp2 (mag x) <= mag x)%Z
H0:bpow (mag x) <= x''
Hx'':x'' = bpow (mag x)
Hf11:fexp1 (mag x + 1)%Z = fexp1 (mag x)
Hbeta:(2 <= beta)%Z

2 < IZR (beta * beta)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
x'':=round beta fexp2 (Znearest choice2) x:R
Nzx'':x'' <> 0
H:(fexp2 (mag x) <= mag x)%Z
H0:bpow (mag x) <= x''
Hx'':x'' = bpow (mag x)
Hf11:fexp1 (mag x + 1)%Z = fexp1 (mag x)
Hbeta:(2 <= beta)%Z

(2 < beta * beta)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
x'':=round beta fexp2 (Znearest choice2) x:R
Nzx'':x'' <> 0
H:(fexp2 (mag x) <= mag x)%Z
H0:bpow (mag x) <= x''
Hx'':x'' = bpow (mag x)
Hf11:fexp1 (mag x + 1)%Z = fexp1 (mag x)
Hbeta:(2 <= beta)%Z

(beta < beta * beta)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
x'':=round beta fexp2 (Znearest choice2) x:R
Nzx'':x'' <> 0
H:(fexp2 (mag x) <= mag x)%Z
H0:bpow (mag x) <= x''
Hx'':x'' = bpow (mag x)
Hf11:fexp1 (mag x + 1)%Z = fexp1 (mag x)
Hbeta:(2 <= beta)%Z

(beta * 1 < beta * beta)%Z
apply Zmult_lt_compat_l; omega.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
x'':=round beta fexp2 (Znearest choice2) x:R
Nzx'':x'' <> 0
H:(mag x < fexp2 (mag x))%Z

round beta fexp1 (Znearest choice1) x'' = 0
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:(mag x <= fexp1 (mag x) - 2)%Z
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
x'':=round beta fexp2 (Znearest choice2) x:R
Nzx'':x'' <> 0
H:(mag x < fexp2 (mag x))%Z

x'' = 0
now apply (round_N_small_pos beta _ _ _ (mag x)). Qed.
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall (choice1 choice2 : Z -> bool) (x : R), 0 < x -> fexp1 (mag x) = (mag x + 1)%Z -> x < bpow (mag x) - / 2 * ulp beta fexp2 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall (choice1 choice2 : Z -> bool) (x : R), 0 < x -> fexp1 (mag x) = (mag x + 1)%Z -> x < bpow (mag x) - / 2 * ulp beta fexp2 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:fexp1 (mag x) = (mag x + 1)%Z

x < bpow (mag x) - / 2 * ulp beta fexp2 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:fexp1 (mag x) = (mag x + 1)%Z

x < bpow (mag x) - / 2 * ulp beta fexp2 x -> round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) x) = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:fexp1 (mag x) = (mag x + 1)%Z
x'':=round beta fexp2 (Znearest choice2) x:R

x < bpow (mag x) - / 2 * ulp beta fexp2 x -> round beta fexp1 (Znearest choice1) x'' = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:fexp1 (mag x) = (mag x + 1)%Z
x'':=round beta fexp2 (Znearest choice2) x:R
u1:=ulp beta fexp1 x:R

x < bpow (mag x) - / 2 * ulp beta fexp2 x -> round beta fexp1 (Znearest choice1) x'' = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:fexp1 (mag x) = (mag x + 1)%Z
x'':=round beta fexp2 (Znearest choice2) x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R

x < bpow (mag x) - / 2 * u2 -> round beta fexp1 (Znearest choice1) x'' = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:fexp1 (mag x) = (mag x + 1)%Z
x'':=round beta fexp2 (Znearest choice2) x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Hx:x < bpow (mag x) - / 2 * u2

round beta fexp1 (Znearest choice1) x'' = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:fexp1 (mag x) = (mag x + 1)%Z
x'':=round beta fexp2 (Znearest choice2) x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Hx:x < bpow (mag x) - / 2 * u2

bpow (mag x - 1) <= x < bpow (mag x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:fexp1 (mag x) = (mag x + 1)%Z
x'':=round beta fexp2 (Znearest choice2) x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Hx:x < bpow (mag x) - / 2 * u2
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
round beta fexp1 (Znearest choice1) x'' = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:fexp1 (mag x) = (mag x + 1)%Z
x'':=round beta fexp2 (Znearest choice2) x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Hx:x < bpow (mag x) - / 2 * u2

bpow (mag x - 1) <= x < bpow (mag x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
ex:Z
Hex:x <> 0 -> bpow (ex - 1) <= Rabs x < bpow ex
Hf1:fexp1 (Build_mag_prop beta x ex Hex) = (Build_mag_prop beta x ex Hex + 1)%Z
x'':=round beta fexp2 (Znearest choice2) x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Hx:x < bpow (Build_mag_prop beta x ex Hex) - / 2 * u2

bpow (ex - 1) <= x < bpow ex
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
ex:Z
Hex:x <> 0 -> bpow (ex - 1) <= Rabs x < bpow ex
Hf1:fexp1 (Build_mag_prop beta x ex Hex) = (Build_mag_prop beta x ex Hex + 1)%Z
x'':=round beta fexp2 (Znearest choice2) x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Hx:x < bpow (Build_mag_prop beta x ex Hex) - / 2 * u2

bpow (ex - 1) <= Rabs x < bpow ex
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
ex:Z
Hex:x <> 0 -> bpow (ex - 1) <= Rabs x < bpow ex
Hf1:fexp1 (Build_mag_prop beta x ex Hex) = (Build_mag_prop beta x ex Hex + 1)%Z
x'':=round beta fexp2 (Znearest choice2) x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Hx:x < bpow (Build_mag_prop beta x ex Hex) - / 2 * u2

x <> 0
now apply Rgt_not_eq.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:fexp1 (mag x) = (mag x + 1)%Z
x'':=round beta fexp2 (Znearest choice2) x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Hx:x < bpow (mag x) - / 2 * u2
Hlx:bpow (mag x - 1) <= x < bpow (mag x)

round beta fexp1 (Znearest choice1) x'' = round beta fexp1 (Znearest choice1) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:fexp1 (mag x) = (mag x + 1)%Z
x'':=round beta fexp2 (Znearest choice2) x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Hx:x < bpow (mag x) - / 2 * u2
Hlx:bpow (mag x - 1) <= x < bpow (mag x)

round beta fexp1 (Znearest choice1) x'' = 0
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:fexp1 (mag x) = (mag x + 1)%Z
x'':=round beta fexp2 (Znearest choice2) x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Hx:x < bpow (mag x) - / 2 * u2
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
Nzx'':x'' <> 0

round beta fexp1 (Znearest choice1) x'' = 0
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:fexp1 (mag x) = (mag x + 1)%Z
x'':=round beta fexp2 (Znearest choice2) x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Hx:x < bpow (mag x) - / 2 * u2
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
Nzx'':x'' <> 0

bpow (mag x - 1) <= x'' < bpow (mag x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:fexp1 (mag x) = (mag x + 1)%Z
x'':=round beta fexp2 (Znearest choice2) x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Hx:x < bpow (mag x) - / 2 * u2
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
Nzx'':x'' <> 0

bpow (mag x - 1) <= x''
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:fexp1 (mag x) = (mag x + 1)%Z
x'':=round beta fexp2 (Znearest choice2) x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Hx:x < bpow (mag x) - / 2 * u2
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
Nzx'':x'' <> 0
x'' < bpow (mag x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:fexp1 (mag x) = (mag x + 1)%Z
x'':=round beta fexp2 (Znearest choice2) x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Hx:x < bpow (mag x) - / 2 * u2
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
Nzx'':x'' <> 0

bpow (mag x - 1) <= x''
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:fexp1 (mag x) = (mag x + 1)%Z
x'':=round beta fexp2 (Znearest choice2) x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Hx:x < bpow (mag x) - / 2 * u2
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
Nzx'':x'' <> 0

Valid_rnd (Znearest choice2)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:fexp1 (mag x) = (mag x + 1)%Z
x'':=round beta fexp2 (Znearest choice2) x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Hx:x < bpow (mag x) - / 2 * u2
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
Nzx'':x'' <> 0
0 < round beta fexp2 (Znearest choice2) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:fexp1 (mag x) = (mag x + 1)%Z
x'':=round beta fexp2 (Znearest choice2) x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Hx:x < bpow (mag x) - / 2 * u2
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
Nzx'':x'' <> 0
bpow (mag x - 1) <= x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:fexp1 (mag x) = (mag x + 1)%Z
x'':=round beta fexp2 (Znearest choice2) x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Hx:x < bpow (mag x) - / 2 * u2
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
Nzx'':x'' <> 0

Valid_rnd (Znearest choice2)
now apply valid_rnd_N.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:fexp1 (mag x) = (mag x + 1)%Z
x'':=round beta fexp2 (Znearest choice2) x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Hx:x < bpow (mag x) - / 2 * u2
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
Nzx'':x'' <> 0

0 < round beta fexp2 (Znearest choice2) x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:fexp1 (mag x) = (mag x + 1)%Z
x'':=round beta fexp2 (Znearest choice2) x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Hx:x < bpow (mag x) - / 2 * u2
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
Nzx'':x'' <> 0

0 <= x''
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:fexp1 (mag x) = (mag x + 1)%Z
x'':=round beta fexp2 (Znearest choice2) x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Hx:x < bpow (mag x) - / 2 * u2
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
Nzx'':x'' <> 0

round beta fexp2 (Znearest choice2) 0 <= x''
now apply round_le; [|apply valid_rnd_N|apply Rlt_le].
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:fexp1 (mag x) = (mag x + 1)%Z
x'':=round beta fexp2 (Znearest choice2) x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Hx:x < bpow (mag x) - / 2 * u2
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
Nzx'':x'' <> 0

bpow (mag x - 1) <= x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:fexp1 (mag x) = (mag x + 1)%Z
x'':=round beta fexp2 (Znearest choice2) x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Hx:x < bpow (mag x) - / 2 * u2
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
Nzx'':x'' <> 0

bpow (mag x - 1) <= Rabs x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
ex:Z
Hex:x <> 0 -> bpow (ex - 1) <= Rabs x < bpow ex
Hf1:fexp1 (Build_mag_prop beta x ex Hex) = (Build_mag_prop beta x ex Hex + 1)%Z
x'':=round beta fexp2 (Znearest choice2) x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Hx:x < bpow (Build_mag_prop beta x ex Hex) - / 2 * u2
Hlx:bpow (Build_mag_prop beta x ex Hex - 1) <= x < bpow (Build_mag_prop beta x ex Hex)
Nzx'':x'' <> 0

x <> 0
now apply Rgt_not_eq.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:fexp1 (mag x) = (mag x + 1)%Z
x'':=round beta fexp2 (Znearest choice2) x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Hx:x < bpow (mag x) - / 2 * u2
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
Nzx'':x'' <> 0

x'' < bpow (mag x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:fexp1 (mag x) = (mag x + 1)%Z
x'':=round beta fexp2 (Znearest choice2) x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Hx:x < bpow (mag x) - / 2 * u2
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
Nzx'':x'' <> 0

x + (x'' - x) < bpow (mag x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:fexp1 (mag x) = (mag x + 1)%Z
x'':=round beta fexp2 (Znearest choice2) x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Hx:x < bpow (mag x) - / 2 * u2
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
Nzx'':x'' <> 0

x + (x'' - x) < bpow (mag x) - / 2 * u2 + / 2 * u2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:fexp1 (mag x) = (mag x + 1)%Z
x'':=round beta fexp2 (Znearest choice2) x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Hx:x < bpow (mag x) - / 2 * u2
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
Nzx'':x'' <> 0

x'' - x <= / 2 * u2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf1:fexp1 (mag x) = (mag x + 1)%Z
x'':=round beta fexp2 (Znearest choice2) x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Hx:x < bpow (mag x) - / 2 * u2
Hlx:bpow (mag x - 1) <= x < bpow (mag x)
Nzx'':x'' <> 0

Rabs (x'' - x) <= / 2 * u2
now apply error_le_half_ulp. Qed.
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall (choice1 choice2 : Z -> bool) (x : R), 0 < x -> (fexp2 (mag x) <= fexp1 (mag x) - 1)%Z -> (fexp1 (mag x) = (mag x + 1)%Z -> bpow (mag x) - / 2 * ulp beta fexp2 x <= x -> round_round_eq fexp1 fexp2 choice1 choice2 x) -> ((fexp1 (mag x) <= mag x)%Z -> midp fexp1 x - / 2 * ulp beta fexp2 x <= x < midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x) -> ((fexp1 (mag x) <= mag x)%Z -> x = midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x) -> ((fexp1 (mag x) <= mag x)%Z -> midp fexp1 x < x <= midp fexp1 x + / 2 * ulp beta fexp2 x -> round_round_eq fexp1 fexp2 choice1 choice2 x) -> round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall (choice1 choice2 : Z -> bool) (x : R), 0 < x -> (fexp2 (mag x) <= fexp1 (mag x) - 1)%Z -> (fexp1 (mag x) = (mag x + 1)%Z -> bpow (mag x) - / 2 * ulp beta fexp2 x <= x -> round_round_eq fexp1 fexp2 choice1 choice2 x) -> ((fexp1 (mag x) <= mag x)%Z -> midp fexp1 x - / 2 * ulp beta fexp2 x <= x < midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x) -> ((fexp1 (mag x) <= mag x)%Z -> x = midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x) -> ((fexp1 (mag x) <= mag x)%Z -> midp fexp1 x < x <= midp fexp1 x + / 2 * ulp beta fexp2 x -> round_round_eq fexp1 fexp2 choice1 choice2 x) -> round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z

(fexp1 (mag x) = (mag x + 1)%Z -> bpow (mag x) - / 2 * ulp beta fexp2 x <= x -> round_round_eq fexp1 fexp2 choice1 choice2 x) -> ((fexp1 (mag x) <= mag x)%Z -> midp fexp1 x - / 2 * ulp beta fexp2 x <= x < midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x) -> ((fexp1 (mag x) <= mag x)%Z -> x = midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x) -> ((fexp1 (mag x) <= mag x)%Z -> midp fexp1 x < x <= midp fexp1 x + / 2 * ulp beta fexp2 x -> round_round_eq fexp1 fexp2 choice1 choice2 x) -> round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
x':=round beta fexp1 Zfloor x:R

(fexp1 (mag x) = (mag x + 1)%Z -> bpow (mag x) - / 2 * ulp beta fexp2 x <= x -> round_round_eq fexp1 fexp2 choice1 choice2 x) -> ((fexp1 (mag x) <= mag x)%Z -> midp fexp1 x - / 2 * ulp beta fexp2 x <= x < midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x) -> ((fexp1 (mag x) <= mag x)%Z -> x = midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x) -> ((fexp1 (mag x) <= mag x)%Z -> midp fexp1 x < x <= midp fexp1 x + / 2 * ulp beta fexp2 x -> round_round_eq fexp1 fexp2 choice1 choice2 x) -> round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
x':=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R

(fexp1 (mag x) = (mag x + 1)%Z -> bpow (mag x) - / 2 * ulp beta fexp2 x <= x -> round_round_eq fexp1 fexp2 choice1 choice2 x) -> ((fexp1 (mag x) <= mag x)%Z -> midp fexp1 x - / 2 * ulp beta fexp2 x <= x < midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x) -> ((fexp1 (mag x) <= mag x)%Z -> x = midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x) -> ((fexp1 (mag x) <= mag x)%Z -> midp fexp1 x < x <= midp fexp1 x + / 2 * ulp beta fexp2 x -> round_round_eq fexp1 fexp2 choice1 choice2 x) -> round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
x':=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R

(fexp1 (mag x) = (mag x + 1)%Z -> bpow (mag x) - / 2 * u2 <= x -> round_round_eq fexp1 fexp2 choice1 choice2 x) -> ((fexp1 (mag x) <= mag x)%Z -> midp fexp1 x - / 2 * u2 <= x < midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x) -> ((fexp1 (mag x) <= mag x)%Z -> x = midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x) -> ((fexp1 (mag x) <= mag x)%Z -> midp fexp1 x < x <= midp fexp1 x + / 2 * u2 -> round_round_eq fexp1 fexp2 choice1 choice2 x) -> round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
x':=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cz:fexp1 (mag x) = (mag x + 1)%Z -> bpow (mag x) - / 2 * u2 <= x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Clt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x - / 2 * u2 <= x < midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Ceq:(fexp1 (mag x) <= mag x)%Z -> x = midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Cgt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x < x <= midp fexp1 x + / 2 * u2 -> round_round_eq fexp1 fexp2 choice1 choice2 x

round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
x':=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cz:fexp1 (mag x) = (mag x + 1)%Z -> bpow (mag x) - / 2 * u2 <= x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Clt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x - / 2 * u2 <= x < midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Ceq:(fexp1 (mag x) <= mag x)%Z -> x = midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Cgt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x < x <= midp fexp1 x + / 2 * u2 -> round_round_eq fexp1 fexp2 choice1 choice2 x
Hlt:(mag x < fexp1 (mag x) - 1)%Z

round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
x':=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cz:fexp1 (mag x) = (mag x + 1)%Z -> bpow (mag x) - / 2 * u2 <= x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Clt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x - / 2 * u2 <= x < midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Ceq:(fexp1 (mag x) <= mag x)%Z -> x = midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Cgt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x < x <= midp fexp1 x + / 2 * u2 -> round_round_eq fexp1 fexp2 choice1 choice2 x
Heq:mag x = (fexp1 (mag x) - 1)%Z
round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
x':=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cz:fexp1 (mag x) = (mag x + 1)%Z -> bpow (mag x) - / 2 * u2 <= x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Clt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x - / 2 * u2 <= x < midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Ceq:(fexp1 (mag x) <= mag x)%Z -> x = midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Cgt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x < x <= midp fexp1 x + / 2 * u2 -> round_round_eq fexp1 fexp2 choice1 choice2 x
Hgt:(mag x > fexp1 (mag x) - 1)%Z
round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
x':=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cz:fexp1 (mag x) = (mag x + 1)%Z -> bpow (mag x) - / 2 * u2 <= x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Clt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x - / 2 * u2 <= x < midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Ceq:(fexp1 (mag x) <= mag x)%Z -> x = midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Cgt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x < x <= midp fexp1 x + / 2 * u2 -> round_round_eq fexp1 fexp2 choice1 choice2 x
Hlt:(mag x < fexp1 (mag x) - 1)%Z

round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
x':=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cz:fexp1 (mag x) = (mag x + 1)%Z -> bpow (mag x) - / 2 * u2 <= x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Clt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x - / 2 * u2 <= x < midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Ceq:(fexp1 (mag x) <= mag x)%Z -> x = midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Cgt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x < x <= midp fexp1 x + / 2 * u2 -> round_round_eq fexp1 fexp2 choice1 choice2 x
Hlt:(mag x < fexp1 (mag x) - 1)%Z
H:(mag x <= fexp1 (mag x) - 2)%Z

round_round_eq fexp1 fexp2 choice1 choice2 x
now apply round_round_really_zero.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
x':=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cz:fexp1 (mag x) = (mag x + 1)%Z -> bpow (mag x) - / 2 * u2 <= x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Clt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x - / 2 * u2 <= x < midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Ceq:(fexp1 (mag x) <= mag x)%Z -> x = midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Cgt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x < x <= midp fexp1 x + / 2 * u2 -> round_round_eq fexp1 fexp2 choice1 choice2 x
Heq:mag x = (fexp1 (mag x) - 1)%Z

round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
x':=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cz:fexp1 (mag x) = (mag x + 1)%Z -> bpow (mag x) - / 2 * u2 <= x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Clt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x - / 2 * u2 <= x < midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Ceq:(fexp1 (mag x) <= mag x)%Z -> x = midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Cgt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x < x <= midp fexp1 x + / 2 * u2 -> round_round_eq fexp1 fexp2 choice1 choice2 x
Heq:mag x = (fexp1 (mag x) - 1)%Z
H:fexp1 (mag x) = (mag x + 1)%Z

round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
x':=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cz:fexp1 (mag x) = (mag x + 1)%Z -> bpow (mag x) - / 2 * u2 <= x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Clt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x - / 2 * u2 <= x < midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Ceq:(fexp1 (mag x) <= mag x)%Z -> x = midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Cgt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x < x <= midp fexp1 x + / 2 * u2 -> round_round_eq fexp1 fexp2 choice1 choice2 x
Heq:mag x = (fexp1 (mag x) - 1)%Z
H:fexp1 (mag x) = (mag x + 1)%Z
Hlt':x < bpow (mag x) - / 2 * u2

round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
x':=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cz:fexp1 (mag x) = (mag x + 1)%Z -> bpow (mag x) - / 2 * u2 <= x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Clt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x - / 2 * u2 <= x < midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Ceq:(fexp1 (mag x) <= mag x)%Z -> x = midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Cgt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x < x <= midp fexp1 x + / 2 * u2 -> round_round_eq fexp1 fexp2 choice1 choice2 x
Heq:mag x = (fexp1 (mag x) - 1)%Z
H:fexp1 (mag x) = (mag x + 1)%Z
Hge':bpow (mag x) - / 2 * u2 <= x
round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
x':=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cz:fexp1 (mag x) = (mag x + 1)%Z -> bpow (mag x) - / 2 * u2 <= x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Clt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x - / 2 * u2 <= x < midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Ceq:(fexp1 (mag x) <= mag x)%Z -> x = midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Cgt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x < x <= midp fexp1 x + / 2 * u2 -> round_round_eq fexp1 fexp2 choice1 choice2 x
Heq:mag x = (fexp1 (mag x) - 1)%Z
H:fexp1 (mag x) = (mag x + 1)%Z
Hlt':x < bpow (mag x) - / 2 * u2

round_round_eq fexp1 fexp2 choice1 choice2 x
now apply round_round_zero.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
x':=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cz:fexp1 (mag x) = (mag x + 1)%Z -> bpow (mag x) - / 2 * u2 <= x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Clt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x - / 2 * u2 <= x < midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Ceq:(fexp1 (mag x) <= mag x)%Z -> x = midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Cgt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x < x <= midp fexp1 x + / 2 * u2 -> round_round_eq fexp1 fexp2 choice1 choice2 x
Heq:mag x = (fexp1 (mag x) - 1)%Z
H:fexp1 (mag x) = (mag x + 1)%Z
Hge':bpow (mag x) - / 2 * u2 <= x

round_round_eq fexp1 fexp2 choice1 choice2 x
now apply Cz.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
x':=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cz:fexp1 (mag x) = (mag x + 1)%Z -> bpow (mag x) - / 2 * u2 <= x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Clt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x - / 2 * u2 <= x < midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Ceq:(fexp1 (mag x) <= mag x)%Z -> x = midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Cgt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x < x <= midp fexp1 x + / 2 * u2 -> round_round_eq fexp1 fexp2 choice1 choice2 x
Hgt:(mag x > fexp1 (mag x) - 1)%Z

round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
x':=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cz:fexp1 (mag x) = (mag x + 1)%Z -> bpow (mag x) - / 2 * u2 <= x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Clt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x - / 2 * u2 <= x < midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Ceq:(fexp1 (mag x) <= mag x)%Z -> x = midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Cgt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x < x <= midp fexp1 x + / 2 * u2 -> round_round_eq fexp1 fexp2 choice1 choice2 x
Hgt:(mag x > fexp1 (mag x) - 1)%Z
H:(fexp1 (mag x) <= mag x)%Z

round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
x':=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cz:fexp1 (mag x) = (mag x + 1)%Z -> bpow (mag x) - / 2 * u2 <= x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Clt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x - / 2 * u2 <= x < midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Ceq:(fexp1 (mag x) <= mag x)%Z -> x = midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Cgt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x < x <= midp fexp1 x + / 2 * u2 -> round_round_eq fexp1 fexp2 choice1 choice2 x
Hgt:(mag x > fexp1 (mag x) - 1)%Z
H:(fexp1 (mag x) <= mag x)%Z
Hlt':x < midp fexp1 x

round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
x':=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cz:fexp1 (mag x) = (mag x + 1)%Z -> bpow (mag x) - / 2 * u2 <= x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Clt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x - / 2 * u2 <= x < midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Ceq:(fexp1 (mag x) <= mag x)%Z -> x = midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Cgt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x < x <= midp fexp1 x + / 2 * u2 -> round_round_eq fexp1 fexp2 choice1 choice2 x
Hgt:(mag x > fexp1 (mag x) - 1)%Z
H:(fexp1 (mag x) <= mag x)%Z
Heq':x = midp fexp1 x
round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
x':=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cz:fexp1 (mag x) = (mag x + 1)%Z -> bpow (mag x) - / 2 * u2 <= x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Clt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x - / 2 * u2 <= x < midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Ceq:(fexp1 (mag x) <= mag x)%Z -> x = midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Cgt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x < x <= midp fexp1 x + / 2 * u2 -> round_round_eq fexp1 fexp2 choice1 choice2 x
Hgt:(mag x > fexp1 (mag x) - 1)%Z
H:(fexp1 (mag x) <= mag x)%Z
Hgt':x > midp fexp1 x
round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
x':=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cz:fexp1 (mag x) = (mag x + 1)%Z -> bpow (mag x) - / 2 * u2 <= x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Clt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x - / 2 * u2 <= x < midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Ceq:(fexp1 (mag x) <= mag x)%Z -> x = midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Cgt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x < x <= midp fexp1 x + / 2 * u2 -> round_round_eq fexp1 fexp2 choice1 choice2 x
Hgt:(mag x > fexp1 (mag x) - 1)%Z
H:(fexp1 (mag x) <= mag x)%Z
Hlt':x < midp fexp1 x

round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
x':=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cz:fexp1 (mag x) = (mag x + 1)%Z -> bpow (mag x) - / 2 * u2 <= x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Clt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x - / 2 * u2 <= x < midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Ceq:(fexp1 (mag x) <= mag x)%Z -> x = midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Cgt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x < x <= midp fexp1 x + / 2 * u2 -> round_round_eq fexp1 fexp2 choice1 choice2 x
Hgt:(mag x > fexp1 (mag x) - 1)%Z
H:(fexp1 (mag x) <= mag x)%Z
Hlt':x < midp fexp1 x
Hlt'':x < midp fexp1 x - / 2 * u2

round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
x':=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cz:fexp1 (mag x) = (mag x + 1)%Z -> bpow (mag x) - / 2 * u2 <= x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Clt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x - / 2 * u2 <= x < midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Ceq:(fexp1 (mag x) <= mag x)%Z -> x = midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Cgt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x < x <= midp fexp1 x + / 2 * u2 -> round_round_eq fexp1 fexp2 choice1 choice2 x
Hgt:(mag x > fexp1 (mag x) - 1)%Z
H:(fexp1 (mag x) <= mag x)%Z
Hlt':x < midp fexp1 x
Hle'':midp fexp1 x - / 2 * u2 <= x
round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
x':=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cz:fexp1 (mag x) = (mag x + 1)%Z -> bpow (mag x) - / 2 * u2 <= x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Clt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x - / 2 * u2 <= x < midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Ceq:(fexp1 (mag x) <= mag x)%Z -> x = midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Cgt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x < x <= midp fexp1 x + / 2 * u2 -> round_round_eq fexp1 fexp2 choice1 choice2 x
Hgt:(mag x > fexp1 (mag x) - 1)%Z
H:(fexp1 (mag x) <= mag x)%Z
Hlt':x < midp fexp1 x
Hlt'':x < midp fexp1 x - / 2 * u2

round_round_eq fexp1 fexp2 choice1 choice2 x
now apply round_round_lt_mid_further_place; [| | |omega| |].
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
x':=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cz:fexp1 (mag x) = (mag x + 1)%Z -> bpow (mag x) - / 2 * u2 <= x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Clt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x - / 2 * u2 <= x < midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Ceq:(fexp1 (mag x) <= mag x)%Z -> x = midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Cgt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x < x <= midp fexp1 x + / 2 * u2 -> round_round_eq fexp1 fexp2 choice1 choice2 x
Hgt:(mag x > fexp1 (mag x) - 1)%Z
H:(fexp1 (mag x) <= mag x)%Z
Hlt':x < midp fexp1 x
Hle'':midp fexp1 x - / 2 * u2 <= x

round_round_eq fexp1 fexp2 choice1 choice2 x
now apply Clt; [|split].
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
x':=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cz:fexp1 (mag x) = (mag x + 1)%Z -> bpow (mag x) - / 2 * u2 <= x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Clt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x - / 2 * u2 <= x < midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Ceq:(fexp1 (mag x) <= mag x)%Z -> x = midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Cgt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x < x <= midp fexp1 x + / 2 * u2 -> round_round_eq fexp1 fexp2 choice1 choice2 x
Hgt:(mag x > fexp1 (mag x) - 1)%Z
H:(fexp1 (mag x) <= mag x)%Z
Heq':x = midp fexp1 x

round_round_eq fexp1 fexp2 choice1 choice2 x
now apply Ceq.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
x':=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cz:fexp1 (mag x) = (mag x + 1)%Z -> bpow (mag x) - / 2 * u2 <= x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Clt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x - / 2 * u2 <= x < midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Ceq:(fexp1 (mag x) <= mag x)%Z -> x = midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Cgt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x < x <= midp fexp1 x + / 2 * u2 -> round_round_eq fexp1 fexp2 choice1 choice2 x
Hgt:(mag x > fexp1 (mag x) - 1)%Z
H:(fexp1 (mag x) <= mag x)%Z
Hgt':x > midp fexp1 x

round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
x':=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cz:fexp1 (mag x) = (mag x + 1)%Z -> bpow (mag x) - / 2 * u2 <= x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Clt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x - / 2 * u2 <= x < midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Ceq:(fexp1 (mag x) <= mag x)%Z -> x = midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Cgt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x < x <= midp fexp1 x + / 2 * u2 -> round_round_eq fexp1 fexp2 choice1 choice2 x
Hgt:(mag x > fexp1 (mag x) - 1)%Z
H:(fexp1 (mag x) <= mag x)%Z
Hgt':x > midp fexp1 x
Hlt'':x <= midp fexp1 x + / 2 * u2

round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
x':=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cz:fexp1 (mag x) = (mag x + 1)%Z -> bpow (mag x) - / 2 * u2 <= x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Clt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x - / 2 * u2 <= x < midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Ceq:(fexp1 (mag x) <= mag x)%Z -> x = midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Cgt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x < x <= midp fexp1 x + / 2 * u2 -> round_round_eq fexp1 fexp2 choice1 choice2 x
Hgt:(mag x > fexp1 (mag x) - 1)%Z
H:(fexp1 (mag x) <= mag x)%Z
Hgt':x > midp fexp1 x
Hle'':midp fexp1 x + / 2 * u2 < x
round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
x':=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cz:fexp1 (mag x) = (mag x + 1)%Z -> bpow (mag x) - / 2 * u2 <= x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Clt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x - / 2 * u2 <= x < midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Ceq:(fexp1 (mag x) <= mag x)%Z -> x = midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Cgt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x < x <= midp fexp1 x + / 2 * u2 -> round_round_eq fexp1 fexp2 choice1 choice2 x
Hgt:(mag x > fexp1 (mag x) - 1)%Z
H:(fexp1 (mag x) <= mag x)%Z
Hgt':x > midp fexp1 x
Hlt'':x <= midp fexp1 x + / 2 * u2

round_round_eq fexp1 fexp2 choice1 choice2 x
now apply Cgt; [|split].
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
x':=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cz:fexp1 (mag x) = (mag x + 1)%Z -> bpow (mag x) - / 2 * u2 <= x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Clt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x - / 2 * u2 <= x < midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Ceq:(fexp1 (mag x) <= mag x)%Z -> x = midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Cgt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x < x <= midp fexp1 x + / 2 * u2 -> round_round_eq fexp1 fexp2 choice1 choice2 x
Hgt:(mag x > fexp1 (mag x) - 1)%Z
H:(fexp1 (mag x) <= mag x)%Z
Hgt':x > midp fexp1 x
Hle'':midp fexp1 x + / 2 * u2 < x

round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
x':=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cz:fexp1 (mag x) = (mag x + 1)%Z -> bpow (mag x) - / 2 * u2 <= x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Clt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x - / 2 * u2 <= x < midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Ceq:(fexp1 (mag x) <= mag x)%Z -> x = midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Cgt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x < x <= midp fexp1 x + / 2 * u2 -> round_round_eq fexp1 fexp2 choice1 choice2 x
Hgt:(mag x > fexp1 (mag x) - 1)%Z
H:(fexp1 (mag x) <= mag x)%Z
Hgt':x > midp fexp1 x
Hle'':midp fexp1 x + / 2 * u2 < x

round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
x':=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cz:fexp1 (mag x) = (mag x + 1)%Z -> bpow (mag x) - / 2 * u2 <= x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Clt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x - / 2 * u2 <= x < midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Ceq:(fexp1 (mag x) <= mag x)%Z -> x = midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Cgt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x < x <= midp fexp1 x + / 2 * u2 -> round_round_eq fexp1 fexp2 choice1 choice2 x
Hgt:(mag x > fexp1 (mag x) - 1)%Z
H:(fexp1 (mag x) <= mag x)%Z
Hgt':x > midp fexp1 x
Hle'':midp fexp1 x + / 2 * u2 < x
Fx:generic_format beta fexp1 x

round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
x':=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cz:fexp1 (mag x) = (mag x + 1)%Z -> bpow (mag x) - / 2 * u2 <= x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Clt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x - / 2 * u2 <= x < midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Ceq:(fexp1 (mag x) <= mag x)%Z -> x = midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Cgt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x < x <= midp fexp1 x + / 2 * u2 -> round_round_eq fexp1 fexp2 choice1 choice2 x
Hgt:(mag x > fexp1 (mag x) - 1)%Z
H:(fexp1 (mag x) <= mag x)%Z
Hgt':x > midp fexp1 x
Hle'':midp fexp1 x + / 2 * u2 < x
Nfx:~ generic_format beta fexp1 x
round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
x':=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cz:fexp1 (mag x) = (mag x + 1)%Z -> bpow (mag x) - / 2 * u2 <= x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Clt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x - / 2 * u2 <= x < midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Ceq:(fexp1 (mag x) <= mag x)%Z -> x = midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Cgt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x < x <= midp fexp1 x + / 2 * u2 -> round_round_eq fexp1 fexp2 choice1 choice2 x
Hgt:(mag x > fexp1 (mag x) - 1)%Z
H:(fexp1 (mag x) <= mag x)%Z
Hgt':x > midp fexp1 x
Hle'':midp fexp1 x + / 2 * u2 < x
Fx:generic_format beta fexp1 x

round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
x':=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cz:fexp1 (mag x) = (mag x + 1)%Z -> bpow (mag x) - / 2 * u2 <= x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Clt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x - / 2 * u2 <= x < midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Ceq:(fexp1 (mag x) <= mag x)%Z -> x = midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Cgt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x < x <= midp fexp1 x + / 2 * u2 -> round_round_eq fexp1 fexp2 choice1 choice2 x
Hgt:(mag x > fexp1 (mag x) - 1)%Z
H:(fexp1 (mag x) <= mag x)%Z
Hgt':x > midp fexp1 x
Hle'':midp fexp1 x + / 2 * u2 < x
Fx:generic_format beta fexp1 x

generic_format beta fexp2 x
now apply (generic_inclusion_mag beta fexp1); [omega|].
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
x':=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cz:fexp1 (mag x) = (mag x + 1)%Z -> bpow (mag x) - / 2 * u2 <= x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Clt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x - / 2 * u2 <= x < midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Ceq:(fexp1 (mag x) <= mag x)%Z -> x = midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Cgt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x < x <= midp fexp1 x + / 2 * u2 -> round_round_eq fexp1 fexp2 choice1 choice2 x
Hgt:(mag x > fexp1 (mag x) - 1)%Z
H:(fexp1 (mag x) <= mag x)%Z
Hgt':x > midp fexp1 x
Hle'':midp fexp1 x + / 2 * u2 < x
Nfx:~ generic_format beta fexp1 x

round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
x':=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cz:fexp1 (mag x) = (mag x + 1)%Z -> bpow (mag x) - / 2 * u2 <= x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Clt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x - / 2 * u2 <= x < midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Ceq:(fexp1 (mag x) <= mag x)%Z -> x = midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Cgt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x < x <= midp fexp1 x + / 2 * u2 -> round_round_eq fexp1 fexp2 choice1 choice2 x
Hgt:(mag x > fexp1 (mag x) - 1)%Z
H:(fexp1 (mag x) <= mag x)%Z
Hgt':x > midp fexp1 x
Hle'':midp fexp1 x + / 2 * u2 < x
Nfx:~ generic_format beta fexp1 x
Hceil:round beta fexp1 Zceil x = x' + u1

round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
x':=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cz:fexp1 (mag x) = (mag x + 1)%Z -> bpow (mag x) - / 2 * u2 <= x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Clt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x - / 2 * u2 <= x < midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Ceq:(fexp1 (mag x) <= mag x)%Z -> x = midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Cgt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x < x <= midp fexp1 x + / 2 * u2 -> round_round_eq fexp1 fexp2 choice1 choice2 x
Hgt:(mag x > fexp1 (mag x) - 1)%Z
H:(fexp1 (mag x) <= mag x)%Z
Hgt':x > midp fexp1 x
Hle'':midp fexp1 x + / 2 * u2 < x
Nfx:~ generic_format beta fexp1 x
Hceil:round beta fexp1 Zceil x = x' + u1
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z

round_round_eq fexp1 fexp2 choice1 choice2 x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
x':=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cz:fexp1 (mag x) = (mag x + 1)%Z -> bpow (mag x) - / 2 * u2 <= x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Clt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x - / 2 * u2 <= x < midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Ceq:(fexp1 (mag x) <= mag x)%Z -> x = midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Cgt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x < x <= midp fexp1 x + / 2 * u2 -> round_round_eq fexp1 fexp2 choice1 choice2 x
Hgt:(mag x > fexp1 (mag x) - 1)%Z
H:(fexp1 (mag x) <= mag x)%Z
Hgt':x > midp fexp1 x
Hle'':midp fexp1 x + / 2 * u2 < x
Nfx:~ generic_format beta fexp1 x
Hceil:round beta fexp1 Zceil x = x' + u1
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z

midp' fexp1 x + / 2 * ulp beta fexp2 x < x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
x':=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cz:fexp1 (mag x) = (mag x + 1)%Z -> bpow (mag x) - / 2 * u2 <= x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Clt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x - / 2 * u2 <= x < midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Ceq:(fexp1 (mag x) <= mag x)%Z -> x = midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Cgt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x < x <= midp fexp1 x + / 2 * u2 -> round_round_eq fexp1 fexp2 choice1 choice2 x
Hgt:(mag x > fexp1 (mag x) - 1)%Z
H:(fexp1 (mag x) <= mag x)%Z
Hgt':x > midp fexp1 x
Nfx:~ generic_format beta fexp1 x
Hceil:round beta fexp1 Zceil x = x' + u1
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z

x' + / 2 * ulp beta fexp1 x + / 2 * u2 < x -> round beta fexp1 Zceil x - / 2 * ulp beta fexp1 x + / 2 * ulp beta fexp2 x < x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
x:R
Px:0 < x
Hf2:(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z
x':=round beta fexp1 Zfloor x:R
u1:=ulp beta fexp1 x:R
u2:=ulp beta fexp2 x:R
Cz:fexp1 (mag x) = (mag x + 1)%Z -> bpow (mag x) - / 2 * u2 <= x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Clt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x - / 2 * u2 <= x < midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Ceq:(fexp1 (mag x) <= mag x)%Z -> x = midp fexp1 x -> round_round_eq fexp1 fexp2 choice1 choice2 x
Cgt:(fexp1 (mag x) <= mag x)%Z -> midp fexp1 x < x <= midp fexp1 x + / 2 * u2 -> round_round_eq fexp1 fexp2 choice1 choice2 x
Hgt:(mag x > fexp1 (mag x) - 1)%Z
H:(fexp1 (mag x) <= mag x)%Z
Hgt':x > midp fexp1 x
Nfx:~ generic_format beta fexp1 x
Hceil:round beta fexp1 Zceil x = x' + u1
Hf2':(fexp2 (mag x) <= fexp1 (mag x) - 1)%Z

x' + / 2 * u1 + / 2 * u2 < x -> x' + u1 - / 2 * u1 + / 2 * u2 < x
lra. } Qed.
beta:radix

forall x y : R, 0 < x -> 0 < y -> mag (x / y) = (mag x - mag y)%Z \/ mag (x / y) = (mag x - mag y + 1)%Z
beta:radix

forall x y : R, 0 < x -> 0 < y -> mag (x / y) = (mag x - mag y)%Z \/ mag (x / y) = (mag x - mag y + 1)%Z
beta:radix
x, y:R
Px:0 < x
Py:0 < y

mag (x / y) = (mag x - mag y)%Z \/ mag (x / y) = (mag x - mag y + 1)%Z
beta:radix
x, y:R
Px:0 < x
Py:0 < y

(mag x - mag y <= mag (x / y) <= mag x - mag y + 1)%Z -> mag (x / y) = (mag x - mag y)%Z \/ mag (x / y) = (mag x - mag y + 1)%Z
omega. Qed. Definition round_round_div_hyp fexp1 fexp2 := (forall ex, (fexp2 ex <= fexp1 ex - 1)%Z) /\ (forall ex ey, (fexp1 ex < ex)%Z -> (fexp1 ey < ey)%Z -> (fexp1 (ex - ey) <= ex - ey + 1)%Z -> (fexp2 (ex - ey) <= fexp1 ex - ey)%Z) /\ (forall ex ey, (fexp1 ex < ex)%Z -> (fexp1 ey < ey)%Z -> (fexp1 (ex - ey + 1) <= ex - ey + 1 + 1)%Z -> (fexp2 (ex - ey + 1) <= fexp1 ex - ey)%Z) /\ (forall ex ey, (fexp1 ex < ex)%Z -> (fexp1 ey < ey)%Z -> (fexp1 (ex - ey) <= ex - ey)%Z -> (fexp2 (ex - ey) <= fexp1 (ex - ey) + fexp1 ey - ey)%Z) /\ (forall ex ey, (fexp1 ex < ex)%Z -> (fexp1 ey < ey)%Z -> (fexp1 (ex - ey) = ex - ey + 1)%Z -> (fexp2 (ex - ey) <= ex - ey - ey + fexp1 ey)%Z).
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> (Z -> bool) -> (Z -> bool) -> round_round_div_hyp fexp1 fexp2 -> forall x y : R, 0 < x -> 0 < y -> generic_format beta fexp1 x -> generic_format beta fexp1 y -> fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z -> ~ bpow (mag (x / y)) - / 2 * ulp beta fexp2 (x / y) <= x / y
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> (Z -> bool) -> (Z -> bool) -> round_round_div_hyp fexp1 fexp2 -> forall x y : R, 0 < x -> 0 < y -> generic_format beta fexp1 x -> generic_format beta fexp1 y -> fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z -> ~ bpow (mag (x / y)) - / 2 * ulp beta fexp2 (x / y) <= x / y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z

~ bpow (mag (x / y)) - / 2 * ulp beta fexp2 (x / y) <= x / y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z

~ bpow (mag (x / y)) - / 2 * ulp beta fexp2 (x / y) <= x / y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z

~ bpow (mag (x / y)) - / 2 * ulp beta fexp2 (x / y) <= x / y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R

~ p - / 2 * ulp beta fexp2 (x / y) <= x / y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R

~ p - / 2 * ulp beta fexp2 (x / y) <= x / y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R

generic_format beta fexp1 x -> generic_format beta fexp1 y -> ~ p - / 2 * ulp beta fexp2 (x / y) <= x / y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R

x = IZR (Ztrunc (x * bpow (- fexp1 (mag x)))) * bpow (fexp1 (mag x)) -> y = IZR (Ztrunc (y * bpow (- fexp1 (mag y)))) * bpow (fexp1 (mag y)) -> ~ p - / 2 * ulp beta fexp2 (x / y) <= x / y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z

x = IZR mx * bpow (fexp1 (mag x)) -> y = IZR (Ztrunc (y * bpow (- fexp1 (mag y)))) * bpow (fexp1 (mag y)) -> ~ p - / 2 * ulp beta fexp2 (x / y) <= x / y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z

x = IZR mx * bpow (fexp1 (mag x)) -> y = IZR my * bpow (fexp1 (mag y)) -> ~ p - / 2 * ulp beta fexp2 (x / y) <= x / y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))

~ p - / 2 * ulp beta fexp2 (x / y) <= x / y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))

~ p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
x / y <> 0
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))

~ p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
/ y <> 0
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))

~ p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y

False
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p

False
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p

p - / 2 * u2 < p - / 2 * u2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p

x / y < p - / 2 * u2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p

x / y * y < (p - / 2 * u2) * y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p

x * (/ y * y) < (p - / 2 * u2) * y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p

x < (p - / 2 * u2) * y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z

x < (p - / 2 * u2) * y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(fexp1 (mag x) - mag (x / y) - fexp1 (mag y) < 0)%Z
x < (p - / 2 * u2) * y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z

x < (p - / 2 * u2) * y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z

x <= p * y - p * bpow (fexp1 (mag y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z
p * y - p * bpow (fexp1 (mag y)) < (p - / 2 * u2) * y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z

x <= p * y - p * bpow (fexp1 (mag y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z

IZR mx * bpow (fexp1 (mag x)) <= p * (IZR my * bpow (fexp1 (mag y))) - p * bpow (fexp1 (mag y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z

IZR mx * bpow (fexp1 (mag x)) <= p * IZR my * bpow (fexp1 (mag y)) - p * bpow (fexp1 (mag y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z

IZR mx * bpow (fexp1 (mag x)) <= IZR my * p * bpow (fexp1 (mag y)) - p * bpow (fexp1 (mag y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z

IZR mx * bpow (fexp1 (mag x)) <= IZR my * bpow (mag (x / y) + fexp1 (mag y)) - bpow (mag (x / y) + fexp1 (mag y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z

IZR mx * bpow (fexp1 (mag x)) * bpow (- mag (x / y) - fexp1 (mag y)) <= (IZR my * bpow (mag (x / y) + fexp1 (mag y)) - bpow (mag (x / y) + fexp1 (mag y))) * bpow (- mag (x / y) - fexp1 (mag y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z

IZR mx * bpow (fexp1 (mag x)) * bpow (- mag (x / y) - fexp1 (mag y)) <= IZR my * bpow (mag (x / y) + fexp1 (mag y)) * bpow (- mag (x / y) - fexp1 (mag y)) - bpow (mag (x / y) + fexp1 (mag y)) * bpow (- mag (x / y) - fexp1 (mag y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z

IZR mx * bpow (fexp1 (mag x) - mag (x / y) - fexp1 (mag y)) <= IZR my - 1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z

IZR mx * IZR (beta ^ (fexp1 (mag x) - mag (x / y) - fexp1 (mag y))) <= IZR my - 1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z

IZR (mx * beta ^ (fexp1 (mag x) - mag (x / y) - fexp1 (mag y))) <= IZR my - 1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z

IZR (mx * beta ^ (fexp1 (mag x) - mag (x / y) - fexp1 (mag y))) <= IZR (my - 1)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z

(mx * beta ^ (fexp1 (mag x) - mag (x / y) - fexp1 (mag y)) <= my - 1)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z

(mx * beta ^ (fexp1 (mag x) - mag (x / y) - fexp1 (mag y)) + 1 <= my)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z

(mx * beta ^ (fexp1 (mag x) - mag (x / y) - fexp1 (mag y)) < my)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z

IZR (mx * beta ^ (fexp1 (mag x) - mag (x / y) - fexp1 (mag y))) < IZR my
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z

IZR mx * IZR (beta ^ (fexp1 (mag x) - mag (x / y) - fexp1 (mag y))) < IZR my
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z

IZR mx * bpow (fexp1 (mag x) - mag (x / y) - fexp1 (mag y)) < IZR my
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z

IZR mx * bpow (fexp1 (mag x) - mag (x / y) - fexp1 (mag y)) * bpow (fexp1 (mag y) + mag (x / y)) < IZR my * bpow (fexp1 (mag y) + mag (x / y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z

IZR mx * bpow (fexp1 (mag x)) < IZR my * bpow (fexp1 (mag y) + mag (x / y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z

x < IZR my * bpow (fexp1 (mag y) + mag (x / y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z

x < IZR my * (bpow (fexp1 (mag y)) * bpow (mag (x / y)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z

x < y * bpow (mag (x / y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z

x < y * p
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z

x * / y < y * p * / y
field_simplify; lra.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z

p * y - p * bpow (fexp1 (mag y)) < (p - / 2 * u2) * y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z

p * y - p * bpow (fexp1 (mag y)) < p * y - / 2 * u2 * y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z

- (p * bpow (fexp1 (mag y))) < - (/ 2 * u2 * y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z

/ 2 * u2 * y < p * bpow (fexp1 (mag y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z

/ 2 * u2 * y < u2 * bpow (mag y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z
u2 * bpow (mag y) <= p * bpow (fexp1 (mag y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z

/ 2 * u2 * y < u2 * bpow (mag y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z

/ 2 * u2 * y < 1 * (u2 * bpow (mag y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z

/ 2 * (u2 * y) < 1 * (u2 * bpow (mag y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z

/ 2 * (u2 * y) < 1 * (u2 * bpow (mag y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z

0 <= / 2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z
0 <= u2 * y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z
/ 2 < 1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z
u2 * y < u2 * bpow (mag y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z

0 <= / 2
lra.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z

0 <= u2 * y
now apply Rmult_le_pos; [apply bpow_ge_0|apply Rlt_le].
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z

/ 2 < 1
lra.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z

u2 * y < u2 * bpow (mag y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z

y < bpow (mag y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z

Rabs y < bpow (mag y)
apply bpow_mag_gt. }
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z

u2 * bpow (mag y) <= p * bpow (fexp1 (mag y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z

(fexp2 (mag (x / y)) + mag y <= mag (x / y) + fexp1 (mag y))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z

(fexp2 (mag (x / y)) <= - mag y + mag (x / y) + fexp1 (mag y))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z

(fexp2 (mag (x / y)) <= mag (x / y) - mag y + fexp1 (mag y))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z
Hxy:mag (x / y) = (mag x - mag y + 1)%Z

(fexp2 (mag x - mag y + 1) <= mag x - mag y + 1 - mag y + fexp1 (mag y))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z
Hxy:mag (x / y) = (mag x - mag y + 1)%Z

(fexp2 (mag x + 1 - mag y) <= mag x + 1 - mag y - mag y + fexp1 (mag y))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z
Hxy:mag (x / y) = (mag x - mag y + 1)%Z

(fexp1 (mag x + 1) < mag x + 1)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z
Hxy:mag (x / y) = (mag x - mag y + 1)%Z
(fexp1 (mag y) < mag y)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z
Hxy:mag (x / y) = (mag x - mag y + 1)%Z
fexp1 (mag x + 1 - mag y)%Z = (mag x + 1 - mag y + 1)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z
Hxy:mag (x / y) = (mag x - mag y + 1)%Z

(fexp1 (mag x + 1) < mag x + 1)%Z
now assert (fexp1 (mag x + 1) <= mag x)%Z; [apply valid_exp|omega].
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z
Hxy:mag (x / y) = (mag x - mag y + 1)%Z

(fexp1 (mag y) < mag y)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z
Hxy:mag (x / y) = (mag x - mag y + 1)%Z
fexp1 (mag x + 1 - mag y)%Z = (mag x + 1 - mag y + 1)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z
Hxy:mag (x / y) = (mag x - mag y + 1)%Z

(fexp1 (mag y) < mag y)%Z
assumption.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z
Hxy:mag (x / y) = (mag x - mag y + 1)%Z

fexp1 (mag x + 1 - mag y)%Z = (mag x + 1 - mag y + 1)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(0 <= fexp1 (mag x) - mag (x / y) - fexp1 (mag y))%Z
Hxy:mag (x / y) = (mag x - mag y + 1)%Z

fexp1 (mag x - mag y + 1)%Z = (mag x - mag y + 1 + 1)%Z
now rewrite <- Hxy.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(fexp1 (mag x) - mag (x / y) - fexp1 (mag y) < 0)%Z

x < (p - / 2 * u2) * y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(fexp1 (mag x) - mag (x / y) - fexp1 (mag y) < 0)%Z

x <= p * y - bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(fexp1 (mag x) - mag (x / y) - fexp1 (mag y) < 0)%Z
p * y - bpow (fexp1 (mag x)) < (p - / 2 * u2) * y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(fexp1 (mag x) - mag (x / y) - fexp1 (mag y) < 0)%Z

x <= p * y - bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(fexp1 (mag x) - mag (x / y) - fexp1 (mag y) < 0)%Z

IZR mx * bpow (fexp1 (mag x)) <= p * (IZR my * bpow (fexp1 (mag y))) - bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(fexp1 (mag x) - mag (x / y) - fexp1 (mag y) < 0)%Z

IZR mx * bpow (fexp1 (mag x)) * bpow (- fexp1 (mag x)) <= (p * (IZR my * bpow (fexp1 (mag y))) - bpow (fexp1 (mag x))) * bpow (- fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(fexp1 (mag x) - mag (x / y) - fexp1 (mag y) < 0)%Z

IZR mx * bpow (fexp1 (mag x)) * bpow (- fexp1 (mag x)) <= p * (IZR my * bpow (fexp1 (mag y))) * bpow (- fexp1 (mag x)) - bpow (fexp1 (mag x)) * bpow (- fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(fexp1 (mag x) - mag (x / y) - fexp1 (mag y) < 0)%Z

IZR mx <= p * IZR my * bpow (fexp1 (mag y) - fexp1 (mag x)) - 1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(fexp1 (mag x) - mag (x / y) - fexp1 (mag y) < 0)%Z

IZR mx <= IZR my * p * bpow (fexp1 (mag y) - fexp1 (mag x)) - 1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(fexp1 (mag x) - mag (x / y) - fexp1 (mag y) < 0)%Z

IZR mx <= IZR my * bpow (mag (x / y) + fexp1 (mag y) - fexp1 (mag x)) - 1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(fexp1 (mag x) - mag (x / y) - fexp1 (mag y) < 0)%Z

IZR mx <= IZR my * IZR (beta ^ (mag (x / y) + fexp1 (mag y) - fexp1 (mag x))) - 1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(fexp1 (mag x) - mag (x / y) - fexp1 (mag y) < 0)%Z

IZR mx <= IZR (my * beta ^ (mag (x / y) + fexp1 (mag y) - fexp1 (mag x))) - 1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(fexp1 (mag x) - mag (x / y) - fexp1 (mag y) < 0)%Z

IZR mx <= IZR (my * beta ^ (mag (x / y) + fexp1 (mag y) - fexp1 (mag x)) - 1)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(fexp1 (mag x) - mag (x / y) - fexp1 (mag y) < 0)%Z

(mx <= my * beta ^ (mag (x / y) + fexp1 (mag y) - fexp1 (mag x)) - 1)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(fexp1 (mag x) - mag (x / y) - fexp1 (mag y) < 0)%Z

(mx + 1 <= my * beta ^ (mag (x / y) + fexp1 (mag y) - fexp1 (mag x)))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(fexp1 (mag x) - mag (x / y) - fexp1 (mag y) < 0)%Z

(mx < my * beta ^ (mag (x / y) + fexp1 (mag y) - fexp1 (mag x)))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(fexp1 (mag x) - mag (x / y) - fexp1 (mag y) < 0)%Z

IZR mx < IZR (my * beta ^ (mag (x / y) + fexp1 (mag y) - fexp1 (mag x)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(fexp1 (mag x) - mag (x / y) - fexp1 (mag y) < 0)%Z

IZR mx < IZR my * IZR (beta ^ (mag (x / y) + fexp1 (mag y) - fexp1 (mag x)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(fexp1 (mag x) - mag (x / y) - fexp1 (mag y) < 0)%Z

IZR mx < IZR my * bpow (mag (x / y) + fexp1 (mag y) - fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(fexp1 (mag x) - mag (x / y) - fexp1 (mag y) < 0)%Z

IZR mx * bpow (fexp1 (mag x)) < IZR my * bpow (mag (x / y) + fexp1 (mag y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(fexp1 (mag x) - mag (x / y) - fexp1 (mag y) < 0)%Z

x < IZR my * bpow (mag (x / y) + fexp1 (mag y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(fexp1 (mag x) - mag (x / y) - fexp1 (mag y) < 0)%Z

x < IZR my * (bpow (fexp1 (mag y)) * bpow (mag (x / y)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(fexp1 (mag x) - mag (x / y) - fexp1 (mag y) < 0)%Z

x < y * bpow (mag (x / y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(fexp1 (mag x) - mag (x / y) - fexp1 (mag y) < 0)%Z

x < y * p
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(fexp1 (mag x) - mag (x / y) - fexp1 (mag y) < 0)%Z

x * / y < y * p * / y
field_simplify; lra.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(fexp1 (mag x) - mag (x / y) - fexp1 (mag y) < 0)%Z

p * y - bpow (fexp1 (mag x)) < (p - / 2 * u2) * y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(fexp1 (mag x) - mag (x / y) - fexp1 (mag y) < 0)%Z

p * y - bpow (fexp1 (mag x)) < p * y - / 2 * u2 * y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(fexp1 (mag x) - mag (x / y) - fexp1 (mag y) < 0)%Z

- bpow (fexp1 (mag x)) < - (/ 2 * u2 * y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(fexp1 (mag x) - mag (x / y) - fexp1 (mag y) < 0)%Z

/ 2 * u2 * y < bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(fexp1 (mag x) - mag (x / y) - fexp1 (mag y) < 0)%Z

/ 2 * u2 * y < u2 * bpow (mag y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(fexp1 (mag x) - mag (x / y) - fexp1 (mag y) < 0)%Z
u2 * bpow (mag y) <= bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(fexp1 (mag x) - mag (x / y) - fexp1 (mag y) < 0)%Z

/ 2 * u2 * y < u2 * bpow (mag y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(fexp1 (mag x) - mag (x / y) - fexp1 (mag y) < 0)%Z

/ 2 * u2 * y < 1 * (u2 * bpow (mag y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(fexp1 (mag x) - mag (x / y) - fexp1 (mag y) < 0)%Z

/ 2 * (u2 * y) < 1 * (u2 * bpow (mag y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(fexp1 (mag x) - mag (x / y) - fexp1 (mag y) < 0)%Z

/ 2 * (u2 * y) < 1 * (u2 * bpow (mag y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(fexp1 (mag x) - mag (x / y) - fexp1 (mag y) < 0)%Z

0 <= / 2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(fexp1 (mag x) - mag (x / y) - fexp1 (mag y) < 0)%Z
0 <= u2 * y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(fexp1 (mag x) - mag (x / y) - fexp1 (mag y) < 0)%Z
/ 2 < 1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(fexp1 (mag x) - mag (x / y) - fexp1 (mag y) < 0)%Z
u2 * y < u2 * bpow (mag y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(fexp1 (mag x) - mag (x / y) - fexp1 (mag y) < 0)%Z

0 <= / 2
lra.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(fexp1 (mag x) - mag (x / y) - fexp1 (mag y) < 0)%Z

0 <= u2 * y
now apply Rmult_le_pos; [apply bpow_ge_0|apply Rlt_le].
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(fexp1 (mag x) - mag (x / y) - fexp1 (mag y) < 0)%Z

/ 2 < 1
lra.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(fexp1 (mag x) - mag (x / y) - fexp1 (mag y) < 0)%Z

u2 * y < u2 * bpow (mag y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(fexp1 (mag x) - mag (x / y) - fexp1 (mag y) < 0)%Z

y < bpow (mag y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(fexp1 (mag x) - mag (x / y) - fexp1 (mag y) < 0)%Z

Rabs y < bpow (mag y)
apply bpow_mag_gt. }
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(fexp1 (mag x) - mag (x / y) - fexp1 (mag y) < 0)%Z

u2 * bpow (mag y) <= bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(fexp1 (mag x) - mag (x / y) - fexp1 (mag y) < 0)%Z

(fexp2 (mag (x / y)) + mag y <= fexp1 (mag x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(fexp1 (mag x) - mag (x / y) - fexp1 (mag y) < 0)%Z

(fexp2 (mag (x / y)) <= - mag y + fexp1 (mag x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
p:=bpow (mag (x / y)):R
u2:=bpow (fexp2 (mag (x / y))):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hl:p - / 2 * bpow (cexp beta fexp2 (x / y)) <= x / y
Hr:x / y < p
He:(fexp1 (mag x) - mag (x / y) - fexp1 (mag y) < 0)%Z

(fexp2 (mag (x / y)) <= fexp1 (mag x) + - mag y)%Z
destruct (mag_div_disj x y Px Py) as [Hxy|Hxy]; rewrite Hxy; apply Hexp; try assumption; rewrite <- Hxy; rewrite Hf1; apply Z.le_refl. Qed.
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> (Z -> bool) -> (Z -> bool) -> round_round_div_hyp fexp1 fexp2 -> forall x y : R, 0 < x -> 0 < y -> generic_format beta fexp1 x -> generic_format beta fexp1 y -> (fexp1 (mag (x / y)) <= mag (x / y))%Z -> ~ midp fexp1 (x / y) - / 2 * ulp beta fexp2 (x / y) <= x / y < midp fexp1 (x / y)
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> (Z -> bool) -> (Z -> bool) -> round_round_div_hyp fexp1 fexp2 -> forall x y : R, 0 < x -> 0 < y -> generic_format beta fexp1 x -> generic_format beta fexp1 y -> (fexp1 (mag (x / y)) <= mag (x / y))%Z -> ~ midp fexp1 (x / y) - / 2 * ulp beta fexp2 (x / y) <= x / y < midp fexp1 (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z

~ midp fexp1 (x / y) - / 2 * ulp beta fexp2 (x / y) <= x / y < midp fexp1 (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z

~ midp fexp1 (x / y) - / 2 * ulp beta fexp2 (x / y) <= x / y < midp fexp1 (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z

~ midp fexp1 (x / y) - / 2 * ulp beta fexp2 (x / y) <= x / y < midp fexp1 (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z

x / y <> 0
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
~ midp fexp1 (x / y) - / 2 * ulp beta fexp2 (x / y) <= x / y < midp fexp1 (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z

/ y <> 0
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
~ midp fexp1 (x / y) - / 2 * ulp beta fexp2 (x / y) <= x / y < midp fexp1 (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0

~ midp fexp1 (x / y) - / 2 * ulp beta fexp2 (x / y) <= x / y < midp fexp1 (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0

~ / 2 * (ulp beta fexp1 (x / y) - ulp beta fexp2 (x / y)) <= x / y - round beta fexp1 Zfloor (x / y) < / 2 * ulp beta fexp1 (x / y) -> ~ midp fexp1 (x / y) - / 2 * ulp beta fexp2 (x / y) <= x / y < midp fexp1 (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
~ / 2 * (ulp beta fexp1 (x / y) - ulp beta fexp2 (x / y)) <= x / y - round beta fexp1 Zfloor (x / y) < / 2 * ulp beta fexp1 (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0

~ / 2 * (ulp beta fexp1 (x / y) - ulp beta fexp2 (x / y)) <= x / y - round beta fexp1 Zfloor (x / y) < / 2 * ulp beta fexp1 (x / y) -> ~ midp fexp1 (x / y) - / 2 * ulp beta fexp2 (x / y) <= x / y < midp fexp1 (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
H:~ / 2 * (ulp beta fexp1 (x / y) - ulp beta fexp2 (x / y)) <= x / y - round beta fexp1 Zfloor (x / y) < / 2 * ulp beta fexp1 (x / y)
H':midp fexp1 (x / y) - / 2 * ulp beta fexp2 (x / y) <= x / y < midp fexp1 (x / y)

/ 2 * (ulp beta fexp1 (x / y) - ulp beta fexp2 (x / y)) <= x / y - round beta fexp1 Zfloor (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
H:~ / 2 * (ulp beta fexp1 (x / y) - ulp beta fexp2 (x / y)) <= x / y - round beta fexp1 Zfloor (x / y) < / 2 * ulp beta fexp1 (x / y)
H':midp fexp1 (x / y) - / 2 * ulp beta fexp2 (x / y) <= x / y < midp fexp1 (x / y)
x / y - round beta fexp1 Zfloor (x / y) < / 2 * ulp beta fexp1 (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
H:~ / 2 * (ulp beta fexp1 (x / y) - ulp beta fexp2 (x / y)) <= x / y - round beta fexp1 Zfloor (x / y) < / 2 * ulp beta fexp1 (x / y)
H':midp fexp1 (x / y) - / 2 * ulp beta fexp2 (x / y) <= x / y < midp fexp1 (x / y)

/ 2 * (ulp beta fexp1 (x / y) - ulp beta fexp2 (x / y)) <= x / y - round beta fexp1 Zfloor (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
H:~ / 2 * (ulp beta fexp1 (x / y) - ulp beta fexp2 (x / y)) <= x / y - round beta fexp1 Zfloor (x / y) < / 2 * ulp beta fexp1 (x / y)
H':midp fexp1 (x / y) - / 2 * ulp beta fexp2 (x / y) <= x / y < midp fexp1 (x / y)

round beta fexp1 Zfloor (x / y) + / 2 * (ulp beta fexp1 (x / y) - ulp beta fexp2 (x / y)) <= round beta fexp1 Zfloor (x / y) + (x / y - round beta fexp1 Zfloor (x / y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
H:~ / 2 * (ulp beta fexp1 (x / y) - ulp beta fexp2 (x / y)) <= x / y - round beta fexp1 Zfloor (x / y) < / 2 * ulp beta fexp1 (x / y)
H':midp fexp1 (x / y) - / 2 * ulp beta fexp2 (x / y) <= x / y < midp fexp1 (x / y)

round beta fexp1 Zfloor (x / y) + / 2 * ulp beta fexp1 (x / y) - / 2 * ulp beta fexp2 (x / y) <= x / y
apply H'.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
H:~ / 2 * (ulp beta fexp1 (x / y) - ulp beta fexp2 (x / y)) <= x / y - round beta fexp1 Zfloor (x / y) < / 2 * ulp beta fexp1 (x / y)
H':midp fexp1 (x / y) - / 2 * ulp beta fexp2 (x / y) <= x / y < midp fexp1 (x / y)

x / y - round beta fexp1 Zfloor (x / y) < / 2 * ulp beta fexp1 (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
H:~ / 2 * (ulp beta fexp1 (x / y) - ulp beta fexp2 (x / y)) <= x / y - round beta fexp1 Zfloor (x / y) < / 2 * ulp beta fexp1 (x / y)
H':midp fexp1 (x / y) - / 2 * ulp beta fexp2 (x / y) <= x / y < midp fexp1 (x / y)

round beta fexp1 Zfloor (x / y) + (x / y - round beta fexp1 Zfloor (x / y)) < round beta fexp1 Zfloor (x / y) + / 2 * ulp beta fexp1 (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
H:~ / 2 * (ulp beta fexp1 (x / y) - ulp beta fexp2 (x / y)) <= x / y - round beta fexp1 Zfloor (x / y) < / 2 * ulp beta fexp1 (x / y)
H':midp fexp1 (x / y) - / 2 * ulp beta fexp2 (x / y) <= x / y < midp fexp1 (x / y)

x / y < round beta fexp1 Zfloor (x / y) + / 2 * ulp beta fexp1 (x / y)
apply H'.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0

~ / 2 * (ulp beta fexp1 (x / y) - ulp beta fexp2 (x / y)) <= x / y - round beta fexp1 Zfloor (x / y) < / 2 * ulp beta fexp1 (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R

~ / 2 * (ulp beta fexp1 (x / y) - ulp beta fexp2 (x / y)) <= x / y - round beta fexp1 Zfloor (x / y) < / 2 * ulp beta fexp1 (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R

~ / 2 * (ulp beta fexp1 (x / y) - ulp beta fexp2 (x / y)) <= x / y - round beta fexp1 Zfloor (x / y) < / 2 * ulp beta fexp1 (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R

~ / 2 * (ulp beta fexp1 (x / y) - ulp beta fexp2 (x / y)) <= x / y - x' < / 2 * ulp beta fexp1 (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R

~ / 2 * (bpow (cexp beta fexp1 (x / y)) - bpow (cexp beta fexp2 (x / y))) <= x / y - x' < / 2 * bpow (cexp beta fexp1 (x / y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R

generic_format beta fexp1 x -> generic_format beta fexp1 y -> ~ / 2 * (bpow (cexp beta fexp1 (x / y)) - bpow (cexp beta fexp2 (x / y))) <= x / y - x' < / 2 * bpow (cexp beta fexp1 (x / y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R

x = IZR (Ztrunc (x * bpow (- fexp1 (mag x)))) * bpow (fexp1 (mag x)) -> y = IZR (Ztrunc (y * bpow (- fexp1 (mag y)))) * bpow (fexp1 (mag y)) -> ~ / 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z

x = IZR mx * bpow (fexp1 (mag x)) -> y = IZR (Ztrunc (y * bpow (- fexp1 (mag y)))) * bpow (fexp1 (mag y)) -> ~ / 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z

x = IZR mx * bpow (fexp1 (mag x)) -> y = IZR my * bpow (fexp1 (mag y)) -> ~ / 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))

~ / 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))

False
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))

/ 2 * (u1 - u2) < / 2 * (u1 - u2)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))

x / y - x' < / 2 * (u1 - u2)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))

x / y < x' + / 2 * u1 - / 2 * u2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))

x / y * y < (x' + / 2 * u1 - / 2 * u2) * y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))

x * (/ y * y) < (x' + / 2 * u1 - / 2 * u2) * y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))

x < (x' + / 2 * u1 - / 2 * u2) * y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))

x < x' * y + / 2 * u1 * y - / 2 * u2 * y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))

2 * x < 2 * (x' * y + / 2 * u1 * y - / 2 * u2 * y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))

2 * x < 2 * (x' * y) + 2 * (/ 2 * u1 * y) - 2 * (/ 2 * u2 * y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))

2 * x < 2 * x' * y + 2 * / 2 * u1 * y - 2 * / 2 * u2 * y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))

2 * x < 2 * x' * y + u1 * y - u2 * y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

2 * x < 2 * x' * y + u1 * y - u2 * y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z
2 * x < 2 * x' * y + u1 * y - u2 * y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

2 * x < 2 * x' * y + u1 * y - u2 * y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

2 * x <= 2 * x' * y + u1 * y - bpow (fexp1 (mag (x / y)) + fexp1 (mag y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z
2 * x' * y + u1 * y - bpow (fexp1 (mag (x / y)) + fexp1 (mag y)) < 2 * x' * y + u1 * y - u2 * y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

2 * x <= 2 * x' * y + u1 * y - bpow (fexp1 (mag (x / y)) + fexp1 (mag y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

2 * (IZR mx * bpow (fexp1 (mag x))) <= 2 * x' * (IZR my * bpow (fexp1 (mag y))) + u1 * (IZR my * bpow (fexp1 (mag y))) - bpow (fexp1 (mag (x / y)) + fexp1 (mag y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

2 * (IZR mx * bpow (fexp1 (mag x))) * bpow (- fexp1 (mag (x / y)) - fexp1 (mag y)) <= (2 * x' * (IZR my * bpow (fexp1 (mag y))) + u1 * (IZR my * bpow (fexp1 (mag y))) - bpow (fexp1 (mag (x / y)) + fexp1 (mag y))) * bpow (- fexp1 (mag (x / y)) - fexp1 (mag y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

2 * (IZR mx * bpow (fexp1 (mag x))) * bpow (- fexp1 (mag (x / y)) - fexp1 (mag y)) <= 2 * x' * (IZR my * bpow (fexp1 (mag y))) * bpow (- fexp1 (mag (x / y)) - fexp1 (mag y)) + u1 * (IZR my * bpow (fexp1 (mag y))) * bpow (- fexp1 (mag (x / y)) - fexp1 (mag y)) - bpow (fexp1 (mag (x / y)) + fexp1 (mag y)) * bpow (- fexp1 (mag (x / y)) - fexp1 (mag y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

2 * IZR mx * bpow (fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y)) <= 2 * x' * IZR my * bpow (- fexp1 (mag (x / y))) + u1 * IZR my * bpow (- fexp1 (mag (x / y))) - 1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

2 * IZR mx * bpow (fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y)) <= 2 * IZR my * x' * bpow (- fexp1 (mag (x / y))) + u1 * IZR my * bpow (- fexp1 (mag (x / y))) - 1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

2 * IZR mx * bpow (fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y)) <= 2 * IZR my * x' * bpow (- fexp1 (mag (x / y))) + IZR my * u1 * bpow (- fexp1 (mag (x / y))) - 1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

2 * IZR mx * bpow (fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y)) <= 2 * IZR my * (IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) * bpow (fexp1 (mag (x / y)))) * bpow (- fexp1 (mag (x / y))) + IZR my * bpow (fexp1 (mag (x / y))) * bpow (- fexp1 (mag (x / y))) - 1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

2 * IZR mx * bpow (fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y)) <= 2 * IZR my * IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) + IZR my - 1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

2 * IZR mx * IZR (beta ^ (fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))) <= 2 * IZR my * IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) + IZR my - 1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

IZR (2 * mx * beta ^ (fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))) <= IZR (2 * my * Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) + IZR my - 1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

IZR (2 * mx * beta ^ (fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))) <= IZR (2 * my * Zfloor (x / y * bpow (- fexp1 (mag (x / y)))) + my) - 1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

IZR (2 * mx * beta ^ (fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))) <= IZR (2 * my * Zfloor (x / y * bpow (- fexp1 (mag (x / y)))) + my - 1)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

(2 * mx * beta ^ (fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y)) <= 2 * my * Zfloor (x / y * bpow (- fexp1 (mag (x / y)))) + my - 1)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

(2 * mx * beta ^ (fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y)) + 1 <= 2 * my * Zfloor (x / y * bpow (- fexp1 (mag (x / y)))) + my)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

(2 * mx * beta ^ (fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y)) < 2 * my * Zfloor (x / y * bpow (- fexp1 (mag (x / y)))) + my)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

IZR (2 * mx * beta ^ (fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))) < IZR (2 * my * Zfloor (x / y * bpow (- fexp1 (mag (x / y)))) + my)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

IZR (2 * mx * beta ^ (fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))) < IZR (2 * my * Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) + IZR my
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

2 * IZR mx * IZR (beta ^ (fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))) < 2 * IZR my * IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) + IZR my
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

2 * IZR mx * bpow (fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y)) < 2 * IZR my * IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) + IZR my
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

2 * IZR mx * bpow (fexp1 (mag x)) < (2 * IZR my * IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) + IZR my) * bpow (fexp1 (mag (x / y)) + fexp1 (mag y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

2 * (IZR mx * bpow (fexp1 (mag x))) < (2 * IZR my * IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) + IZR my) * bpow (fexp1 (mag (x / y)) + fexp1 (mag y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

2 * x < (2 * IZR my * IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) + IZR my) * bpow (fexp1 (mag (x / y)) + fexp1 (mag y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

2 * x < 2 * IZR my * IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) * bpow (fexp1 (mag (x / y)) + fexp1 (mag y)) + IZR my * bpow (fexp1 (mag (x / y)) + fexp1 (mag y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

2 * x < 2 * IZR my * (IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) * bpow (fexp1 (mag (x / y)) + fexp1 (mag y))) + IZR my * bpow (fexp1 (mag (x / y)) + fexp1 (mag y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

2 * x < 2 * IZR my * (IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) * (bpow (fexp1 (mag (x / y))) * bpow (fexp1 (mag y)))) + IZR my * (bpow (fexp1 (mag (x / y))) * bpow (fexp1 (mag y)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

2 * x < 2 * IZR my * (IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) * bpow (fexp1 (mag (x / y))) * bpow (fexp1 (mag y))) + IZR my * (bpow (fexp1 (mag (x / y))) * bpow (fexp1 (mag y)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

2 * x < 2 * IZR my * (x' * bpow (fexp1 (mag y))) + IZR my * (bpow (fexp1 (mag (x / y))) * bpow (fexp1 (mag y)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

2 * x < 2 * IZR my * (bpow (fexp1 (mag y)) * x') + IZR my * (bpow (fexp1 (mag y)) * bpow (fexp1 (mag (x / y))))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

2 * x < 2 * (IZR my * (bpow (fexp1 (mag y)) * x')) + IZR my * (bpow (fexp1 (mag y)) * bpow (fexp1 (mag (x / y))))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

2 * x < 2 * (IZR my * bpow (fexp1 (mag y)) * x') + IZR my * bpow (fexp1 (mag y)) * bpow (fexp1 (mag (x / y)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

2 * x < 2 * (y * x') + y * bpow (fexp1 (mag (x / y)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

2 * x < 2 * (y * x') + y * u1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

/ 2 * (2 * x) < / 2 * (2 * (y * x') + y * u1)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

/ 2 * (2 * x) < / 2 * (2 * (y * x')) + / 2 * (y * u1)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

/ 2 * 2 * x < / 2 * 2 * y * x' + / 2 * y * u1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

x < y * x' + / 2 * y * u1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

x - y * x' < y * / 2 * u1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

/ y * (x - y * x') < / y * (y * / 2 * u1)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

/ y * x - / y * (y * x') < / y * (y * / 2 * u1)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

/ y * x - / y * y * x' < / y * y * / 2 * u1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

/ y * x - x' < / 2 * u1
now rewrite Rmult_comm.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

2 * x' * y + u1 * y - bpow (fexp1 (mag (x / y)) + fexp1 (mag y)) < 2 * x' * y + u1 * y - u2 * y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

- bpow (fexp1 (mag (x / y)) + fexp1 (mag y)) < - (u2 * y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

u2 * y < bpow (fexp1 (mag (x / y)) + fexp1 (mag y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

u2 * y < u2 * bpow (mag y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z
u2 * bpow (mag y) <= bpow (fexp1 (mag (x / y)) + fexp1 (mag y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

u2 * y < u2 * bpow (mag y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

u2 * y < u2 * bpow (mag y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

0 < u2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z
y < bpow (mag y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

0 < u2
apply bpow_gt_0.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

y < bpow (mag y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

Rabs y < bpow (mag y)
apply bpow_mag_gt. }
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

u2 * bpow (mag y) <= bpow (fexp1 (mag (x / y)) + fexp1 (mag y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

(fexp2 (mag (x / y)) + mag y <= fexp1 (mag (x / y)) + fexp1 (mag y))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

(fexp2 (mag (x / y)) <= - mag y + fexp1 (mag (x / y)) + fexp1 (mag y))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

(fexp2 (mag (x / y)) <= fexp1 (mag (x / y)) + fexp1 (mag y) + - mag y)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z
Hxy:mag (x / y) = (mag x - mag y + 1)%Z

(fexp2 (mag x - mag y + 1) <= fexp1 (mag x - mag y + 1) + fexp1 (mag y) + - mag y)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z
Hxy:mag (x / y) = (mag x - mag y + 1)%Z

(fexp2 (mag x + 1 - mag y) <= fexp1 (mag x + 1 - mag y) + fexp1 (mag y) + - mag y)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z
Hxy:mag (x / y) = (mag x - mag y + 1)%Z

(fexp1 (mag x + 1) < mag x + 1)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z
Hxy:mag (x / y) = (mag x - mag y + 1)%Z
(fexp1 (mag y) < mag y)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z
Hxy:mag (x / y) = (mag x - mag y + 1)%Z
(fexp1 (mag x + 1 - mag y) <= mag x + 1 - mag y)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z
Hxy:mag (x / y) = (mag x - mag y + 1)%Z

(fexp1 (mag x + 1) < mag x + 1)%Z
now assert (fexp1 (mag x + 1) <= mag x)%Z; [apply valid_exp|omega].
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z
Hxy:mag (x / y) = (mag x - mag y + 1)%Z

(fexp1 (mag y) < mag y)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z
Hxy:mag (x / y) = (mag x - mag y + 1)%Z
(fexp1 (mag x + 1 - mag y) <= mag x + 1 - mag y)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z
Hxy:mag (x / y) = (mag x - mag y + 1)%Z

(fexp1 (mag y) < mag y)%Z
assumption.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z
Hxy:mag (x / y) = (mag x - mag y + 1)%Z

(fexp1 (mag x + 1 - mag y) <= mag x + 1 - mag y)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z
Hxy:mag (x / y) = (mag x - mag y + 1)%Z

(fexp1 (mag x - mag y + 1) <= mag x - mag y + 1)%Z
now rewrite <- Hxy.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

2 * x < 2 * x' * y + u1 * y - u2 * y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

2 * x <= 2 * x' * y + u1 * y - bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z
2 * x' * y + u1 * y - bpow (fexp1 (mag x)) < 2 * x' * y + u1 * y - u2 * y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

2 * x <= 2 * x' * y + u1 * y - bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

2 * (IZR mx * bpow (fexp1 (mag x))) <= 2 * x' * (IZR my * bpow (fexp1 (mag y))) + u1 * (IZR my * bpow (fexp1 (mag y))) - bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

2 * (IZR mx * bpow (fexp1 (mag x))) * bpow (- fexp1 (mag x)) <= (2 * x' * (IZR my * bpow (fexp1 (mag y))) + u1 * (IZR my * bpow (fexp1 (mag y))) - bpow (fexp1 (mag x))) * bpow (- fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

2 * (IZR mx * bpow (fexp1 (mag x))) * bpow (- fexp1 (mag x)) <= 2 * x' * (IZR my * bpow (fexp1 (mag y))) * bpow (- fexp1 (mag x)) + u1 * (IZR my * bpow (fexp1 (mag y))) * bpow (- fexp1 (mag x)) - bpow (fexp1 (mag x)) * bpow (- fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

2 * IZR mx <= 2 * x' * IZR my * bpow (fexp1 (mag y) - fexp1 (mag x)) + u1 * IZR my * bpow (fexp1 (mag y) - fexp1 (mag x)) - 1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

2 * IZR mx <= 2 * IZR my * x' * bpow (fexp1 (mag y) - fexp1 (mag x)) + u1 * IZR my * bpow (fexp1 (mag y) - fexp1 (mag x)) - 1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

2 * IZR mx <= 2 * IZR my * x' * bpow (fexp1 (mag y) - fexp1 (mag x)) + IZR my * u1 * bpow (fexp1 (mag y) - fexp1 (mag x)) - 1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

2 * IZR mx <= 2 * IZR my * (IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) * bpow (fexp1 (mag (x / y)))) * bpow (fexp1 (mag y) - fexp1 (mag x)) + IZR my * bpow (fexp1 (mag (x / y))) * bpow (fexp1 (mag y) - fexp1 (mag x)) - 1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

2 * IZR mx <= 2 * IZR my * IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) * bpow (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x)) + IZR my * bpow (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x)) - 1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

2 * IZR mx <= 2 * IZR my * IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) * IZR (beta ^ (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x))) + IZR my * IZR (beta ^ (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x))) - 1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

IZR (2 * mx) <= IZR (2 * my * Zfloor (x / y * bpow (- fexp1 (mag (x / y)))) * beta ^ (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x))) + IZR (my * beta ^ (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x))) - 1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

IZR (2 * mx) <= IZR (2 * my * Zfloor (x / y * bpow (- fexp1 (mag (x / y)))) * beta ^ (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x)) + my * beta ^ (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x))) - 1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

IZR (2 * mx) <= IZR (2 * my * Zfloor (x / y * bpow (- fexp1 (mag (x / y)))) * beta ^ (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x)) + my * beta ^ (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x)) - 1)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

(2 * mx <= 2 * my * Zfloor (x / y * bpow (- fexp1 (mag (x / y)))) * beta ^ (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x)) + my * beta ^ (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x)) - 1)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

(2 * mx + 1 <= 2 * my * Zfloor (x / y * bpow (- fexp1 (mag (x / y)))) * beta ^ (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x)) + my * beta ^ (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x)))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

(2 * mx < 2 * my * Zfloor (x / y * bpow (- fexp1 (mag (x / y)))) * beta ^ (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x)) + my * beta ^ (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x)))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

IZR (2 * mx) < IZR (2 * my * Zfloor (x / y * bpow (- fexp1 (mag (x / y)))) * beta ^ (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x)) + my * beta ^ (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

IZR (2 * mx) < IZR (2 * my * Zfloor (x / y * bpow (- fexp1 (mag (x / y)))) * beta ^ (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x))) + IZR (my * beta ^ (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

2 * IZR mx < 2 * IZR my * IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) * IZR (beta ^ (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x))) + IZR my * IZR (beta ^ (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

2 * IZR mx < 2 * IZR my * IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) * bpow (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x)) + IZR my * bpow (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

2 * IZR mx * bpow (fexp1 (mag x)) < (2 * IZR my * IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) * bpow (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x)) + IZR my * bpow (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x))) * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

2 * (IZR mx * bpow (fexp1 (mag x))) < (2 * IZR my * IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) * bpow (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x)) + IZR my * bpow (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x))) * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

2 * x < (2 * IZR my * IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) * bpow (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x)) + IZR my * bpow (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x))) * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

2 * x < 2 * IZR my * IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) * bpow (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x)) * bpow (fexp1 (mag x)) + IZR my * bpow (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x)) * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

2 * x < 2 * IZR my * IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) * bpow (fexp1 (mag (x / y)) + fexp1 (mag y)) + IZR my * bpow (fexp1 (mag (x / y)) + fexp1 (mag y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

2 * x < 2 * IZR my * (IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) * bpow (fexp1 (mag (x / y)) + fexp1 (mag y))) + IZR my * bpow (fexp1 (mag (x / y)) + fexp1 (mag y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

2 * x < 2 * IZR my * (IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) * (bpow (fexp1 (mag (x / y))) * bpow (fexp1 (mag y)))) + IZR my * (bpow (fexp1 (mag (x / y))) * bpow (fexp1 (mag y)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

2 * x < 2 * IZR my * (IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) * bpow (fexp1 (mag (x / y))) * bpow (fexp1 (mag y))) + IZR my * (bpow (fexp1 (mag (x / y))) * bpow (fexp1 (mag y)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

2 * x < 2 * IZR my * (x' * bpow (fexp1 (mag y))) + IZR my * (bpow (fexp1 (mag (x / y))) * bpow (fexp1 (mag y)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

2 * x < 2 * IZR my * (bpow (fexp1 (mag y)) * x') + IZR my * (bpow (fexp1 (mag y)) * bpow (fexp1 (mag (x / y))))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

2 * x < 2 * (IZR my * (bpow (fexp1 (mag y)) * x')) + IZR my * (bpow (fexp1 (mag y)) * bpow (fexp1 (mag (x / y))))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

2 * x < 2 * (IZR my * bpow (fexp1 (mag y)) * x') + IZR my * bpow (fexp1 (mag y)) * bpow (fexp1 (mag (x / y)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

2 * x < 2 * (y * x') + y * bpow (fexp1 (mag (x / y)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

2 * x < 2 * (y * x') + y * u1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

/ 2 * (2 * x) < / 2 * (2 * (y * x') + y * u1)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

/ 2 * (2 * x) < / 2 * (2 * (y * x')) + / 2 * (y * u1)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

/ 2 * 2 * x < / 2 * 2 * y * x' + / 2 * y * u1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

x < y * x' + / 2 * y * u1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

x - y * x' < y * / 2 * u1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

/ y * (x - y * x') < / y * (y * / 2 * u1)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

/ y * x - / y * (y * x') < / y * (y * / 2 * u1)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

/ y * x - / y * y * x' < / y * y * / 2 * u1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

/ y * x - x' < / 2 * u1
now rewrite Rmult_comm.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

2 * x' * y + u1 * y - bpow (fexp1 (mag x)) < 2 * x' * y + u1 * y - u2 * y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

- bpow (fexp1 (mag x)) < - (u2 * y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

u2 * y < bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

u2 * y < u2 * bpow (mag y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z
u2 * bpow (mag y) <= bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

u2 * y < u2 * bpow (mag y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

u2 * y < u2 * bpow (mag y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

0 < u2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z
y < bpow (mag y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

0 < u2
apply bpow_gt_0.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

y < bpow (mag y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

Rabs y < bpow (mag y)
apply bpow_mag_gt. }
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

u2 * bpow (mag y) <= bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

(fexp2 (mag (x / y)) + mag y <= fexp1 (mag x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

(fexp2 (mag (x / y)) <= - mag y + fexp1 (mag x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
S:x / y <> 0
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * (bpow (fexp1 (mag (x / y))) - bpow (fexp2 (mag (x / y)))) <= x / y - x' < / 2 * bpow (fexp1 (mag (x / y)))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

(fexp2 (mag (x / y)) <= fexp1 (mag x) + - mag y)%Z
destruct (mag_div_disj x y Px Py) as [Hxy|Hxy]; rewrite Hxy; apply Hexp; try assumption; rewrite <- Hxy; omega. Qed.
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> (Z -> bool) -> (Z -> bool) -> round_round_div_hyp fexp1 fexp2 -> forall x y : R, 0 < x -> 0 < y -> generic_format beta fexp1 x -> generic_format beta fexp1 y -> (fexp1 (mag (x / y)) <= mag (x / y))%Z -> ~ midp fexp1 (x / y) < x / y <= midp fexp1 (x / y) + / 2 * ulp beta fexp2 (x / y)
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> (Z -> bool) -> (Z -> bool) -> round_round_div_hyp fexp1 fexp2 -> forall x y : R, 0 < x -> 0 < y -> generic_format beta fexp1 x -> generic_format beta fexp1 y -> (fexp1 (mag (x / y)) <= mag (x / y))%Z -> ~ midp fexp1 (x / y) < x / y <= midp fexp1 (x / y) + / 2 * ulp beta fexp2 (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z

~ midp fexp1 (x / y) < x / y <= midp fexp1 (x / y) + / 2 * ulp beta fexp2 (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z

~ midp fexp1 (x / y) < x / y <= midp fexp1 (x / y) + / 2 * ulp beta fexp2 (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z

~ midp fexp1 (x / y) < x / y <= midp fexp1 (x / y) + / 2 * ulp beta fexp2 (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z

~ / 2 * ulp beta fexp1 (x / y) < x / y - round beta fexp1 Zfloor (x / y) <= / 2 * (ulp beta fexp1 (x / y) + ulp beta fexp2 (x / y)) -> ~ midp fexp1 (x / y) < x / y <= midp fexp1 (x / y) + / 2 * ulp beta fexp2 (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
~ / 2 * ulp beta fexp1 (x / y) < x / y - round beta fexp1 Zfloor (x / y) <= / 2 * (ulp beta fexp1 (x / y) + ulp beta fexp2 (x / y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z

~ / 2 * ulp beta fexp1 (x / y) < x / y - round beta fexp1 Zfloor (x / y) <= / 2 * (ulp beta fexp1 (x / y) + ulp beta fexp2 (x / y)) -> ~ midp fexp1 (x / y) < x / y <= midp fexp1 (x / y) + / 2 * ulp beta fexp2 (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
H:~ / 2 * ulp beta fexp1 (x / y) < x / y - round beta fexp1 Zfloor (x / y) <= / 2 * (ulp beta fexp1 (x / y) + ulp beta fexp2 (x / y))
H':midp fexp1 (x / y) < x / y <= midp fexp1 (x / y) + / 2 * ulp beta fexp2 (x / y)

/ 2 * ulp beta fexp1 (x / y) < x / y - round beta fexp1 Zfloor (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
H:~ / 2 * ulp beta fexp1 (x / y) < x / y - round beta fexp1 Zfloor (x / y) <= / 2 * (ulp beta fexp1 (x / y) + ulp beta fexp2 (x / y))
H':midp fexp1 (x / y) < x / y <= midp fexp1 (x / y) + / 2 * ulp beta fexp2 (x / y)
x / y - round beta fexp1 Zfloor (x / y) <= / 2 * (ulp beta fexp1 (x / y) + ulp beta fexp2 (x / y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
H:~ / 2 * ulp beta fexp1 (x / y) < x / y - round beta fexp1 Zfloor (x / y) <= / 2 * (ulp beta fexp1 (x / y) + ulp beta fexp2 (x / y))
H':midp fexp1 (x / y) < x / y <= midp fexp1 (x / y) + / 2 * ulp beta fexp2 (x / y)

/ 2 * ulp beta fexp1 (x / y) < x / y - round beta fexp1 Zfloor (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
H:~ / 2 * ulp beta fexp1 (x / y) < x / y - round beta fexp1 Zfloor (x / y) <= / 2 * (ulp beta fexp1 (x / y) + ulp beta fexp2 (x / y))
H':midp fexp1 (x / y) < x / y <= midp fexp1 (x / y) + / 2 * ulp beta fexp2 (x / y)

round beta fexp1 Zfloor (x / y) + / 2 * ulp beta fexp1 (x / y) < round beta fexp1 Zfloor (x / y) + (x / y - round beta fexp1 Zfloor (x / y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
H:~ / 2 * ulp beta fexp1 (x / y) < x / y - round beta fexp1 Zfloor (x / y) <= / 2 * (ulp beta fexp1 (x / y) + ulp beta fexp2 (x / y))
H':midp fexp1 (x / y) < x / y <= midp fexp1 (x / y) + / 2 * ulp beta fexp2 (x / y)

round beta fexp1 Zfloor (x / y) + / 2 * ulp beta fexp1 (x / y) < x / y
apply H'.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
H:~ / 2 * ulp beta fexp1 (x / y) < x / y - round beta fexp1 Zfloor (x / y) <= / 2 * (ulp beta fexp1 (x / y) + ulp beta fexp2 (x / y))
H':midp fexp1 (x / y) < x / y <= midp fexp1 (x / y) + / 2 * ulp beta fexp2 (x / y)

x / y - round beta fexp1 Zfloor (x / y) <= / 2 * (ulp beta fexp1 (x / y) + ulp beta fexp2 (x / y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
H:~ / 2 * ulp beta fexp1 (x / y) < x / y - round beta fexp1 Zfloor (x / y) <= / 2 * (ulp beta fexp1 (x / y) + ulp beta fexp2 (x / y))
H':midp fexp1 (x / y) < x / y <= midp fexp1 (x / y) + / 2 * ulp beta fexp2 (x / y)

round beta fexp1 Zfloor (x / y) + (x / y - round beta fexp1 Zfloor (x / y)) <= round beta fexp1 Zfloor (x / y) + / 2 * (ulp beta fexp1 (x / y) + ulp beta fexp2 (x / y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
H:~ / 2 * ulp beta fexp1 (x / y) < x / y - round beta fexp1 Zfloor (x / y) <= / 2 * (ulp beta fexp1 (x / y) + ulp beta fexp2 (x / y))
H':midp fexp1 (x / y) < x / y <= midp fexp1 (x / y) + / 2 * ulp beta fexp2 (x / y)

x / y <= round beta fexp1 Zfloor (x / y) + / 2 * ulp beta fexp1 (x / y) + / 2 * ulp beta fexp2 (x / y)
apply H'.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z

~ / 2 * ulp beta fexp1 (x / y) < x / y - round beta fexp1 Zfloor (x / y) <= / 2 * (ulp beta fexp1 (x / y) + ulp beta fexp2 (x / y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R

~ / 2 * ulp beta fexp1 (x / y) < x / y - round beta fexp1 Zfloor (x / y) <= / 2 * (ulp beta fexp1 (x / y) + ulp beta fexp2 (x / y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R

~ / 2 * ulp beta fexp1 (x / y) < x / y - round beta fexp1 Zfloor (x / y) <= / 2 * (ulp beta fexp1 (x / y) + ulp beta fexp2 (x / y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R

~ / 2 * ulp beta fexp1 (x / y) < x / y - x' <= / 2 * (ulp beta fexp1 (x / y) + ulp beta fexp2 (x / y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R

x / y <> 0
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
~ / 2 * ulp beta fexp1 (x / y) < x / y - x' <= / 2 * (ulp beta fexp1 (x / y) + ulp beta fexp2 (x / y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R

/ y <> 0
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
~ / 2 * ulp beta fexp1 (x / y) < x / y - x' <= / 2 * (ulp beta fexp1 (x / y) + ulp beta fexp2 (x / y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0

~ / 2 * ulp beta fexp1 (x / y) < x / y - x' <= / 2 * (ulp beta fexp1 (x / y) + ulp beta fexp2 (x / y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0

~ / 2 * bpow (cexp beta fexp1 (x / y)) < x / y - x' <= / 2 * (bpow (cexp beta fexp1 (x / y)) + bpow (cexp beta fexp2 (x / y)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0

generic_format beta fexp1 x -> generic_format beta fexp1 y -> ~ / 2 * bpow (cexp beta fexp1 (x / y)) < x / y - x' <= / 2 * (bpow (cexp beta fexp1 (x / y)) + bpow (cexp beta fexp2 (x / y)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0

x = IZR (Ztrunc (x * bpow (- fexp1 (mag x)))) * bpow (fexp1 (mag x)) -> y = IZR (Ztrunc (y * bpow (- fexp1 (mag y)))) * bpow (fexp1 (mag y)) -> ~ / 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z

x = IZR mx * bpow (fexp1 (mag x)) -> y = IZR (Ztrunc (y * bpow (- fexp1 (mag y)))) * bpow (fexp1 (mag y)) -> ~ / 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z

x = IZR mx * bpow (fexp1 (mag x)) -> y = IZR my * bpow (fexp1 (mag y)) -> ~ / 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))

~ / 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))

False
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))

/ 2 * (u1 + u2) < / 2 * (u1 + u2)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))

/ 2 * (u1 + u2) < x / y - x'
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))

/ 2 * u1 + / 2 * u2 + x' < x / y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))

(/ 2 * u1 + / 2 * u2 + x') * y < x / y * y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))

(/ 2 * u1 + / 2 * u2 + x') * y < x * (/ y * y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))

(/ 2 * u1 + / 2 * u2 + x') * y < x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))

/ 2 * u1 * y + / 2 * u2 * y + x' * y < x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))

2 * (/ 2 * u1 * y + / 2 * u2 * y + x' * y) < 2 * x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))

2 * (/ 2 * u1 * y) + 2 * (/ 2 * u2 * y) + 2 * (x' * y) < 2 * x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))

2 * / 2 * u1 * y + 2 * / 2 * u2 * y + 2 * x' * y < 2 * x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))

u1 * y + u2 * y + 2 * x' * y < 2 * x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

u1 * y + u2 * y + 2 * x' * y < 2 * x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z
u1 * y + u2 * y + 2 * x' * y < 2 * x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

u1 * y + u2 * y + 2 * x' * y < 2 * x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

u1 * y + u2 * y + 2 * x' * y < u1 * y + bpow (fexp1 (mag (x / y)) + fexp1 (mag y)) + 2 * x' * y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z
u1 * y + bpow (fexp1 (mag (x / y)) + fexp1 (mag y)) + 2 * x' * y <= 2 * x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

u1 * y + u2 * y + 2 * x' * y < u1 * y + bpow (fexp1 (mag (x / y)) + fexp1 (mag y)) + 2 * x' * y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

u2 * y < bpow (fexp1 (mag (x / y)) + fexp1 (mag y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

u2 * y < u2 * bpow (mag y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z
u2 * bpow (mag y) <= bpow (fexp1 (mag (x / y)) + fexp1 (mag y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

u2 * y < u2 * bpow (mag y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

u2 * y < u2 * bpow (mag y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

0 < u2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z
y < bpow (mag y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

0 < u2
apply bpow_gt_0.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

y < bpow (mag y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

Rabs y < bpow (mag y)
apply bpow_mag_gt. }
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

u2 * bpow (mag y) <= bpow (fexp1 (mag (x / y)) + fexp1 (mag y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

(fexp2 (mag (x / y)) + mag y <= fexp1 (mag (x / y)) + fexp1 (mag y))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

(fexp2 (mag (x / y)) <= - mag y + fexp1 (mag (x / y)) + fexp1 (mag y))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

(fexp2 (mag (x / y)) <= fexp1 (mag (x / y)) + fexp1 (mag y) + - mag y)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z
Hxy:mag (x / y) = (mag x - mag y + 1)%Z

(fexp2 (mag x - mag y + 1) <= fexp1 (mag x - mag y + 1) + fexp1 (mag y) + - mag y)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z
Hxy:mag (x / y) = (mag x - mag y + 1)%Z

(fexp2 (mag x + 1 - mag y) <= fexp1 (mag x + 1 - mag y) + fexp1 (mag y) + - mag y)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z
Hxy:mag (x / y) = (mag x - mag y + 1)%Z

(fexp1 (mag x + 1) < mag x + 1)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z
Hxy:mag (x / y) = (mag x - mag y + 1)%Z
(fexp1 (mag y) < mag y)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z
Hxy:mag (x / y) = (mag x - mag y + 1)%Z
(fexp1 (mag x + 1 - mag y) <= mag x + 1 - mag y)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z
Hxy:mag (x / y) = (mag x - mag y + 1)%Z

(fexp1 (mag x + 1) < mag x + 1)%Z
now assert (fexp1 (mag x + 1) <= mag x)%Z; [apply valid_exp|omega].
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z
Hxy:mag (x / y) = (mag x - mag y + 1)%Z

(fexp1 (mag y) < mag y)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z
Hxy:mag (x / y) = (mag x - mag y + 1)%Z
(fexp1 (mag x + 1 - mag y) <= mag x + 1 - mag y)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z
Hxy:mag (x / y) = (mag x - mag y + 1)%Z

(fexp1 (mag y) < mag y)%Z
assumption.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z
Hxy:mag (x / y) = (mag x - mag y + 1)%Z

(fexp1 (mag x + 1 - mag y) <= mag x + 1 - mag y)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z
Hxy:mag (x / y) = (mag x - mag y + 1)%Z

(fexp1 (mag x - mag y + 1) <= mag x - mag y + 1)%Z
now rewrite <- Hxy.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

u1 * y + bpow (fexp1 (mag (x / y)) + fexp1 (mag y)) + 2 * x' * y <= 2 * x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

u1 * y + bpow (fexp1 (mag (x / y)) + fexp1 (mag y)) + 2 * x' * y <= 2 * (IZR mx * bpow (fexp1 (mag x)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

u1 * (IZR my * bpow (fexp1 (mag y))) + bpow (fexp1 (mag (x / y)) + fexp1 (mag y)) + 2 * x' * y <= 2 * (IZR mx * bpow (fexp1 (mag x)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

u1 * (IZR my * bpow (fexp1 (mag y))) + bpow (fexp1 (mag (x / y)) + fexp1 (mag y)) + 2 * x' * (IZR my * bpow (fexp1 (mag y))) <= 2 * (IZR mx * bpow (fexp1 (mag x)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

(u1 * (IZR my * bpow (fexp1 (mag y))) + bpow (fexp1 (mag (x / y)) + fexp1 (mag y)) + 2 * x' * (IZR my * bpow (fexp1 (mag y)))) * bpow (- fexp1 (mag (x / y)) - fexp1 (mag y)) <= 2 * (IZR mx * bpow (fexp1 (mag x))) * bpow (- fexp1 (mag (x / y)) - fexp1 (mag y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

u1 * (IZR my * bpow (fexp1 (mag y))) * bpow (- fexp1 (mag (x / y)) - fexp1 (mag y)) + bpow (fexp1 (mag (x / y)) + fexp1 (mag y)) * bpow (- fexp1 (mag (x / y)) - fexp1 (mag y)) + 2 * x' * (IZR my * bpow (fexp1 (mag y))) * bpow (- fexp1 (mag (x / y)) - fexp1 (mag y)) <= 2 * (IZR mx * bpow (fexp1 (mag x))) * bpow (- fexp1 (mag (x / y)) - fexp1 (mag y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

u1 * IZR my * bpow (- fexp1 (mag (x / y))) + 1 + 2 * x' * IZR my * bpow (- fexp1 (mag (x / y))) <= 2 * IZR mx * bpow (fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

IZR my * u1 * bpow (- fexp1 (mag (x / y))) + 1 + 2 * x' * IZR my * bpow (- fexp1 (mag (x / y))) <= 2 * IZR mx * bpow (fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

IZR my + 1 + 2 * x' * IZR my * bpow (- fexp1 (mag (x / y))) <= 2 * IZR mx * bpow (fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

IZR my + 1 + 2 * (x' * IZR my) * bpow (- fexp1 (mag (x / y))) <= 2 * IZR mx * bpow (fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

IZR my + 1 + 2 * (IZR my * x') * bpow (- fexp1 (mag (x / y))) <= 2 * IZR mx * bpow (fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

IZR my + 1 + 2 * (IZR my * x' * bpow (- fexp1 (mag (x / y)))) <= 2 * IZR mx * bpow (fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

IZR my + 1 + 2 * (IZR my * (IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) * bpow (fexp1 (mag (x / y)))) * bpow (- fexp1 (mag (x / y)))) <= 2 * IZR mx * bpow (fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

IZR my + 1 + 2 * (IZR my * IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y)))))) <= 2 * IZR mx * bpow (fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

IZR my + 1 + 2 * (IZR my * IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y)))))) <= 2 * IZR mx * IZR (beta ^ (fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

IZR my + 1 + IZR (2 * (my * Zfloor (x / y * bpow (- fexp1 (mag (x / y)))))) <= IZR (2 * mx * beta ^ (fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

IZR (my + 1 + 2 * (my * Zfloor (x / y * bpow (- fexp1 (mag (x / y)))))) <= IZR (2 * mx * beta ^ (fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

(my + 1 + 2 * (my * Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) <= 2 * mx * beta ^ (fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y)))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

(2 * (my * Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) + my + 1 <= 2 * mx * beta ^ (fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y)))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

(2 * (my * Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) + my < 2 * mx * beta ^ (fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y)))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

IZR (2 * (my * Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) + my) < IZR (2 * mx * beta ^ (fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

IZR (2 * (my * Zfloor (x / y * bpow (- fexp1 (mag (x / y)))))) + IZR my < IZR (2 * mx * beta ^ (fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

2 * (IZR my * IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y)))))) + IZR my < 2 * IZR mx * IZR (beta ^ (fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

2 * (IZR my * IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y)))))) + IZR my < 2 * IZR mx * bpow (fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

(2 * (IZR my * IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y)))))) + IZR my) * bpow (fexp1 (mag y)) < 2 * IZR mx * bpow (fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y)) * bpow (fexp1 (mag y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

2 * (IZR my * IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y)))))) * bpow (fexp1 (mag y)) + IZR my * bpow (fexp1 (mag y)) < 2 * IZR mx * bpow (fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y)) * bpow (fexp1 (mag y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

2 * (IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) * IZR my) * bpow (fexp1 (mag y)) + IZR my * bpow (fexp1 (mag y)) < 2 * IZR mx * bpow (fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y)) * bpow (fexp1 (mag y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

2 * (IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) * (IZR my * bpow (fexp1 (mag y)))) + IZR my * bpow (fexp1 (mag y)) < 2 * IZR mx * bpow (fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y)) * bpow (fexp1 (mag y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

2 * (IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) * y) + y < 2 * IZR mx * bpow (fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y)) * bpow (fexp1 (mag y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

2 * (IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) * y) + y < 2 * IZR mx * bpow (fexp1 (mag x) - fexp1 (mag (x / y)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

2 * (IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) * y) + y < 2 * IZR mx * (bpow (fexp1 (mag x)) * bpow (- fexp1 (mag (x / y))))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

2 * (IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) * y) + y < 2 * (IZR mx * (bpow (fexp1 (mag x)) * bpow (- fexp1 (mag (x / y)))))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

2 * (IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) * y) + y < 2 * (IZR mx * bpow (fexp1 (mag x)) * bpow (- fexp1 (mag (x / y))))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

2 * (IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) * y) + y < 2 * (x * bpow (- fexp1 (mag (x / y))))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

(2 * (IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) * y) + y) * bpow (fexp1 (mag (x / y))) < 2 * (x * bpow (- fexp1 (mag (x / y)))) * bpow (fexp1 (mag (x / y)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

2 * (IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) * y) * bpow (fexp1 (mag (x / y))) + y * bpow (fexp1 (mag (x / y))) < 2 * (x * bpow (- fexp1 (mag (x / y)))) * bpow (fexp1 (mag (x / y)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

2 * (IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) * y) * bpow (fexp1 (mag (x / y))) + y * bpow (fexp1 (mag (x / y))) < 2 * x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

2 * (y * IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y)))))) * bpow (fexp1 (mag (x / y))) + y * bpow (fexp1 (mag (x / y))) < 2 * x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

2 * (y * (IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) * bpow (fexp1 (mag (x / y))))) + y * bpow (fexp1 (mag (x / y))) < 2 * x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

2 * (y * x') + y * bpow (fexp1 (mag (x / y))) < 2 * x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

2 * (y * x') + y * u1 < 2 * x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

/ 2 * (2 * (y * x') + y * u1) < / 2 * (2 * x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

/ 2 * (2 * (y * x')) + / 2 * (y * u1) < / 2 * (2 * x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

/ 2 * 2 * y * x' + / 2 * y * u1 < / 2 * 2 * x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

y * x' + / 2 * y * u1 < x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

y * / 2 * u1 < - y * x' + x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

/ y * (y * / 2 * u1) < / y * (- y * x' + x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

/ y * (y * / 2 * u1) < / y * (- y * x') + / y * x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

/ y * y * / 2 * u1 < / y * - y * x' + / y * x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

/ y * y * / 2 * u1 < - (/ y * y) * x' + / y * x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

/ y * y * / 2 * u1 < - (/ y * y * x') + / y * x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

/ 2 * u1 < - x' + / y * x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(0 <= fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y))%Z

/ 2 * u1 < - x' + x * / y
now rewrite (Rplus_comm (- x')).
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

u1 * y + u2 * y + 2 * x' * y < 2 * x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

u1 * y + u2 * y + 2 * x' * y < 2 * x' * y + u1 * y + bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z
2 * x' * y + u1 * y + bpow (fexp1 (mag x)) <= 2 * x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

u1 * y + u2 * y + 2 * x' * y < 2 * x' * y + u1 * y + bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

u2 * y < bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

u2 * y < u2 * bpow (mag y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z
u2 * bpow (mag y) <= bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

u2 * y < u2 * bpow (mag y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

0 < u2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z
y < bpow (mag y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

y < bpow (mag y)
now apply Rabs_lt_inv; apply bpow_mag_gt.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

u2 * bpow (mag y) <= bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

(fexp2 (mag (x / y)) + mag y <= fexp1 (mag x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

(fexp2 (mag (x / y)) <= - mag y + fexp1 (mag x))%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

(fexp2 (mag (x / y)) <= fexp1 (mag x) + - mag y)%Z
destruct (mag_div_disj x y Px Py) as [Hxy|Hxy]; rewrite Hxy; apply Hexp; try assumption; rewrite <- Hxy; omega.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

2 * x' * y + u1 * y + bpow (fexp1 (mag x)) <= 2 * x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

2 * x' * y + u1 * y + bpow (fexp1 (mag x)) <= 2 * (IZR mx * bpow (fexp1 (mag x)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

2 * x' * (IZR my * bpow (fexp1 (mag y))) + u1 * (IZR my * bpow (fexp1 (mag y))) + bpow (fexp1 (mag x)) <= 2 * (IZR mx * bpow (fexp1 (mag x)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

(2 * x' * (IZR my * bpow (fexp1 (mag y))) + u1 * (IZR my * bpow (fexp1 (mag y))) + bpow (fexp1 (mag x))) * bpow (- fexp1 (mag x)) <= 2 * (IZR mx * bpow (fexp1 (mag x))) * bpow (- fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

2 * x' * (IZR my * bpow (fexp1 (mag y))) * bpow (- fexp1 (mag x)) + u1 * (IZR my * bpow (fexp1 (mag y))) * bpow (- fexp1 (mag x)) + bpow (fexp1 (mag x)) * bpow (- fexp1 (mag x)) <= 2 * (IZR mx * bpow (fexp1 (mag x))) * bpow (- fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

2 * x' * IZR my * bpow (fexp1 (mag y) - fexp1 (mag x)) + u1 * IZR my * bpow (fexp1 (mag y) - fexp1 (mag x)) + 1 <= 2 * IZR mx
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

2 * IZR my * x' * bpow (fexp1 (mag y) - fexp1 (mag x)) + u1 * IZR my * bpow (fexp1 (mag y) - fexp1 (mag x)) + 1 <= 2 * IZR mx
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

2 * IZR my * x' * bpow (fexp1 (mag y) - fexp1 (mag x)) + IZR my * u1 * bpow (fexp1 (mag y) - fexp1 (mag x)) + 1 <= 2 * IZR mx
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

2 * IZR my * (IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) * bpow (fexp1 (mag (x / y)))) * bpow (fexp1 (mag y) - fexp1 (mag x)) + IZR my * bpow (fexp1 (mag (x / y))) * bpow (fexp1 (mag y) - fexp1 (mag x)) + 1 <= 2 * IZR mx
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

2 * IZR my * IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) * bpow (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x)) + IZR my * bpow (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x)) + 1 <= 2 * IZR mx
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

2 * IZR my * IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) * IZR (beta ^ (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x))) + IZR my * IZR (beta ^ (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x))) + 1 <= 2 * IZR mx
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

IZR (2 * my * Zfloor (x / y * bpow (- fexp1 (mag (x / y)))) * beta ^ (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x))) + IZR (my * beta ^ (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x))) + 1 <= IZR (2 * mx)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

IZR (2 * my * Zfloor (x / y * bpow (- fexp1 (mag (x / y)))) * beta ^ (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x)) + my * beta ^ (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x)) + 1) <= IZR (2 * mx)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

(2 * my * Zfloor (x / y * bpow (- fexp1 (mag (x / y)))) * beta ^ (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x)) + my * beta ^ (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x)) + 1 <= 2 * mx)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

(2 * my * Zfloor (x / y * bpow (- fexp1 (mag (x / y)))) * beta ^ (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x)) + my * beta ^ (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x)) < 2 * mx)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

IZR (2 * my * Zfloor (x / y * bpow (- fexp1 (mag (x / y)))) * beta ^ (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x)) + my * beta ^ (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x))) < IZR (2 * mx)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

IZR (2 * my * Zfloor (x / y * bpow (- fexp1 (mag (x / y)))) * beta ^ (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x))) + IZR (my * beta ^ (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x))) < IZR (2 * mx)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

2 * IZR my * IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) * IZR (beta ^ (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x))) + IZR my * IZR (beta ^ (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x))) < 2 * IZR mx
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

2 * IZR my * IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) * bpow (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x)) + IZR my * bpow (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x)) < 2 * IZR mx
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

(2 * IZR my * IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) * bpow (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x)) + IZR my * bpow (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x))) * bpow (fexp1 (mag x)) < 2 * IZR mx * bpow (fexp1 (mag x))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

(2 * IZR my * IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) * bpow (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x)) + IZR my * bpow (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x))) * bpow (fexp1 (mag x)) < 2 * (IZR mx * bpow (fexp1 (mag x)))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

(2 * IZR my * IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) * bpow (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x)) + IZR my * bpow (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x))) * bpow (fexp1 (mag x)) < 2 * x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

2 * IZR my * IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) * bpow (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x)) * bpow (fexp1 (mag x)) + IZR my * bpow (fexp1 (mag (x / y)) + fexp1 (mag y) - fexp1 (mag x)) * bpow (fexp1 (mag x)) < 2 * x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

2 * IZR my * IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) * bpow (fexp1 (mag (x / y)) + fexp1 (mag y)) + IZR my * bpow (fexp1 (mag (x / y)) + fexp1 (mag y)) < 2 * x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

2 * IZR my * IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) * (bpow (fexp1 (mag (x / y))) * bpow (fexp1 (mag y))) + IZR my * (bpow (fexp1 (mag (x / y))) * bpow (fexp1 (mag y))) < 2 * x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

2 * IZR my * (IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) * (bpow (fexp1 (mag (x / y))) * bpow (fexp1 (mag y)))) + IZR my * (bpow (fexp1 (mag (x / y))) * bpow (fexp1 (mag y))) < 2 * x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

2 * IZR my * (IZR (Zfloor (x / y * bpow (- fexp1 (mag (x / y))))) * bpow (fexp1 (mag (x / y))) * bpow (fexp1 (mag y))) + IZR my * (bpow (fexp1 (mag (x / y))) * bpow (fexp1 (mag y))) < 2 * x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

2 * IZR my * (x' * bpow (fexp1 (mag y))) + IZR my * (bpow (fexp1 (mag (x / y))) * bpow (fexp1 (mag y))) < 2 * x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

2 * IZR my * (bpow (fexp1 (mag y)) * x') + IZR my * (bpow (fexp1 (mag y)) * bpow (fexp1 (mag (x / y)))) < 2 * x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

2 * (IZR my * (bpow (fexp1 (mag y)) * x')) + IZR my * (bpow (fexp1 (mag y)) * bpow (fexp1 (mag (x / y)))) < 2 * x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

2 * (IZR my * bpow (fexp1 (mag y)) * x') + IZR my * bpow (fexp1 (mag y)) * bpow (fexp1 (mag (x / y))) < 2 * x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

2 * (y * x') + y * bpow (fexp1 (mag (x / y))) < 2 * x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

2 * (y * x') + y * u1 < 2 * x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

/ 2 * (2 * (y * x') + y * u1) < / 2 * (2 * x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

/ 2 * (2 * (y * x')) + / 2 * (y * u1) < / 2 * (2 * x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

/ 2 * 2 * y * x' + / 2 * y * u1 < / 2 * 2 * x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

y * x' + / 2 * y * u1 < x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

y * / 2 * u1 < - y * x' + x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

/ y * (y * / 2 * u1) < / y * (- y * x' + x)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

/ y * (y * / 2 * u1) < / y * (- y * x') + / y * x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

/ y * y * / 2 * u1 < / y * - y * x' + / y * x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

/ y * y * / 2 * u1 < - (/ y * y) * x' + / y * x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

/ y * y * / 2 * u1 < - (/ y * y * x') + / y * x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

/ 2 * u1 < - x' + / y * x
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hfx:(fexp1 (mag x) < mag x)%Z
Hfy:(fexp1 (mag y) < mag y)%Z
u1:=bpow (fexp1 (mag (x / y))):R
u2:=bpow (fexp2 (mag (x / y))):R
x':=round beta fexp1 Zfloor (x / y):R
S:x / y <> 0
mx:=Ztrunc (x * bpow (- fexp1 (mag x))):Z
my:=Ztrunc (y * bpow (- fexp1 (mag y))):Z
Fx:x = IZR mx * bpow (fexp1 (mag x))
Fy:y = IZR my * bpow (fexp1 (mag y))
Hlr:/ 2 * bpow (fexp1 (mag (x / y))) < x / y - x' <= / 2 * (bpow (fexp1 (mag (x / y))) + bpow (fexp2 (mag (x / y))))
He:(fexp1 (mag x) - fexp1 (mag (x / y)) - fexp1 (mag y) < 0)%Z

/ 2 * u1 < - x' + x * / y
now rewrite (Rplus_comm (- x')). Qed.
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall choice1 choice2 : Z -> bool, (exists n : Z, beta = (2 * n)%Z) -> round_round_div_hyp fexp1 fexp2 -> forall x y : R, 0 < x -> 0 < y -> generic_format beta fexp1 x -> generic_format beta fexp1 y -> round_round_eq fexp1 fexp2 choice1 choice2 (x / y)
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall choice1 choice2 : Z -> bool, (exists n : Z, beta = (2 * n)%Z) -> round_round_div_hyp fexp1 fexp2 -> forall x y : R, 0 < x -> 0 < y -> generic_format beta fexp1 x -> generic_format beta fexp1 y -> round_round_eq fexp1 fexp2 choice1 choice2 (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y

round_round_eq fexp1 fexp2 choice1 choice2 (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y

0 < x / y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Pxy:0 < x / y
round_round_eq fexp1 fexp2 choice1 choice2 (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y

0 < x / y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y

0 < / y
now apply Rinv_0_lt_compat.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Pxy:0 < x / y

round_round_eq fexp1 fexp2 choice1 choice2 (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Pxy:0 < x / y

Valid_exp fexp1
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Pxy:0 < x / y
Valid_exp fexp2
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Pxy:0 < x / y
0 < x / y
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Pxy:0 < x / y
(fexp2 (mag (x / y)) <= fexp1 (mag (x / y)) - 1)%Z
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Pxy:0 < x / y
fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z -> bpow (mag (x / y)) - / 2 * ulp beta fexp2 (x / y) <= x / y -> round_round_eq fexp1 fexp2 choice1 choice2 (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Pxy:0 < x / y
(fexp1 (mag (x / y)) <= mag (x / y))%Z -> midp fexp1 (x / y) - / 2 * ulp beta fexp2 (x / y) <= x / y < midp fexp1 (x / y) -> round_round_eq fexp1 fexp2 choice1 choice2 (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Pxy:0 < x / y
(fexp1 (mag (x / y)) <= mag (x / y))%Z -> x / y = midp fexp1 (x / y) -> round_round_eq fexp1 fexp2 choice1 choice2 (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Pxy:0 < x / y
(fexp1 (mag (x / y)) <= mag (x / y))%Z -> midp fexp1 (x / y) < x / y <= midp fexp1 (x / y) + / 2 * ulp beta fexp2 (x / y) -> round_round_eq fexp1 fexp2 choice1 choice2 (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Pxy:0 < x / y

Valid_exp fexp1
exact Vfexp1.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Pxy:0 < x / y

Valid_exp fexp2
exact Vfexp2.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Pxy:0 < x / y

0 < x / y
exact Pxy.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Pxy:0 < x / y

(fexp2 (mag (x / y)) <= fexp1 (mag (x / y)) - 1)%Z
apply Hexp.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Pxy:0 < x / y

fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z -> bpow (mag (x / y)) - / 2 * ulp beta fexp2 (x / y) <= x / y -> round_round_eq fexp1 fexp2 choice1 choice2 (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Pxy:0 < x / y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hlxy:bpow (mag (x / y)) - / 2 * ulp beta fexp2 (x / y) <= x / y

round_round_eq fexp1 fexp2 choice1 choice2 (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Pxy:0 < x / y
Hf1:fexp1 (mag (x / y)) = (mag (x / y) + 1)%Z
Hlxy:bpow (mag (x / y)) - / 2 * ulp beta fexp2 (x / y) <= x / y

False
now apply (round_round_div_aux0 fexp1 fexp2 _ _ choice1 choice2 Hexp x y).
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Pxy:0 < x / y

(fexp1 (mag (x / y)) <= mag (x / y))%Z -> midp fexp1 (x / y) - / 2 * ulp beta fexp2 (x / y) <= x / y < midp fexp1 (x / y) -> round_round_eq fexp1 fexp2 choice1 choice2 (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Pxy:0 < x / y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hlxy:midp fexp1 (x / y) - / 2 * ulp beta fexp2 (x / y) <= x / y < midp fexp1 (x / y)

round_round_eq fexp1 fexp2 choice1 choice2 (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Pxy:0 < x / y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hlxy:midp fexp1 (x / y) - / 2 * ulp beta fexp2 (x / y) <= x / y < midp fexp1 (x / y)

False
now apply (round_round_div_aux1 fexp1 fexp2 _ _ choice1 choice2 Hexp x y).
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Pxy:0 < x / y

(fexp1 (mag (x / y)) <= mag (x / y))%Z -> x / y = midp fexp1 (x / y) -> round_round_eq fexp1 fexp2 choice1 choice2 (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Pxy:0 < x / y
H:(fexp1 (mag (x / y)) <= mag (x / y))%Z

x / y = midp fexp1 (x / y) -> round_round_eq fexp1 fexp2 choice1 choice2 (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Pxy:0 < x / y
H:(fexp1 (mag (x / y)) <= mag (x / y))%Z

(fexp2 (mag (x / y)) <= fexp1 (mag (x / y)) - 1)%Z
apply Hexp.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Pxy:0 < x / y

(fexp1 (mag (x / y)) <= mag (x / y))%Z -> midp fexp1 (x / y) < x / y <= midp fexp1 (x / y) + / 2 * ulp beta fexp2 (x / y) -> round_round_eq fexp1 fexp2 choice1 choice2 (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Pxy:0 < x / y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hlxy:midp fexp1 (x / y) < x / y <= midp fexp1 (x / y) + / 2 * ulp beta fexp2 (x / y)

round_round_eq fexp1 fexp2 choice1 choice2 (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Px:0 < x
Py:0 < y
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Pxy:0 < x / y
Hf1:(fexp1 (mag (x / y)) <= mag (x / y))%Z
Hlxy:midp fexp1 (x / y) < x / y <= midp fexp1 (x / y) + / 2 * ulp beta fexp2 (x / y)

False
now apply (round_round_div_aux2 fexp1 fexp2 _ _ choice1 choice2 Hexp x y). Qed.
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall choice1 choice2 : Z -> bool, (exists n : Z, beta = (2 * n)%Z) -> round_round_div_hyp fexp1 fexp2 -> forall x y : R, y <> 0 -> generic_format beta fexp1 x -> generic_format beta fexp1 y -> round_round_eq fexp1 fexp2 choice1 choice2 (x / y)
beta:radix

forall fexp1 fexp2 : Z -> Z, Valid_exp fexp1 -> Valid_exp fexp2 -> forall choice1 choice2 : Z -> bool, (exists n : Z, beta = (2 * n)%Z) -> round_round_div_hyp fexp1 fexp2 -> forall x y : R, y <> 0 -> generic_format beta fexp1 x -> generic_format beta fexp1 y -> round_round_eq fexp1 fexp2 choice1 choice2 (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Nzy:y <> 0
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y

round_round_eq fexp1 fexp2 choice1 choice2 (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Nzy:y <> 0
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x / y)) = round beta fexp1 (Znearest choice1) (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Nzy:y <> 0
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nx:x < 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x / y)) = round beta fexp1 (Znearest choice1) (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Nzy:y <> 0
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0
round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x / y)) = round beta fexp1 (Znearest choice1) (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Nzy:y <> 0
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:x > 0
round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x / y)) = round beta fexp1 (Znearest choice1) (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Nzy:y <> 0
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nx:x < 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x / y)) = round beta fexp1 (Znearest choice1) (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Nzy:y <> 0
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nx:x < 0
Ny:y < 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x / y)) = round beta fexp1 (Znearest choice1) (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Nzy:y <> 0
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nx:x < 0
Zy:y = 0
round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x / y)) = round beta fexp1 (Znearest choice1) (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Nzy:y <> 0
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nx:x < 0
Py:y > 0
round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x / y)) = round beta fexp1 (Znearest choice1) (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Nzy:y <> 0
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nx:x < 0
Ny:y < 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x / y)) = round beta fexp1 (Znearest choice1) (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Nzy:y <> 0
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nx:x < 0
Ny:y < 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (- - x / y)) = round beta fexp1 (Znearest choice1) (- - x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Nzy:y <> 0
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nx:x < 0
Ny:y < 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (- - x / - - y)) = round beta fexp1 (Znearest choice1) (- - x / - - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Nzy:y <> 0
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nx:x < 0
Ny:y < 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (- (- x / - - y))) = round beta fexp1 (Znearest choice1) (- (- x / - - y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Nzy:y <> 0
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nx:x < 0
Ny:y < 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (- (- x * - / - y))) = round beta fexp1 (Znearest choice1) (- (- x * - / - y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Nzy:y <> 0
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nx:x < 0
Ny:y < 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (- - (- x * / - y))) = round beta fexp1 (Znearest choice1) (- - (- x * / - y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Nzy:y <> 0
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nx:x < 0
Ny:y < 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (- x * / - y)) = round beta fexp1 (Znearest choice1) (- x * / - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Nzy:y <> 0
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nx:x < 0
Ny:y < 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (- x / - y)) = round beta fexp1 (Znearest choice1) (- x / - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Nzy:y <> 0
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nx:- 0 < - x
Ny:y < 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (- x / - y)) = round beta fexp1 (Znearest choice1) (- x / - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Nzy:y <> 0
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nx:- 0 < - x
Ny:- 0 < - y

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (- x / - y)) = round beta fexp1 (Znearest choice1) (- x / - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Nzy:y <> 0
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nx:0 < - x
Ny:0 < - y

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (- x / - y)) = round beta fexp1 (Znearest choice1) (- x / - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Nzy:y <> 0
Fx:generic_format beta fexp1 (- x)
Fy:generic_format beta fexp1 y
Nx:0 < - x
Ny:0 < - y

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (- x / - y)) = round beta fexp1 (Znearest choice1) (- x / - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Nzy:y <> 0
Fx:generic_format beta fexp1 (- x)
Fy:generic_format beta fexp1 (- y)
Nx:0 < - x
Ny:0 < - y

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (- x / - y)) = round beta fexp1 (Znearest choice1) (- x / - y)
now apply round_round_div_aux.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Nzy:y <> 0
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nx:x < 0
Zy:y = 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x / y)) = round beta fexp1 (Znearest choice1) (x / y)
now casetype False; apply Nzy.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Nzy:y <> 0
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nx:x < 0
Py:y > 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x / y)) = round beta fexp1 (Znearest choice1) (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Nzy:y <> 0
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nx:x < 0
Py:y > 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (- - x / y)) = round beta fexp1 (Znearest choice1) (- - x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Nzy:y <> 0
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nx:x < 0
Py:y > 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (- (- x / y))) = round beta fexp1 (Znearest choice1) (- (- x / y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Nzy:y <> 0
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nx:x < 0
Py:y > 0

- round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) (round beta fexp2 (Znearest (fun t : Z => negb (choice2 (- (t + 1))%Z))) (- x / y)) = - round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) (- x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Nzy:y <> 0
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nx:x < 0
Py:y > 0

round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) (round beta fexp2 (Znearest (fun t : Z => negb (choice2 (- (t + 1))%Z))) (- x / y)) = round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) (- x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Nzy:y <> 0
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nx:- 0 < - x
Py:y > 0

round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) (round beta fexp2 (Znearest (fun t : Z => negb (choice2 (- (t + 1))%Z))) (- x / y)) = round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) (- x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Nzy:y <> 0
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Nx:0 < - x
Py:y > 0

round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) (round beta fexp2 (Znearest (fun t : Z => negb (choice2 (- (t + 1))%Z))) (- x / y)) = round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) (- x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Nzy:y <> 0
Fx:generic_format beta fexp1 (- x)
Fy:generic_format beta fexp1 y
Nx:0 < - x
Py:y > 0

round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) (round beta fexp2 (Znearest (fun t : Z => negb (choice2 (- (t + 1))%Z))) (- x / y)) = round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) (- x / y)
now apply round_round_div_aux.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Nzy:y <> 0
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x / y)) = round beta fexp1 (Znearest choice1) (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Nzy:y <> 0
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (0 / y)) = round beta fexp1 (Znearest choice1) (0 / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Nzy:y <> 0
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Zx:x = 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) 0) = round beta fexp1 (Znearest choice1) 0
now rewrite round_0; [|apply valid_rnd_N].
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Nzy:y <> 0
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:x > 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x / y)) = round beta fexp1 (Znearest choice1) (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Nzy:y <> 0
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:x > 0
Ny:y < 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x / y)) = round beta fexp1 (Znearest choice1) (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Nzy:y <> 0
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:x > 0
Zy:y = 0
round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x / y)) = round beta fexp1 (Znearest choice1) (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Nzy:y <> 0
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:x > 0
Py:y > 0
round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x / y)) = round beta fexp1 (Znearest choice1) (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Nzy:y <> 0
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:x > 0
Ny:y < 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x / y)) = round beta fexp1 (Znearest choice1) (x / y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Nzy:y <> 0
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:x > 0
Ny:y < 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x / - - y)) = round beta fexp1 (Znearest choice1) (x / - - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Nzy:y <> 0
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:x > 0
Ny:y < 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x * - / - y)) = round beta fexp1 (Znearest choice1) (x * - / - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Nzy:y <> 0
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:x > 0
Ny:y < 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (- (x * / - y))) = round beta fexp1 (Znearest choice1) (- (x * / - y))
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Nzy:y <> 0
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:x > 0
Ny:y < 0

- round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) (round beta fexp2 (Znearest (fun t : Z => negb (choice2 (- (t + 1))%Z))) (x * / - y)) = - round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) (x * / - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Nzy:y <> 0
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:x > 0
Ny:y < 0

round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) (round beta fexp2 (Znearest (fun t : Z => negb (choice2 (- (t + 1))%Z))) (x * / - y)) = round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) (x * / - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Nzy:y <> 0
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:x > 0
Ny:- 0 < - y

round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) (round beta fexp2 (Znearest (fun t : Z => negb (choice2 (- (t + 1))%Z))) (x * / - y)) = round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) (x * / - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Nzy:y <> 0
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:x > 0
Ny:0 < - y

round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) (round beta fexp2 (Znearest (fun t : Z => negb (choice2 (- (t + 1))%Z))) (x * / - y)) = round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) (x * / - y)
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Nzy:y <> 0
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 (- y)
Px:x > 0
Ny:0 < - y

round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) (round beta fexp2 (Znearest (fun t : Z => negb (choice2 (- (t + 1))%Z))) (x * / - y)) = round beta fexp1 (Znearest (fun t : Z => negb (choice1 (- (t + 1))%Z))) (x * / - y)
now apply round_round_div_aux.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Nzy:y <> 0
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:x > 0
Zy:y = 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x / y)) = round beta fexp1 (Znearest choice1) (x / y)
now casetype False; apply Nzy.
beta:radix
fexp1, fexp2:Z -> Z
Vfexp1:Valid_exp fexp1
Vfexp2:Valid_exp fexp2
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hexp:round_round_div_hyp fexp1 fexp2
x, y:R
Nzy:y <> 0
Fx:generic_format beta fexp1 x
Fy:generic_format beta fexp1 y
Px:x > 0
Py:y > 0

round beta fexp1 (Znearest choice1) (round beta fexp2 (Znearest choice2) (x / y)) = round beta fexp1 (Znearest choice1) (x / y)
now apply round_round_div_aux. Qed. Section Double_round_div_FLX. Variable prec : Z. Variable prec' : Z. Context { prec_gt_0_ : Prec_gt_0 prec }. Context { prec_gt_0_' : Prec_gt_0 prec' }.
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(2 * prec <= prec')%Z -> round_round_div_hyp (FLX_exp prec) (FLX_exp prec')
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(2 * prec <= prec')%Z -> round_round_div_hyp (FLX_exp prec) (FLX_exp prec')
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hprec:(2 * prec <= prec')%Z

round_round_div_hyp (FLX_exp prec) (FLX_exp prec')
beta:radix
prec, prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hprec:(2 * prec <= prec')%Z

round_round_div_hyp (FLX_exp prec) (FLX_exp prec')
beta:radix
prec, prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hprec:(2 * prec <= prec')%Z

round_round_div_hyp (fun e : Z => (e - prec)%Z) (fun e : Z => (e - prec')%Z)
beta:radix
prec, prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hprec:(2 * prec <= prec')%Z

(forall ex : Z, (ex - prec' <= ex - prec - 1)%Z) /\ (forall ex ey : Z, (ex - prec < ex)%Z -> (ey - prec < ey)%Z -> (ex - ey - prec <= ex - ey + 1)%Z -> (ex - ey - prec' <= ex - prec - ey)%Z) /\ (forall ex ey : Z, (ex - prec < ex)%Z -> (ey - prec < ey)%Z -> (ex - ey + 1 - prec <= ex - ey + 1 + 1)%Z -> (ex - ey + 1 - prec' <= ex - prec - ey)%Z) /\ (forall ex ey : Z, (ex - prec < ex)%Z -> (ey - prec < ey)%Z -> (ex - ey - prec <= ex - ey)%Z -> (ex - ey - prec' <= ex - ey - prec + (ey - prec) - ey)%Z) /\ (forall ex ey : Z, (ex - prec < ex)%Z -> (ey - prec < ey)%Z -> (ex - ey - prec)%Z = (ex - ey + 1)%Z -> (ex - ey - prec' <= ex - ey - ey + (ey - prec))%Z)
beta:radix
prec, prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hprec:(2 * prec <= prec')%Z

(forall ex ey : Z, (ex - prec < ex)%Z -> (ey - prec < ey)%Z -> (ex - ey - prec <= ex - ey + 1)%Z -> (ex - ey - prec' <= ex - prec - ey)%Z) /\ (forall ex ey : Z, (ex - prec < ex)%Z -> (ey - prec < ey)%Z -> (ex - ey + 1 - prec <= ex - ey + 1 + 1)%Z -> (ex - ey + 1 - prec' <= ex - prec - ey)%Z) /\ (forall ex ey : Z, (ex - prec < ex)%Z -> (ey - prec < ey)%Z -> (ex - ey - prec <= ex - ey)%Z -> (ex - ey - prec' <= ex - ey - prec + (ey - prec) - ey)%Z) /\ (forall ex ey : Z, (ex - prec < ex)%Z -> (ey - prec < ey)%Z -> (ex - ey - prec)%Z = (ex - ey + 1)%Z -> (ex - ey - prec' <= ex - ey - ey + (ey - prec))%Z)
split; [|split; [|split]]; intros ex ey; omega. Qed.
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

forall choice1 choice2 : Z -> bool, (exists n : Z, beta = (2 * n)%Z) -> (2 * prec <= prec')%Z -> forall x y : R, y <> 0 -> FLX_format beta prec x -> FLX_format beta prec y -> round_round_eq (FLX_exp prec) (FLX_exp prec') choice1 choice2 (x / y)
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

forall choice1 choice2 : Z -> bool, (exists n : Z, beta = (2 * n)%Z) -> (2 * prec <= prec')%Z -> forall x y : R, y <> 0 -> FLX_format beta prec x -> FLX_format beta prec y -> round_round_eq (FLX_exp prec) (FLX_exp prec') choice1 choice2 (x / y)
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Nzy:y <> 0
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y

round_round_eq (FLX_exp prec) (FLX_exp prec') choice1 choice2 (x / y)
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Nzy:y <> 0
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y

Valid_exp (FLX_exp prec)
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Nzy:y <> 0
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y
Valid_exp (FLX_exp prec')
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Nzy:y <> 0
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y
exists n : Z, beta = (2 * n)%Z
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Nzy:y <> 0
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y
round_round_div_hyp (FLX_exp prec) (FLX_exp prec')
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Nzy:y <> 0
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y
y <> 0
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Nzy:y <> 0
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y
generic_format beta (FLX_exp prec) x
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Nzy:y <> 0
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y
generic_format beta (FLX_exp prec) y
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Nzy:y <> 0
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y

Valid_exp (FLX_exp prec)
now apply FLX_exp_valid.
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Nzy:y <> 0
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y

Valid_exp (FLX_exp prec')
now apply FLX_exp_valid.
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Nzy:y <> 0
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y

exists n : Z, beta = (2 * n)%Z
exact Ebeta.
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Nzy:y <> 0
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y

round_round_div_hyp (FLX_exp prec) (FLX_exp prec')
now apply FLX_round_round_div_hyp.
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Nzy:y <> 0
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y

y <> 0
exact Nzy.
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Nzy:y <> 0
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y

generic_format beta (FLX_exp prec) x
now apply generic_format_FLX.
beta:radix
prec, prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Nzy:y <> 0
Fx:FLX_format beta prec x
Fy:FLX_format beta prec y

generic_format beta (FLX_exp prec) y
now apply generic_format_FLX. Qed. End Double_round_div_FLX. Section Double_round_div_FLT. Variable emin prec : Z. Variable emin' prec' : Z. Context { prec_gt_0_ : Prec_gt_0 prec }. Context { prec_gt_0_' : Prec_gt_0 prec' }.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(emin' <= emin - prec - 2)%Z -> (2 * prec <= prec')%Z -> round_round_div_hyp (FLT_exp emin prec) (FLT_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(emin' <= emin - prec - 2)%Z -> (2 * prec <= prec')%Z -> round_round_div_hyp (FLT_exp emin prec) (FLT_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin - prec - 2)%Z
Hprec:(2 * prec <= prec')%Z

round_round_div_hyp (FLT_exp emin prec) (FLT_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin - prec - 2)%Z
Hprec:(2 * prec <= prec')%Z

round_round_div_hyp (fun e : Z => Z.max (e - prec) emin) (fun e : Z => Z.max (e - prec') emin')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin - prec - 2)%Z
Hprec:(2 * prec <= prec')%Z

round_round_div_hyp (fun e : Z => Z.max (e - prec) emin) (fun e : Z => Z.max (e - prec') emin')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin - prec - 2)%Z
Hprec:(2 * prec <= prec')%Z

(forall ex : Z, (Z.max (ex - prec') emin' <= Z.max (ex - prec) emin - 1)%Z) /\ (forall ex ey : Z, (Z.max (ex - prec) emin < ex)%Z -> (Z.max (ey - prec) emin < ey)%Z -> (Z.max (ex - ey - prec) emin <= ex - ey + 1)%Z -> (Z.max (ex - ey - prec') emin' <= Z.max (ex - prec) emin - ey)%Z) /\ (forall ex ey : Z, (Z.max (ex - prec) emin < ex)%Z -> (Z.max (ey - prec) emin < ey)%Z -> (Z.max (ex - ey + 1 - prec) emin <= ex - ey + 1 + 1)%Z -> (Z.max (ex - ey + 1 - prec') emin' <= Z.max (ex - prec) emin - ey)%Z) /\ (forall ex ey : Z, (Z.max (ex - prec) emin < ex)%Z -> (Z.max (ey - prec) emin < ey)%Z -> (Z.max (ex - ey - prec) emin <= ex - ey)%Z -> (Z.max (ex - ey - prec') emin' <= Z.max (ex - ey - prec) emin + Z.max (ey - prec) emin - ey)%Z) /\ (forall ex ey : Z, (Z.max (ex - prec) emin < ex)%Z -> (Z.max (ey - prec) emin < ey)%Z -> Z.max (ex - ey - prec) emin = (ex - ey + 1)%Z -> (Z.max (ex - ey - prec') emin' <= ex - ey - ey + Z.max (ey - prec) emin)%Z)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin - prec - 2)%Z
Hprec:(2 * prec <= prec')%Z
ex:Z

(Z.max (ex - prec') emin' <= Z.max (ex - prec) emin - 1)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin - prec - 2)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z
(Z.max (ex - prec) emin < ex)%Z -> (Z.max (ey - prec) emin < ey)%Z -> (Z.max (ex - ey - prec) emin <= ex - ey + 1)%Z -> (Z.max (ex - ey - prec') emin' <= Z.max (ex - prec) emin - ey)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin - prec - 2)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z
(Z.max (ex - prec) emin < ex)%Z -> (Z.max (ey - prec) emin < ey)%Z -> (Z.max (ex - ey + 1 - prec) emin <= ex - ey + 1 + 1)%Z -> (Z.max (ex - ey + 1 - prec') emin' <= Z.max (ex - prec) emin - ey)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin - prec - 2)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z
(Z.max (ex - prec) emin < ex)%Z -> (Z.max (ey - prec) emin < ey)%Z -> (Z.max (ex - ey - prec) emin <= ex - ey)%Z -> (Z.max (ex - ey - prec') emin' <= Z.max (ex - ey - prec) emin + Z.max (ey - prec) emin - ey)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin - prec - 2)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z
(Z.max (ex - prec) emin < ex)%Z -> (Z.max (ey - prec) emin < ey)%Z -> Z.max (ex - ey - prec) emin = (ex - ey + 1)%Z -> (Z.max (ex - ey - prec') emin' <= ex - ey - ey + Z.max (ey - prec) emin)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin - prec - 2)%Z
Hprec:(2 * prec <= prec')%Z
ex:Z

(Z.max (ex - prec') emin' <= Z.max (ex - prec) emin - 1)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin - prec - 2)%Z
Hprec:(2 * prec <= prec')%Z
ex:Z

(ex - prec' >= emin')%Z /\ Z.max (ex - prec') emin' = (ex - prec')%Z \/ (ex - prec' < emin')%Z /\ Z.max (ex - prec') emin' = emin' -> (Z.max (ex - prec') emin' <= Z.max (ex - prec) emin - 1)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin - prec - 2)%Z
Hprec:(2 * prec <= prec')%Z
ex:Z

(ex - prec >= emin)%Z /\ Z.max (ex - prec) emin = (ex - prec)%Z \/ (ex - prec < emin)%Z /\ Z.max (ex - prec) emin = emin -> (ex - prec' >= emin')%Z /\ Z.max (ex - prec') emin' = (ex - prec')%Z \/ (ex - prec' < emin')%Z /\ Z.max (ex - prec') emin' = emin' -> (Z.max (ex - prec') emin' <= Z.max (ex - prec) emin - 1)%Z
omega.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin - prec - 2)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z

(Z.max (ex - prec) emin < ex)%Z -> (Z.max (ey - prec) emin < ey)%Z -> (Z.max (ex - ey - prec) emin <= ex - ey + 1)%Z -> (Z.max (ex - ey - prec') emin' <= Z.max (ex - prec) emin - ey)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin - prec - 2)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z

(ex - prec >= emin)%Z /\ Z.max (ex - prec) emin = (ex - prec)%Z \/ (ex - prec < emin)%Z /\ Z.max (ex - prec) emin = emin -> (Z.max (ex - prec) emin < ex)%Z -> (Z.max (ey - prec) emin < ey)%Z -> (Z.max (ex - ey - prec) emin <= ex - ey + 1)%Z -> (Z.max (ex - ey - prec') emin' <= Z.max (ex - prec) emin - ey)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin - prec - 2)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z

(ey - prec >= emin)%Z /\ Z.max (ey - prec) emin = (ey - prec)%Z \/ (ey - prec < emin)%Z /\ Z.max (ey - prec) emin = emin -> (ex - prec >= emin)%Z /\ Z.max (ex - prec) emin = (ex - prec)%Z \/ (ex - prec < emin)%Z /\ Z.max (ex - prec) emin = emin -> (Z.max (ex - prec) emin < ex)%Z -> (Z.max (ey - prec) emin < ey)%Z -> (Z.max (ex - ey - prec) emin <= ex - ey + 1)%Z -> (Z.max (ex - ey - prec') emin' <= Z.max (ex - prec) emin - ey)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin - prec - 2)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z

(ex - ey - prec >= emin)%Z /\ Z.max (ex - ey - prec) emin = (ex - ey - prec)%Z \/ (ex - ey - prec < emin)%Z /\ Z.max (ex - ey - prec) emin = emin -> (ey - prec >= emin)%Z /\ Z.max (ey - prec) emin = (ey - prec)%Z \/ (ey - prec < emin)%Z /\ Z.max (ey - prec) emin = emin -> (ex - prec >= emin)%Z /\ Z.max (ex - prec) emin = (ex - prec)%Z \/ (ex - prec < emin)%Z /\ Z.max (ex - prec) emin = emin -> (Z.max (ex - prec) emin < ex)%Z -> (Z.max (ey - prec) emin < ey)%Z -> (Z.max (ex - ey - prec) emin <= ex - ey + 1)%Z -> (Z.max (ex - ey - prec') emin' <= Z.max (ex - prec) emin - ey)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin - prec - 2)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z

(ex - ey - prec' >= emin')%Z /\ Z.max (ex - ey - prec') emin' = (ex - ey - prec')%Z \/ (ex - ey - prec' < emin')%Z /\ Z.max (ex - ey - prec') emin' = emin' -> (ex - ey - prec >= emin)%Z /\ Z.max (ex - ey - prec) emin = (ex - ey - prec)%Z \/ (ex - ey - prec < emin)%Z /\ Z.max (ex - ey - prec) emin = emin -> (ey - prec >= emin)%Z /\ Z.max (ey - prec) emin = (ey - prec)%Z \/ (ey - prec < emin)%Z /\ Z.max (ey - prec) emin = emin -> (ex - prec >= emin)%Z /\ Z.max (ex - prec) emin = (ex - prec)%Z \/ (ex - prec < emin)%Z /\ Z.max (ex - prec) emin = emin -> (Z.max (ex - prec) emin < ex)%Z -> (Z.max (ey - prec) emin < ey)%Z -> (Z.max (ex - ey - prec) emin <= ex - ey + 1)%Z -> (Z.max (ex - ey - prec') emin' <= Z.max (ex - prec) emin - ey)%Z
omega.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin - prec - 2)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z

(Z.max (ex - prec) emin < ex)%Z -> (Z.max (ey - prec) emin < ey)%Z -> (Z.max (ex - ey + 1 - prec) emin <= ex - ey + 1 + 1)%Z -> (Z.max (ex - ey + 1 - prec') emin' <= Z.max (ex - prec) emin - ey)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin - prec - 2)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z

(ex - prec >= emin)%Z /\ Z.max (ex - prec) emin = (ex - prec)%Z \/ (ex - prec < emin)%Z /\ Z.max (ex - prec) emin = emin -> (Z.max (ex - prec) emin < ex)%Z -> (Z.max (ey - prec) emin < ey)%Z -> (Z.max (ex - ey + 1 - prec) emin <= ex - ey + 1 + 1)%Z -> (Z.max (ex - ey + 1 - prec') emin' <= Z.max (ex - prec) emin - ey)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin - prec - 2)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z

(ey - prec >= emin)%Z /\ Z.max (ey - prec) emin = (ey - prec)%Z \/ (ey - prec < emin)%Z /\ Z.max (ey - prec) emin = emin -> (ex - prec >= emin)%Z /\ Z.max (ex - prec) emin = (ex - prec)%Z \/ (ex - prec < emin)%Z /\ Z.max (ex - prec) emin = emin -> (Z.max (ex - prec) emin < ex)%Z -> (Z.max (ey - prec) emin < ey)%Z -> (Z.max (ex - ey + 1 - prec) emin <= ex - ey + 1 + 1)%Z -> (Z.max (ex - ey + 1 - prec') emin' <= Z.max (ex - prec) emin - ey)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin - prec - 2)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z

(ex - ey + 1 - prec >= emin)%Z /\ Z.max (ex - ey + 1 - prec) emin = (ex - ey + 1 - prec)%Z \/ (ex - ey + 1 - prec < emin)%Z /\ Z.max (ex - ey + 1 - prec) emin = emin -> (ey - prec >= emin)%Z /\ Z.max (ey - prec) emin = (ey - prec)%Z \/ (ey - prec < emin)%Z /\ Z.max (ey - prec) emin = emin -> (ex - prec >= emin)%Z /\ Z.max (ex - prec) emin = (ex - prec)%Z \/ (ex - prec < emin)%Z /\ Z.max (ex - prec) emin = emin -> (Z.max (ex - prec) emin < ex)%Z -> (Z.max (ey - prec) emin < ey)%Z -> (Z.max (ex - ey + 1 - prec) emin <= ex - ey + 1 + 1)%Z -> (Z.max (ex - ey + 1 - prec') emin' <= Z.max (ex - prec) emin - ey)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin - prec - 2)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z

(ex - ey + 1 - prec' >= emin')%Z /\ Z.max (ex - ey + 1 - prec') emin' = (ex - ey + 1 - prec')%Z \/ (ex - ey + 1 - prec' < emin')%Z /\ Z.max (ex - ey + 1 - prec') emin' = emin' -> (ex - ey + 1 - prec >= emin)%Z /\ Z.max (ex - ey + 1 - prec) emin = (ex - ey + 1 - prec)%Z \/ (ex - ey + 1 - prec < emin)%Z /\ Z.max (ex - ey + 1 - prec) emin = emin -> (ey - prec >= emin)%Z /\ Z.max (ey - prec) emin = (ey - prec)%Z \/ (ey - prec < emin)%Z /\ Z.max (ey - prec) emin = emin -> (ex - prec >= emin)%Z /\ Z.max (ex - prec) emin = (ex - prec)%Z \/ (ex - prec < emin)%Z /\ Z.max (ex - prec) emin = emin -> (Z.max (ex - prec) emin < ex)%Z -> (Z.max (ey - prec) emin < ey)%Z -> (Z.max (ex - ey + 1 - prec) emin <= ex - ey + 1 + 1)%Z -> (Z.max (ex - ey + 1 - prec') emin' <= Z.max (ex - prec) emin - ey)%Z
omega.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin - prec - 2)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z

(Z.max (ex - prec) emin < ex)%Z -> (Z.max (ey - prec) emin < ey)%Z -> (Z.max (ex - ey - prec) emin <= ex - ey)%Z -> (Z.max (ex - ey - prec') emin' <= Z.max (ex - ey - prec) emin + Z.max (ey - prec) emin - ey)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin - prec - 2)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z

(ex - prec >= emin)%Z /\ Z.max (ex - prec) emin = (ex - prec)%Z \/ (ex - prec < emin)%Z /\ Z.max (ex - prec) emin = emin -> (Z.max (ex - prec) emin < ex)%Z -> (Z.max (ey - prec) emin < ey)%Z -> (Z.max (ex - ey - prec) emin <= ex - ey)%Z -> (Z.max (ex - ey - prec') emin' <= Z.max (ex - ey - prec) emin + Z.max (ey - prec) emin - ey)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin - prec - 2)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z

(ey - prec >= emin)%Z /\ Z.max (ey - prec) emin = (ey - prec)%Z \/ (ey - prec < emin)%Z /\ Z.max (ey - prec) emin = emin -> (ex - prec >= emin)%Z /\ Z.max (ex - prec) emin = (ex - prec)%Z \/ (ex - prec < emin)%Z /\ Z.max (ex - prec) emin = emin -> (Z.max (ex - prec) emin < ex)%Z -> (Z.max (ey - prec) emin < ey)%Z -> (Z.max (ex - ey - prec) emin <= ex - ey)%Z -> (Z.max (ex - ey - prec') emin' <= Z.max (ex - ey - prec) emin + Z.max (ey - prec) emin - ey)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin - prec - 2)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z

(ex - ey - prec >= emin)%Z /\ Z.max (ex - ey - prec) emin = (ex - ey - prec)%Z \/ (ex - ey - prec < emin)%Z /\ Z.max (ex - ey - prec) emin = emin -> (ey - prec >= emin)%Z /\ Z.max (ey - prec) emin = (ey - prec)%Z \/ (ey - prec < emin)%Z /\ Z.max (ey - prec) emin = emin -> (ex - prec >= emin)%Z /\ Z.max (ex - prec) emin = (ex - prec)%Z \/ (ex - prec < emin)%Z /\ Z.max (ex - prec) emin = emin -> (Z.max (ex - prec) emin < ex)%Z -> (Z.max (ey - prec) emin < ey)%Z -> (Z.max (ex - ey - prec) emin <= ex - ey)%Z -> (Z.max (ex - ey - prec') emin' <= Z.max (ex - ey - prec) emin + Z.max (ey - prec) emin - ey)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin - prec - 2)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z

(ex - ey - prec' >= emin')%Z /\ Z.max (ex - ey - prec') emin' = (ex - ey - prec')%Z \/ (ex - ey - prec' < emin')%Z /\ Z.max (ex - ey - prec') emin' = emin' -> (ex - ey - prec >= emin)%Z /\ Z.max (ex - ey - prec) emin = (ex - ey - prec)%Z \/ (ex - ey - prec < emin)%Z /\ Z.max (ex - ey - prec) emin = emin -> (ey - prec >= emin)%Z /\ Z.max (ey - prec) emin = (ey - prec)%Z \/ (ey - prec < emin)%Z /\ Z.max (ey - prec) emin = emin -> (ex - prec >= emin)%Z /\ Z.max (ex - prec) emin = (ex - prec)%Z \/ (ex - prec < emin)%Z /\ Z.max (ex - prec) emin = emin -> (Z.max (ex - prec) emin < ex)%Z -> (Z.max (ey - prec) emin < ey)%Z -> (Z.max (ex - ey - prec) emin <= ex - ey)%Z -> (Z.max (ex - ey - prec') emin' <= Z.max (ex - ey - prec) emin + Z.max (ey - prec) emin - ey)%Z
omega.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin - prec - 2)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z

(Z.max (ex - prec) emin < ex)%Z -> (Z.max (ey - prec) emin < ey)%Z -> Z.max (ex - ey - prec) emin = (ex - ey + 1)%Z -> (Z.max (ex - ey - prec') emin' <= ex - ey - ey + Z.max (ey - prec) emin)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin - prec - 2)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z

(ex - prec >= emin)%Z /\ Z.max (ex - prec) emin = (ex - prec)%Z \/ (ex - prec < emin)%Z /\ Z.max (ex - prec) emin = emin -> (Z.max (ex - prec) emin < ex)%Z -> (Z.max (ey - prec) emin < ey)%Z -> Z.max (ex - ey - prec) emin = (ex - ey + 1)%Z -> (Z.max (ex - ey - prec') emin' <= ex - ey - ey + Z.max (ey - prec) emin)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin - prec - 2)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z

(ey - prec >= emin)%Z /\ Z.max (ey - prec) emin = (ey - prec)%Z \/ (ey - prec < emin)%Z /\ Z.max (ey - prec) emin = emin -> (ex - prec >= emin)%Z /\ Z.max (ex - prec) emin = (ex - prec)%Z \/ (ex - prec < emin)%Z /\ Z.max (ex - prec) emin = emin -> (Z.max (ex - prec) emin < ex)%Z -> (Z.max (ey - prec) emin < ey)%Z -> Z.max (ex - ey - prec) emin = (ex - ey + 1)%Z -> (Z.max (ex - ey - prec') emin' <= ex - ey - ey + Z.max (ey - prec) emin)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin - prec - 2)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z

(ex - ey - prec >= emin)%Z /\ Z.max (ex - ey - prec) emin = (ex - ey - prec)%Z \/ (ex - ey - prec < emin)%Z /\ Z.max (ex - ey - prec) emin = emin -> (ey - prec >= emin)%Z /\ Z.max (ey - prec) emin = (ey - prec)%Z \/ (ey - prec < emin)%Z /\ Z.max (ey - prec) emin = emin -> (ex - prec >= emin)%Z /\ Z.max (ex - prec) emin = (ex - prec)%Z \/ (ex - prec < emin)%Z /\ Z.max (ex - prec) emin = emin -> (Z.max (ex - prec) emin < ex)%Z -> (Z.max (ey - prec) emin < ey)%Z -> Z.max (ex - ey - prec) emin = (ex - ey + 1)%Z -> (Z.max (ex - ey - prec') emin' <= ex - ey - ey + Z.max (ey - prec) emin)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' <= emin - prec - 2)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z

(ex - ey - prec' >= emin')%Z /\ Z.max (ex - ey - prec') emin' = (ex - ey - prec')%Z \/ (ex - ey - prec' < emin')%Z /\ Z.max (ex - ey - prec') emin' = emin' -> (ex - ey - prec >= emin)%Z /\ Z.max (ex - ey - prec) emin = (ex - ey - prec)%Z \/ (ex - ey - prec < emin)%Z /\ Z.max (ex - ey - prec) emin = emin -> (ey - prec >= emin)%Z /\ Z.max (ey - prec) emin = (ey - prec)%Z \/ (ey - prec < emin)%Z /\ Z.max (ey - prec) emin = emin -> (ex - prec >= emin)%Z /\ Z.max (ex - prec) emin = (ex - prec)%Z \/ (ex - prec < emin)%Z /\ Z.max (ex - prec) emin = emin -> (Z.max (ex - prec) emin < ex)%Z -> (Z.max (ey - prec) emin < ey)%Z -> Z.max (ex - ey - prec) emin = (ex - ey + 1)%Z -> (Z.max (ex - ey - prec') emin' <= ex - ey - ey + Z.max (ey - prec) emin)%Z
omega. Qed.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

forall choice1 choice2 : Z -> bool, (exists n : Z, beta = (2 * n)%Z) -> (emin' <= emin - prec - 2)%Z -> (2 * prec <= prec')%Z -> forall x y : R, y <> 0 -> FLT_format beta emin prec x -> FLT_format beta emin prec y -> round_round_eq (FLT_exp emin prec) (FLT_exp emin' prec') choice1 choice2 (x / y)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

forall choice1 choice2 : Z -> bool, (exists n : Z, beta = (2 * n)%Z) -> (emin' <= emin - prec - 2)%Z -> (2 * prec <= prec')%Z -> forall x y : R, y <> 0 -> FLT_format beta emin prec x -> FLT_format beta emin prec y -> round_round_eq (FLT_exp emin prec) (FLT_exp emin' prec') choice1 choice2 (x / y)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hemin:(emin' <= emin - prec - 2)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Nzy:y <> 0
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y

round_round_eq (FLT_exp emin prec) (FLT_exp emin' prec') choice1 choice2 (x / y)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hemin:(emin' <= emin - prec - 2)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Nzy:y <> 0
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y

Valid_exp (FLT_exp emin prec)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hemin:(emin' <= emin - prec - 2)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Nzy:y <> 0
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y
Valid_exp (FLT_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hemin:(emin' <= emin - prec - 2)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Nzy:y <> 0
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y
exists n : Z, beta = (2 * n)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hemin:(emin' <= emin - prec - 2)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Nzy:y <> 0
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y
round_round_div_hyp (FLT_exp emin prec) (FLT_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hemin:(emin' <= emin - prec - 2)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Nzy:y <> 0
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y
y <> 0
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hemin:(emin' <= emin - prec - 2)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Nzy:y <> 0
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y
generic_format beta (FLT_exp emin prec) x
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hemin:(emin' <= emin - prec - 2)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Nzy:y <> 0
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y
generic_format beta (FLT_exp emin prec) y
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hemin:(emin' <= emin - prec - 2)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Nzy:y <> 0
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y

Valid_exp (FLT_exp emin prec)
now apply FLT_exp_valid.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hemin:(emin' <= emin - prec - 2)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Nzy:y <> 0
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y

Valid_exp (FLT_exp emin' prec')
now apply FLT_exp_valid.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hemin:(emin' <= emin - prec - 2)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Nzy:y <> 0
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y

exists n : Z, beta = (2 * n)%Z
exact Ebeta.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hemin:(emin' <= emin - prec - 2)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Nzy:y <> 0
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y

round_round_div_hyp (FLT_exp emin prec) (FLT_exp emin' prec')
now apply FLT_round_round_div_hyp.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hemin:(emin' <= emin - prec - 2)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Nzy:y <> 0
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y

y <> 0
exact Nzy.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hemin:(emin' <= emin - prec - 2)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Nzy:y <> 0
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y

generic_format beta (FLT_exp emin prec) x
now apply generic_format_FLT.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hemin:(emin' <= emin - prec - 2)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Nzy:y <> 0
Fx:FLT_format beta emin prec x
Fy:FLT_format beta emin prec y

generic_format beta (FLT_exp emin prec) y
now apply generic_format_FLT. Qed. End Double_round_div_FLT. Section Double_round_div_FTZ. Variable emin prec : Z. Variable emin' prec' : Z. Context { prec_gt_0_ : Prec_gt_0 prec }. Context { prec_gt_0_' : Prec_gt_0 prec' }.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(emin' + prec' <= emin - 1)%Z -> (2 * prec <= prec')%Z -> round_round_div_hyp (FTZ_exp emin prec) (FTZ_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

(emin' + prec' <= emin - 1)%Z -> (2 * prec <= prec')%Z -> round_round_div_hyp (FTZ_exp emin prec) (FTZ_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' + prec' <= emin - 1)%Z
Hprec:(2 * prec <= prec')%Z

round_round_div_hyp (FTZ_exp emin prec) (FTZ_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' + prec' <= emin - 1)%Z
Hprec:(2 * prec <= prec')%Z

round_round_div_hyp (fun e : Z => if (e - prec <? emin)%Z then (emin + prec - 1)%Z else (e - prec)%Z) (fun e : Z => if (e - prec' <? emin')%Z then (emin' + prec' - 1)%Z else (e - prec')%Z)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' + prec' <= emin - 1)%Z
Hprec:(2 * prec <= prec')%Z

round_round_div_hyp (fun e : Z => if (e - prec <? emin)%Z then (emin + prec - 1)%Z else (e - prec)%Z) (fun e : Z => if (e - prec' <? emin')%Z then (emin' + prec' - 1)%Z else (e - prec')%Z)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' + prec' <= emin - 1)%Z
Hprec:(2 * prec <= prec')%Z

round_round_div_hyp (fun e : Z => if (e - prec <? emin)%Z then (emin + prec - 1)%Z else (e - prec)%Z) (fun e : Z => if (e - prec' <? emin')%Z then (emin' + prec' - 1)%Z else (e - prec')%Z)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' + prec' <= emin - 1)%Z
Hprec:(2 * prec <= prec')%Z

(forall ex : Z, ((if ex - prec' <? emin' then emin' + prec' - 1 else ex - prec') <= (if ex - prec <? emin then emin + prec - 1 else ex - prec) - 1)%Z) /\ (forall ex ey : Z, ((if ex - prec <? emin then emin + prec - 1 else ex - prec) < ex)%Z -> ((if ey - prec <? emin then emin + prec - 1 else ey - prec) < ey)%Z -> ((if ex - ey - prec <? emin then emin + prec - 1 else ex - ey - prec) <= ex - ey + 1)%Z -> ((if ex - ey - prec' <? emin' then emin' + prec' - 1 else ex - ey - prec') <= (if ex - prec <? emin then emin + prec - 1 else ex - prec) - ey)%Z) /\ (forall ex ey : Z, ((if ex - prec <? emin then emin + prec - 1 else ex - prec) < ex)%Z -> ((if ey - prec <? emin then emin + prec - 1 else ey - prec) < ey)%Z -> ((if ex - ey + 1 - prec <? emin then emin + prec - 1 else ex - ey + 1 - prec) <= ex - ey + 1 + 1)%Z -> ((if ex - ey + 1 - prec' <? emin' then emin' + prec' - 1 else ex - ey + 1 - prec') <= (if ex - prec <? emin then emin + prec - 1 else ex - prec) - ey)%Z) /\ (forall ex ey : Z, ((if ex - prec <? emin then emin + prec - 1 else ex - prec) < ex)%Z -> ((if ey - prec <? emin then emin + prec - 1 else ey - prec) < ey)%Z -> ((if ex - ey - prec <? emin then emin + prec - 1 else ex - ey - prec) <= ex - ey)%Z -> ((if ex - ey - prec' <? emin' then emin' + prec' - 1 else ex - ey - prec') <= (if ex - ey - prec <? emin then emin + prec - 1 else ex - ey - prec) + (if ey - prec <? emin then emin + prec - 1 else ey - prec) - ey)%Z) /\ (forall ex ey : Z, ((if ex - prec <? emin then emin + prec - 1 else ex - prec) < ex)%Z -> ((if ey - prec <? emin then emin + prec - 1 else ey - prec) < ey)%Z -> (if (ex - ey - prec <? emin)%Z then (emin + prec - 1)%Z else (ex - ey - prec)%Z) = (ex - ey + 1)%Z -> ((if ex - ey - prec' <? emin' then emin' + prec' - 1 else ex - ey - prec') <= ex - ey - ey + (if ey - prec <? emin then emin + prec - 1 else ey - prec))%Z)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' + prec' <= emin - 1)%Z
Hprec:(2 * prec <= prec')%Z
ex:Z

((if ex - prec' <? emin' then emin' + prec' - 1 else ex - prec') <= (if ex - prec <? emin then emin + prec - 1 else ex - prec) - 1)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' + prec' <= emin - 1)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z
((if ex - prec <? emin then emin + prec - 1 else ex - prec) < ex)%Z -> ((if ey - prec <? emin then emin + prec - 1 else ey - prec) < ey)%Z -> ((if ex - ey - prec <? emin then emin + prec - 1 else ex - ey - prec) <= ex - ey + 1)%Z -> ((if ex - ey - prec' <? emin' then emin' + prec' - 1 else ex - ey - prec') <= (if ex - prec <? emin then emin + prec - 1 else ex - prec) - ey)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' + prec' <= emin - 1)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z
((if ex - prec <? emin then emin + prec - 1 else ex - prec) < ex)%Z -> ((if ey - prec <? emin then emin + prec - 1 else ey - prec) < ey)%Z -> ((if ex - ey + 1 - prec <? emin then emin + prec - 1 else ex - ey + 1 - prec) <= ex - ey + 1 + 1)%Z -> ((if ex - ey + 1 - prec' <? emin' then emin' + prec' - 1 else ex - ey + 1 - prec') <= (if ex - prec <? emin then emin + prec - 1 else ex - prec) - ey)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' + prec' <= emin - 1)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z
((if ex - prec <? emin then emin + prec - 1 else ex - prec) < ex)%Z -> ((if ey - prec <? emin then emin + prec - 1 else ey - prec) < ey)%Z -> ((if ex - ey - prec <? emin then emin + prec - 1 else ex - ey - prec) <= ex - ey)%Z -> ((if ex - ey - prec' <? emin' then emin' + prec' - 1 else ex - ey - prec') <= (if ex - ey - prec <? emin then emin + prec - 1 else ex - ey - prec) + (if ey - prec <? emin then emin + prec - 1 else ey - prec) - ey)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' + prec' <= emin - 1)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z
((if ex - prec <? emin then emin + prec - 1 else ex - prec) < ex)%Z -> ((if ey - prec <? emin then emin + prec - 1 else ey - prec) < ey)%Z -> (if (ex - ey - prec <? emin)%Z then (emin + prec - 1)%Z else (ex - ey - prec)%Z) = (ex - ey + 1)%Z -> ((if ex - ey - prec' <? emin' then emin' + prec' - 1 else ex - ey - prec') <= ex - ey - ey + (if ey - prec <? emin then emin + prec - 1 else ey - prec))%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' + prec' <= emin - 1)%Z
Hprec:(2 * prec <= prec')%Z
ex:Z

((if ex - prec' <? emin' then emin' + prec' - 1 else ex - prec') <= (if ex - prec <? emin then emin + prec - 1 else ex - prec) - 1)%Z
destruct (Z.ltb_spec (ex - prec') emin'); destruct (Z.ltb_spec (ex - prec) emin); omega.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' + prec' <= emin - 1)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z

((if ex - prec <? emin then emin + prec - 1 else ex - prec) < ex)%Z -> ((if ey - prec <? emin then emin + prec - 1 else ey - prec) < ey)%Z -> ((if ex - ey - prec <? emin then emin + prec - 1 else ex - ey - prec) <= ex - ey + 1)%Z -> ((if ex - ey - prec' <? emin' then emin' + prec' - 1 else ex - ey - prec') <= (if ex - prec <? emin then emin + prec - 1 else ex - prec) - ey)%Z
destruct (Z.ltb_spec (ex - prec) emin); destruct (Z.ltb_spec (ey - prec) emin); destruct (Z.ltb_spec (ex - ey - prec) emin); destruct (Z.ltb_spec (ex - ey - prec') emin'); omega.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' + prec' <= emin - 1)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z

((if ex - prec <? emin then emin + prec - 1 else ex - prec) < ex)%Z -> ((if ey - prec <? emin then emin + prec - 1 else ey - prec) < ey)%Z -> ((if ex - ey + 1 - prec <? emin then emin + prec - 1 else ex - ey + 1 - prec) <= ex - ey + 1 + 1)%Z -> ((if ex - ey + 1 - prec' <? emin' then emin' + prec' - 1 else ex - ey + 1 - prec') <= (if ex - prec <? emin then emin + prec - 1 else ex - prec) - ey)%Z
destruct (Z.ltb_spec (ex - prec) emin); destruct (Z.ltb_spec (ey - prec) emin); destruct (Z.ltb_spec (ex - ey + 1 - prec) emin); destruct (Z.ltb_spec (ex - ey + 1 - prec') emin'); omega.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' + prec' <= emin - 1)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z

((if ex - prec <? emin then emin + prec - 1 else ex - prec) < ex)%Z -> ((if ey - prec <? emin then emin + prec - 1 else ey - prec) < ey)%Z -> ((if ex - ey - prec <? emin then emin + prec - 1 else ex - ey - prec) <= ex - ey)%Z -> ((if ex - ey - prec' <? emin' then emin' + prec' - 1 else ex - ey - prec') <= (if ex - ey - prec <? emin then emin + prec - 1 else ex - ey - prec) + (if ey - prec <? emin then emin + prec - 1 else ey - prec) - ey)%Z
destruct (Z.ltb_spec (ex - prec) emin); destruct (Z.ltb_spec (ey - prec) emin); destruct (Z.ltb_spec (ex - ey - prec) emin); destruct (Z.ltb_spec (ex - ey - prec') emin'); omega.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:(0 < prec)%Z
prec_gt_0_':Prec_gt_0 prec'
Hemin:(emin' + prec' <= emin - 1)%Z
Hprec:(2 * prec <= prec')%Z
ex, ey:Z

((if ex - prec <? emin then emin + prec - 1 else ex - prec) < ex)%Z -> ((if ey - prec <? emin then emin + prec - 1 else ey - prec) < ey)%Z -> (if (ex - ey - prec <? emin)%Z then (emin + prec - 1)%Z else (ex - ey - prec)%Z) = (ex - ey + 1)%Z -> ((if ex - ey - prec' <? emin' then emin' + prec' - 1 else ex - ey - prec') <= ex - ey - ey + (if ey - prec <? emin then emin + prec - 1 else ey - prec))%Z
destruct (Z.ltb_spec (ex - prec) emin); destruct (Z.ltb_spec (ey - prec) emin); destruct (Z.ltb_spec (ex - ey - prec) emin); destruct (Z.ltb_spec (ex - ey - prec') emin'); omega. Qed.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

forall choice1 choice2 : Z -> bool, (exists n : Z, beta = (2 * n)%Z) -> (emin' + prec' <= emin - 1)%Z -> (2 * prec <= prec')%Z -> forall x y : R, y <> 0 -> FTZ_format beta emin prec x -> FTZ_format beta emin prec y -> round_round_eq (FTZ_exp emin prec) (FTZ_exp emin' prec') choice1 choice2 (x / y)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'

forall choice1 choice2 : Z -> bool, (exists n : Z, beta = (2 * n)%Z) -> (emin' + prec' <= emin - 1)%Z -> (2 * prec <= prec')%Z -> forall x y : R, y <> 0 -> FTZ_format beta emin prec x -> FTZ_format beta emin prec y -> round_round_eq (FTZ_exp emin prec) (FTZ_exp emin' prec') choice1 choice2 (x / y)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hemin:(emin' + prec' <= emin - 1)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Nzy:y <> 0
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y

round_round_eq (FTZ_exp emin prec) (FTZ_exp emin' prec') choice1 choice2 (x / y)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hemin:(emin' + prec' <= emin - 1)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Nzy:y <> 0
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y

Valid_exp (FTZ_exp emin prec)
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hemin:(emin' + prec' <= emin - 1)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Nzy:y <> 0
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y
Valid_exp (FTZ_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hemin:(emin' + prec' <= emin - 1)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Nzy:y <> 0
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y
exists n : Z, beta = (2 * n)%Z
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hemin:(emin' + prec' <= emin - 1)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Nzy:y <> 0
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y
round_round_div_hyp (FTZ_exp emin prec) (FTZ_exp emin' prec')
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hemin:(emin' + prec' <= emin - 1)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Nzy:y <> 0
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y
y <> 0
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hemin:(emin' + prec' <= emin - 1)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Nzy:y <> 0
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y
generic_format beta (FTZ_exp emin prec) x
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hemin:(emin' + prec' <= emin - 1)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Nzy:y <> 0
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y
generic_format beta (FTZ_exp emin prec) y
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hemin:(emin' + prec' <= emin - 1)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Nzy:y <> 0
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y

Valid_exp (FTZ_exp emin prec)
now apply FTZ_exp_valid.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hemin:(emin' + prec' <= emin - 1)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Nzy:y <> 0
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y

Valid_exp (FTZ_exp emin' prec')
now apply FTZ_exp_valid.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hemin:(emin' + prec' <= emin - 1)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Nzy:y <> 0
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y

exists n : Z, beta = (2 * n)%Z
exact Ebeta.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hemin:(emin' + prec' <= emin - 1)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Nzy:y <> 0
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y

round_round_div_hyp (FTZ_exp emin prec) (FTZ_exp emin' prec')
now apply FTZ_round_round_div_hyp.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hemin:(emin' + prec' <= emin - 1)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Nzy:y <> 0
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y

y <> 0
exact Nzy.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hemin:(emin' + prec' <= emin - 1)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Nzy:y <> 0
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y

generic_format beta (FTZ_exp emin prec) x
now apply generic_format_FTZ.
beta:radix
emin, prec, emin', prec':Z
prec_gt_0_:Prec_gt_0 prec
prec_gt_0_':Prec_gt_0 prec'
choice1, choice2:Z -> bool
Ebeta:exists n : Z, beta = (2 * n)%Z
Hemin:(emin' + prec' <= emin - 1)%Z
Hprec:(2 * prec <= prec')%Z
x, y:R
Nzy:y <> 0
Fx:FTZ_format beta emin prec x
Fy:FTZ_format beta emin prec y

generic_format beta (FTZ_exp emin prec) y
now apply generic_format_FTZ. Qed. End Double_round_div_FTZ. End Double_round_div. End Double_round.