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

Missing definitions/lemmas

Require Import Psatz.
Require Export Reals ZArith.
Require Export Zaux.

Section Rmissing.
About R

forall x y : R, (x <= y)%R -> (0 <= y - x)%R

forall x y : R, (x <= y)%R -> (0 <= y - x)%R
x, y:R
H:(x <= y)%R

(0 <= y - x)%R
x, y:R
H:(x <= y)%R

(y - x >= 0)%R
x, y:R
H:(x <= y)%R

(y >= x)%R
now apply Rle_ge. Qed.

forall x y : R, Rabs x = Rabs y -> x = y \/ x = (- y)%R

forall x y : R, Rabs x = Rabs y -> x = y \/ x = (- y)%R
x, y:R
H:Rabs x = Rabs y

x = y \/ x = (- y)%R
x, y:R
H:(if Rcase_abs x then (- x)%R else x) = (if Rcase_abs y then (- y)%R else y)

x = y \/ x = (- y)%R
x, y:R
H:(- x)%R = (if Rcase_abs y then (- y)%R else y)

x = y \/ x = (- y)%R
x, y:R
H:x = (if Rcase_abs y then (- y)%R else y)
x = y \/ x = (- y)%R
x, y:R
H:(- x)%R = (if Rcase_abs y then (- y)%R else y)
H':(- - x)%R = (- (if Rcase_abs y then - y else y))%R

x = y \/ x = (- y)%R
x, y:R
H:x = (if Rcase_abs y then (- y)%R else y)
x = y \/ x = (- y)%R
x, y:R
H:(- x)%R = (if Rcase_abs y then (- y)%R else y)
H':x = (- (if Rcase_abs y then - y else y))%R

x = y \/ x = (- y)%R
x, y:R
H:x = (if Rcase_abs y then (- y)%R else y)
x = y \/ x = (- y)%R
x, y:R
H:(- x)%R = (if Rcase_abs y then (- y)%R else y)
H':x = (- (if Rcase_abs y then - y else y))%R

(- (if Rcase_abs y then - y else y))%R = y \/ (- (if Rcase_abs y then - y else y))%R = (- y)%R
x, y:R
H:x = (if Rcase_abs y then (- y)%R else y)
x = y \/ x = (- y)%R
x, y:R
H:(- x)%R = (- y)%R
H':x = (- - y)%R

(- - y)%R = y \/ (- - y)%R = (- y)%R
x, y:R
H:(- x)%R = y
H':x = (- y)%R
(- y)%R = y \/ (- y)%R = (- y)%R
x, y:R
H:x = (if Rcase_abs y then (- y)%R else y)
x = y \/ x = (- y)%R
x, y:R
H:(- x)%R = (- y)%R
H':x = (- - y)%R

(- - y)%R = y
x, y:R
H:(- x)%R = y
H':x = (- y)%R
(- y)%R = y \/ (- y)%R = (- y)%R
x, y:R
H:x = (if Rcase_abs y then (- y)%R else y)
x = y \/ x = (- y)%R
x, y:R
H:(- x)%R = y
H':x = (- y)%R

(- y)%R = y \/ (- y)%R = (- y)%R
x, y:R
H:x = (if Rcase_abs y then (- y)%R else y)
x = y \/ x = (- y)%R
x, y:R
H:x = (if Rcase_abs y then (- y)%R else y)

x = y \/ x = (- y)%R
x, y:R
H:x = (if Rcase_abs y then (- y)%R else y)

(if Rcase_abs y then (- y)%R else y) = y \/ (if Rcase_abs y then (- y)%R else y) = (- y)%R
now destruct (Rcase_abs y) as [_|_] ; [right|left]. Qed.

forall x y : R, (0 <= y)%R -> (y <= 2 * x)%R -> (Rabs (x - y) <= x)%R

forall x y : R, (0 <= y)%R -> (y <= 2 * x)%R -> (Rabs (x - y) <= x)%R
x, y:R
Hx:(0 <= y)%R
Hy:(y <= 2 * x)%R

(Rabs (x - y) <= x)%R
x, y:R
Hx:(0 <= y)%R
Hy:(y <= 2 * x)%R

(- x <= x - y <= x)%R
lra. Qed.
x:R

Rabs x = 0%R -> x = 0%R
x:R

Rabs x = 0%R -> x = 0%R
split_Rabs; lra. Qed.

forall r r1 r2 : R, (r1 + r)%R = (r2 + r)%R -> r1 = r2

forall r r1 r2 : R, (r1 + r)%R = (r2 + r)%R -> r1 = r2
r, r1, r2:R
H:(r1 + r)%R = (r2 + r)%R

r1 = r2
r, r1, r2:R
H:(r1 + r)%R = (r2 + r)%R

(r + r1)%R = (r + r2)%R
now rewrite 2!(Rplus_comm r). Qed.

forall r1 r2 r3 r4 : R, (0 <= r1)%R -> (0 <= r3)%R -> (r1 < r2)%R -> (r3 < r4)%R -> (r1 * r3 < r2 * r4)%R

forall r1 r2 r3 r4 : R, (0 <= r1)%R -> (0 <= r3)%R -> (r1 < r2)%R -> (r3 < r4)%R -> (r1 * r3 < r2 * r4)%R
r1, r2, r3, r4:R
Pr1:(0 <= r1)%R
Pr3:(0 <= r3)%R
H12:(r1 < r2)%R
H34:(r3 < r4)%R

(r1 * r3 < r2 * r4)%R
r1, r2, r3, r4:R
Pr1:(0 <= r1)%R
Pr3:(0 <= r3)%R
H12:(r1 < r2)%R
H34:(r3 < r4)%R

(r1 * r3 <= r1 * r4)%R
r1, r2, r3, r4:R
Pr1:(0 <= r1)%R
Pr3:(0 <= r3)%R
H12:(r1 < r2)%R
H34:(r3 < r4)%R
(r1 * r4 < r2 * r4)%R
r1, r2, r3, r4:R
Pr1:(0 <= r1)%R
Pr3:(0 <= r3)%R
H12:(r1 < r2)%R
H34:(r3 < r4)%R

(r1 * r3 <= r1 * r4)%R
r1, r2, r3, r4:R
Pr1:(0 <= r1)%R
Pr3:(0 <= r3)%R
H12:(r1 < r2)%R
H34:(r3 < r4)%R

(0 <= r1)%R
r1, r2, r3, r4:R
Pr1:(0 <= r1)%R
Pr3:(0 <= r3)%R
H12:(r1 < r2)%R
H34:(r3 < r4)%R
(r3 <= r4)%R
r1, r2, r3, r4:R
Pr1:(0 <= r1)%R
Pr3:(0 <= r3)%R
H12:(r1 < r2)%R
H34:(r3 < r4)%R

(0 <= r1)%R
exact Pr1.
r1, r2, r3, r4:R
Pr1:(0 <= r1)%R
Pr3:(0 <= r3)%R
H12:(r1 < r2)%R
H34:(r3 < r4)%R

(r3 <= r4)%R
now apply Rlt_le.
r1, r2, r3, r4:R
Pr1:(0 <= r1)%R
Pr3:(0 <= r3)%R
H12:(r1 < r2)%R
H34:(r3 < r4)%R

(r1 * r4 < r2 * r4)%R
r1, r2, r3, r4:R
Pr1:(0 <= r1)%R
Pr3:(0 <= r3)%R
H12:(r1 < r2)%R
H34:(r3 < r4)%R

(0 < r4)%R
r1, r2, r3, r4:R
Pr1:(0 <= r1)%R
Pr3:(0 <= r3)%R
H12:(r1 < r2)%R
H34:(r3 < r4)%R
(r1 < r2)%R
r1, r2, r3, r4:R
Pr1:(0 <= r1)%R
Pr3:(0 <= r3)%R
H12:(r1 < r2)%R
H34:(r3 < r4)%R

(0 < r4)%R
now apply Rle_lt_trans with r3.
r1, r2, r3, r4:R
Pr1:(0 <= r1)%R
Pr3:(0 <= r3)%R
H12:(r1 < r2)%R
H34:(r3 < r4)%R

(r1 < r2)%R
exact H12. Qed.

forall r r1 r2 : R, ((r1 - r2) * r)%R = (r1 * r - r2 * r)%R

forall r r1 r2 : R, ((r1 - r2) * r)%R = (r1 * r - r2 * r)%R
r, r1, r2:R

((r1 - r2) * r)%R = (r1 * r - r2 * r)%R
r, r1, r2:R

(r * (r1 - r2))%R = (r * r1 - r * r2)%R
apply Rmult_minus_distr_l. Qed.

forall r1 r2 r3 : R, (r2 * r1)%R <> (r3 * r1)%R -> r2 <> r3

forall r1 r2 r3 : R, (r2 * r1)%R <> (r3 * r1)%R -> r2 <> r3
r1, r2, r3:R
H1:(r2 * r1)%R <> (r3 * r1)%R
H2:r2 = r3

False
apply H1; rewrite H2; ring. Qed.

forall r1 r2 r3 : R, r1 <> 0%R -> r2 <> r3 -> (r2 * r1)%R <> (r3 * r1)%R

forall r1 r2 r3 : R, r1 <> 0%R -> r2 <> r3 -> (r2 * r1)%R <> (r3 * r1)%R
r1, r2, r3:R
H:r1 <> 0%R
H1:r2 <> r3
H2:(r2 * r1)%R = (r3 * r1)%R

False
now apply H1, Rmult_eq_reg_r with r1. Qed.

forall r r1 r2 : R, (0 <= r)%R -> (Rmin r1 r2 * r)%R = Rmin (r1 * r) (r2 * r)

forall r r1 r2 : R, (0 <= r)%R -> (Rmin r1 r2 * r)%R = Rmin (r1 * r) (r2 * r)
r, r1, r2:R
Hr:(0 < r)%R

(Rmin r1 r2 * r)%R = Rmin (r1 * r) (r2 * r)
r, r1, r2:R
Hr:0%R = r
(Rmin r1 r2 * r)%R = Rmin (r1 * r) (r2 * r)
r, r1, r2:R
Hr:(0 < r)%R

((if Rle_dec r1 r2 then r1 else r2) * r)%R = (if Rle_dec (r1 * r) (r2 * r) then (r1 * r)%R else (r2 * r)%R)
r, r1, r2:R
Hr:0%R = r
(Rmin r1 r2 * r)%R = Rmin (r1 * r) (r2 * r)
r, r1, r2:R
Hr:(0 < r)%R
H1:(r1 <= r2)%R
H2:~ (r1 * r <= r2 * r)%R

(r1 * r)%R = (r2 * r)%R
r, r1, r2:R
Hr:(0 < r)%R
H1:~ (r1 <= r2)%R
H2:(r1 * r <= r2 * r)%R
(r2 * r)%R = (r1 * r)%R
r, r1, r2:R
Hr:0%R = r
(Rmin r1 r2 * r)%R = Rmin (r1 * r) (r2 * r)
r, r1, r2:R
Hr:(0 < r)%R
H1:(r1 <= r2)%R
H2:~ (r1 * r <= r2 * r)%R

r1 = r2
r, r1, r2:R
Hr:(0 < r)%R
H1:~ (r1 <= r2)%R
H2:(r1 * r <= r2 * r)%R
(r2 * r)%R = (r1 * r)%R
r, r1, r2:R
Hr:0%R = r
(Rmin r1 r2 * r)%R = Rmin (r1 * r) (r2 * r)
r, r1, r2:R
Hr:(0 < r)%R
H1:(r1 <= r2)%R
H2:~ (r1 * r <= r2 * r)%R

(r1 <= r2)%R
r, r1, r2:R
Hr:(0 < r)%R
H1:(r1 <= r2)%R
H2:~ (r1 * r <= r2 * r)%R
(r2 <= r1)%R
r, r1, r2:R
Hr:(0 < r)%R
H1:~ (r1 <= r2)%R
H2:(r1 * r <= r2 * r)%R
(r2 * r)%R = (r1 * r)%R
r, r1, r2:R
Hr:0%R = r
(Rmin r1 r2 * r)%R = Rmin (r1 * r) (r2 * r)
r, r1, r2:R
Hr:(0 < r)%R
H1:(r1 <= r2)%R
H2:~ (r1 * r <= r2 * r)%R

(r2 <= r1)%R
r, r1, r2:R
Hr:(0 < r)%R
H1:~ (r1 <= r2)%R
H2:(r1 * r <= r2 * r)%R
(r2 * r)%R = (r1 * r)%R
r, r1, r2:R
Hr:0%R = r
(Rmin r1 r2 * r)%R = Rmin (r1 * r) (r2 * r)
r, r1, r2:R
Hr:(0 < r)%R
H1:(r1 <= r2)%R
H2:~ (r1 * r <= r2 * r)%R

(r2 * r <= r1 * r)%R
r, r1, r2:R
Hr:(0 < r)%R
H1:~ (r1 <= r2)%R
H2:(r1 * r <= r2 * r)%R
(r2 * r)%R = (r1 * r)%R
r, r1, r2:R
Hr:0%R = r
(Rmin r1 r2 * r)%R = Rmin (r1 * r) (r2 * r)
r, r1, r2:R
Hr:(0 < r)%R
H1:(r1 <= r2)%R
H2:~ (r1 * r <= r2 * r)%R

(r2 * r < r1 * r)%R
r, r1, r2:R
Hr:(0 < r)%R
H1:~ (r1 <= r2)%R
H2:(r1 * r <= r2 * r)%R
(r2 * r)%R = (r1 * r)%R
r, r1, r2:R
Hr:0%R = r
(Rmin r1 r2 * r)%R = Rmin (r1 * r) (r2 * r)
r, r1, r2:R
Hr:(0 < r)%R
H1:~ (r1 <= r2)%R
H2:(r1 * r <= r2 * r)%R

(r2 * r)%R = (r1 * r)%R
r, r1, r2:R
Hr:0%R = r
(Rmin r1 r2 * r)%R = Rmin (r1 * r) (r2 * r)
r, r1, r2:R
Hr:(0 < r)%R
H1:~ (r1 <= r2)%R
H2:(r1 * r <= r2 * r)%R

(r2 * r <= r1 * r)%R
r, r1, r2:R
Hr:(0 < r)%R
H1:~ (r1 <= r2)%R
H2:(r1 * r <= r2 * r)%R
(r1 * r <= r2 * r)%R
r, r1, r2:R
Hr:0%R = r
(Rmin r1 r2 * r)%R = Rmin (r1 * r) (r2 * r)
r, r1, r2:R
Hr:(0 < r)%R
H1:~ (r1 <= r2)%R
H2:(r1 * r <= r2 * r)%R

(0 <= r)%R
r, r1, r2:R
Hr:(0 < r)%R
H1:~ (r1 <= r2)%R
H2:(r1 * r <= r2 * r)%R
(r2 <= r1)%R
r, r1, r2:R
Hr:(0 < r)%R
H1:~ (r1 <= r2)%R
H2:(r1 * r <= r2 * r)%R
(r1 * r <= r2 * r)%R
r, r1, r2:R
Hr:0%R = r
(Rmin r1 r2 * r)%R = Rmin (r1 * r) (r2 * r)
r, r1, r2:R
Hr:(0 < r)%R
H1:~ (r1 <= r2)%R
H2:(r1 * r <= r2 * r)%R

(r2 <= r1)%R
r, r1, r2:R
Hr:(0 < r)%R
H1:~ (r1 <= r2)%R
H2:(r1 * r <= r2 * r)%R
(r1 * r <= r2 * r)%R
r, r1, r2:R
Hr:0%R = r
(Rmin r1 r2 * r)%R = Rmin (r1 * r) (r2 * r)
r, r1, r2:R
Hr:(0 < r)%R
H1:~ (r1 <= r2)%R
H2:(r1 * r <= r2 * r)%R

(r2 < r1)%R
r, r1, r2:R
Hr:(0 < r)%R
H1:~ (r1 <= r2)%R
H2:(r1 * r <= r2 * r)%R
(r1 * r <= r2 * r)%R
r, r1, r2:R
Hr:0%R = r
(Rmin r1 r2 * r)%R = Rmin (r1 * r) (r2 * r)
r, r1, r2:R
Hr:(0 < r)%R
H1:~ (r1 <= r2)%R
H2:(r1 * r <= r2 * r)%R

(r1 * r <= r2 * r)%R
r, r1, r2:R
Hr:0%R = r
(Rmin r1 r2 * r)%R = Rmin (r1 * r) (r2 * r)
r, r1, r2:R
Hr:0%R = r

(Rmin r1 r2 * r)%R = Rmin (r1 * r) (r2 * r)
r, r1, r2:R
Hr:0%R = r

(Rmin r1 r2 * 0)%R = Rmin (r1 * 0) (r2 * 0)
r, r1, r2:R
Hr:0%R = r

0%R = Rmin 0 0
r, r1, r2:R
Hr:0%R = r

0%R = (if Rle_dec 0 0 then 0%R else 0%R)
r, r1, r2:R
Hr:0%R = r
H0:(0 <= 0)%R

0%R = 0%R
r, r1, r2:R
Hr:0%R = r
H0:~ (0 <= 0)%R
0%R = 0%R
r, r1, r2:R
Hr:0%R = r
H0:~ (0 <= 0)%R

0%R = 0%R
r, r1, r2:R
Hr:0%R = r
H0:~ (0 <= 0)%R

(0 <= 0)%R
apply Rle_refl. Qed.

forall r r1 r2 : R, (0 <= r)%R -> (r * Rmin r1 r2)%R = Rmin (r * r1) (r * r2)

forall r r1 r2 : R, (0 <= r)%R -> (r * Rmin r1 r2)%R = Rmin (r * r1) (r * r2)
r, r1, r2:R
Hr:(0 <= r)%R

(r * Rmin r1 r2)%R = Rmin (r * r1) (r * r2)
r, r1, r2:R
Hr:(0 <= r)%R

(Rmin r1 r2 * r)%R = Rmin (r1 * r) (r2 * r)
now apply Rmult_min_distr_r. Qed.

forall x y : R, Rmin (- x) (- y) = (- Rmax x y)%R

forall x y : R, Rmin (- x) (- y) = (- Rmax x y)%R
x, y:R

Rmin (- x) (- y) = (- Rmax x y)%R
x, y:R
H:(y <= x)%R

Rmin (- x) (- y) = (- x)%R
x, y:R
H:(x <= y)%R
Rmin (- x) (- y) = (- y)%R
x, y:R
H:(y <= x)%R

(- x <= - y)%R
x, y:R
H:(x <= y)%R
Rmin (- x) (- y) = (- y)%R
x, y:R
H:(x <= y)%R

Rmin (- x) (- y) = (- y)%R
x, y:R
H:(x <= y)%R

(- y <= - x)%R
now apply Ropp_le_contravar. Qed.

forall x y : R, Rmax (- x) (- y) = (- Rmin x y)%R

forall x y : R, Rmax (- x) (- y) = (- Rmin x y)%R
x, y:R

Rmax (- x) (- y) = (- Rmin x y)%R
x, y:R
H:(x <= y)%R

Rmax (- x) (- y) = (- x)%R
x, y:R
H:(y <= x)%R
Rmax (- x) (- y) = (- y)%R
x, y:R
H:(x <= y)%R

(- y <= - x)%R
x, y:R
H:(y <= x)%R
Rmax (- x) (- y) = (- y)%R
x, y:R
H:(y <= x)%R

Rmax (- x) (- y) = (- y)%R
x, y:R
H:(y <= x)%R

(- x <= - y)%R
now apply Ropp_le_contravar. Qed.

forall x y : R, (x <= y)%R -> (exp x <= exp y)%R

forall x y : R, (x <= y)%R -> (exp x <= exp y)%R
x, y:R
H:(x < y)%R

(exp x <= exp y)%R
x, y:R
H:x = y
(exp x <= exp y)%R
x, y:R
H:(x < y)%R

(exp x < exp y)%R
x, y:R
H:x = y
(exp x <= exp y)%R
x, y:R
H:x = y

(exp x <= exp y)%R
x, y:R
H:x = y

(exp y <= exp y)%R
apply Rle_refl. Qed.

forall x y : R, (0 < x)%R -> (x < y)%R -> (/ y < / x)%R

forall x y : R, (0 < x)%R -> (x < y)%R -> (/ y < / x)%R
x, y:R
Hx:(0 < x)%R
Hxy:(x < y)%R

(/ y < / x)%R
x, y:R
Hx:(0 < x)%R
Hxy:(x < y)%R

(0 < x * y)%R
x, y:R
Hx:(0 < x)%R
Hxy:(x < y)%R
(x < y)%R
x, y:R
Hx:(0 < x)%R
Hxy:(x < y)%R

(0 < x)%R
x, y:R
Hx:(0 < x)%R
Hxy:(x < y)%R
(0 < y)%R
x, y:R
Hx:(0 < x)%R
Hxy:(x < y)%R
(x < y)%R
x, y:R
Hx:(0 < x)%R
Hxy:(x < y)%R

(0 < y)%R
x, y:R
Hx:(0 < x)%R
Hxy:(x < y)%R
(x < y)%R
x, y:R
Hx:(0 < x)%R
Hxy:(x < y)%R

(x < y)%R
exact Hxy. Qed.

forall x y : R, (0 < x)%R -> (x <= y)%R -> (/ y <= / x)%R

forall x y : R, (0 < x)%R -> (x <= y)%R -> (/ y <= / x)%R
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R

(/ y <= / x)%R
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R

(0 < x)%R
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
(0 < y)%R
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
(x <= y)%R
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R

(0 < y)%R
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R
(x <= y)%R
x, y:R
Hx:(0 < x)%R
Hxy:(x <= y)%R

(x <= y)%R
exact Hxy. Qed.

forall x : R, (0 <= sqrt x)%R

forall x : R, (0 <= sqrt x)%R
x:R

(0 <= sqrt x)%R
x:R

(0 <= match Rcase_abs x with | left _ => 0 | right a => Rsqrt {| nonneg := x; cond_nonneg := Rge_le x 0 a |} end)%R
x:R

(0 <= 0)%R
x:R
H:(x >= 0)%R
(0 <= Rsqrt {| nonneg := x; cond_nonneg := Rge_le x 0 H |})%R
x:R
H:(x >= 0)%R

(0 <= Rsqrt {| nonneg := x; cond_nonneg := Rge_le x 0 H |})%R
apply Rsqrt_positivity. Qed.

forall x : R, (x <= 0)%R -> sqrt x = 0%R

forall x : R, (x <= 0)%R -> sqrt x = 0%R
x:R
Npx:(x <= 0)%R

sqrt x = 0%R
x:R
Npx:(x <= 0)%R
Zx:x = 0%R

sqrt x = 0%R
x:R
Npx:(x <= 0)%R
Nzx:x <> 0%R
sqrt x = 0%R
x:R
Npx:(x <= 0)%R
Zx:x = 0%R

sqrt x = 0%R
x:R
Npx:(x <= 0)%R
Zx:x = 0%R

sqrt 0 = 0%R
exact sqrt_0.
x:R
Npx:(x <= 0)%R
Nzx:x <> 0%R

sqrt x = 0%R
x:R
Npx:(x <= 0)%R
Nzx:x <> 0%R

match Rcase_abs x with | left _ => 0%R | right a => Rsqrt {| nonneg := x; cond_nonneg := Rge_le x 0 a |} end = 0%R
x:R
Npx:(x <= 0)%R
Nzx:x <> 0%R
r:(x < 0)%R

0%R = 0%R
x:R
Npx:(x <= 0)%R
Nzx:x <> 0%R
r:(x >= 0)%R
Rsqrt {| nonneg := x; cond_nonneg := Rge_le x 0 r |} = 0%R
x:R
Npx:(x <= 0)%R
Nzx:x <> 0%R
r:(x < 0)%R

0%R = 0%R
reflexivity.
x:R
Npx:(x <= 0)%R
Nzx:x <> 0%R
r:(x >= 0)%R

Rsqrt {| nonneg := x; cond_nonneg := Rge_le x 0 r |} = 0%R
x:R
Npx:(x <= 0)%R
Nzx:x <> 0%R
r:(x >= 0)%R

False
now apply Nzx, Rle_antisym; [|apply Rge_le]. Qed.

forall x y : R, (x² <= y²)%R -> (x <= Rabs y)%R

forall x y : R, (x² <= y²)%R -> (x <= Rabs y)%R
x, y:R
H:(x² <= y²)%R

(x <= Rabs y)%R
apply (Rle_trans _ (Rabs x)); [apply Rle_abs|apply (Rsqr_le_abs_0 _ _ H)]. Qed.

forall x y : R, (- y <= x <= y)%R -> (Rabs x <= y)%R

forall x y : R, (- y <= x <= y)%R -> (Rabs x <= y)%R
x, y:R
Hyx:(- y <= x)%R
Hxy:(x <= y)%R

(Rabs x <= y)%R
x, y:R
Hyx:(- y <= x)%R
Hxy:(x <= y)%R

((if Rcase_abs x then - x else x) <= y)%R
x, y:R
Hyx:(- y <= x)%R
Hxy:(x <= y)%R
Hx:(x < 0)%R

(- x <= y)%R
x, y:R
Hyx:(- y <= x)%R
Hxy:(x <= y)%R
Hx:(x >= 0)%R
(x <= y)%R
x, y:R
Hyx:(- y <= x)%R
Hxy:(x <= y)%R
Hx:(x < 0)%R

(- y <= - - x)%R
x, y:R
Hyx:(- y <= x)%R
Hxy:(x <= y)%R
Hx:(x >= 0)%R
(x <= y)%R
x, y:R
Hyx:(- y <= x)%R
Hxy:(x <= y)%R
Hx:(x >= 0)%R

(x <= y)%R
exact Hxy. Qed.

forall x y : R, (Rabs x <= y)%R -> (- y <= x <= y)%R

forall x y : R, (Rabs x <= y)%R -> (- y <= x <= y)%R
x, y:R
Hxy:(Rabs x <= y)%R

(- y <= x <= y)%R
x, y:R
Hxy:(Rabs x <= y)%R

(- y <= x)%R
x, y:R
Hxy:(Rabs x <= y)%R
(x <= y)%R
x, y:R
Hxy:(Rabs x <= y)%R

(- y <= - Rabs x)%R
x, y:R
Hxy:(Rabs x <= y)%R
(- Rabs x <= x)%R
x, y:R
Hxy:(Rabs x <= y)%R
(x <= y)%R
x, y:R
Hxy:(Rabs x <= y)%R

(- Rabs x <= x)%R
x, y:R
Hxy:(Rabs x <= y)%R
(x <= y)%R
x, y:R
Hxy:(Rabs x <= y)%R

(- x <= - - Rabs x)%R
x, y:R
Hxy:(Rabs x <= y)%R
(x <= y)%R
x, y:R
Hxy:(Rabs x <= y)%R

(- x <= Rabs (- x))%R
x, y:R
Hxy:(Rabs x <= y)%R
(x <= y)%R
x, y:R
Hxy:(Rabs x <= y)%R

(x <= y)%R
x, y:R
Hxy:(Rabs x <= y)%R

(x <= Rabs x)%R
apply RRle_abs. Qed.

forall x y : R, (y <= - x)%R \/ (x <= y)%R -> (x <= Rabs y)%R

forall x y : R, (y <= - x)%R \/ (x <= y)%R -> (x <= Rabs y)%R
x, y:R
Hyx:(y <= - x)%R

(x <= Rabs y)%R
x, y:R
Hxy:(x <= y)%R
(x <= Rabs y)%R
x, y:R
Hyx:(y <= - x)%R

(x <= - y)%R
x, y:R
Hyx:(y <= - x)%R
(- y <= Rabs y)%R
x, y:R
Hxy:(x <= y)%R
(x <= Rabs y)%R
x, y:R
Hyx:(y <= - x)%R

(- - y <= - x)%R
x, y:R
Hyx:(y <= - x)%R
(- y <= Rabs y)%R
x, y:R
Hxy:(x <= y)%R
(x <= Rabs y)%R
x, y:R
Hyx:(y <= - x)%R

(- y <= Rabs y)%R
x, y:R
Hxy:(x <= y)%R
(x <= Rabs y)%R
x, y:R
Hyx:(y <= - x)%R

(- y <= Rabs (- y))%R
x, y:R
Hxy:(x <= y)%R
(x <= Rabs y)%R
x, y:R
Hxy:(x <= y)%R

(x <= Rabs y)%R
x, y:R
Hxy:(x <= y)%R

(y <= Rabs y)%R
apply RRle_abs. Qed.

forall x y : R, (x <= Rabs y)%R -> (y <= - x)%R \/ (x <= y)%R

forall x y : R, (x <= Rabs y)%R -> (y <= - x)%R \/ (x <= y)%R
x, y:R

(x <= Rabs y)%R -> (y <= - x)%R \/ (x <= y)%R
x, y:R

(x <= (if Rcase_abs y then - y else y))%R -> (y <= - x)%R \/ (x <= y)%R
x, y:R
Hy:(y < 0)%R
Hxy:(x <= - y)%R

(y <= - x)%R \/ (x <= y)%R
x, y:R
Hy:(y >= 0)%R
Hxy:(x <= y)%R
(y <= - x)%R \/ (x <= y)%R
x, y:R
Hy:(y < 0)%R
Hxy:(x <= - y)%R

(y <= - x)%R
x, y:R
Hy:(y >= 0)%R
Hxy:(x <= y)%R
(y <= - x)%R \/ (x <= y)%R
x, y:R
Hy:(y < 0)%R
Hxy:(x <= - y)%R

(- - x <= - y)%R
x, y:R
Hy:(y >= 0)%R
Hxy:(x <= y)%R
(y <= - x)%R \/ (x <= y)%R
x, y:R
Hy:(y >= 0)%R
Hxy:(x <= y)%R

(y <= - x)%R \/ (x <= y)%R
now right. Qed.

forall x y : R, (- y < x < y)%R -> (Rabs x < y)%R

forall x y : R, (- y < x < y)%R -> (Rabs x < y)%R
x, y:R
Hyx:(- y < x)%R
Hxy:(x < y)%R

(Rabs x < y)%R
now apply Rabs_def1. Qed.

forall x y : R, (Rabs x < y)%R -> (- y < x < y)%R

forall x y : R, (Rabs x < y)%R -> (- y < x < y)%R
x, y:R
H:(Rabs x < y)%R

(- y < x < y)%R
now split ; eapply Rabs_def2. Qed.

forall x y : R, (y < - x)%R \/ (x < y)%R -> (x < Rabs y)%R

forall x y : R, (y < - x)%R \/ (x < y)%R -> (x < Rabs y)%R
x, y:R
Hyx:(y < - x)%R

(x < Rabs y)%R
x, y:R
Hxy:(x < y)%R
(x < Rabs y)%R
x, y:R
Hyx:(y < - x)%R

(x < Rabs (- y))%R
x, y:R
Hxy:(x < y)%R
(x < Rabs y)%R
x, y:R
Hyx:(y < - x)%R

(x < - y)%R
x, y:R
Hyx:(y < - x)%R
(- y <= Rabs (- y))%R
x, y:R
Hxy:(x < y)%R
(x < Rabs y)%R
x, y:R
Hyx:(y < - x)%R

(- - y < - x)%R
x, y:R
Hyx:(y < - x)%R
(- y <= Rabs (- y))%R
x, y:R
Hxy:(x < y)%R
(x < Rabs y)%R
x, y:R
Hyx:(y < - x)%R

(- y <= Rabs (- y))%R
x, y:R
Hxy:(x < y)%R
(x < Rabs y)%R
x, y:R
Hxy:(x < y)%R

(x < Rabs y)%R
x, y:R
Hxy:(x < y)%R

(y <= Rabs y)%R
apply RRle_abs. Qed.

forall x y : R, (x < Rabs y)%R -> (y < - x)%R \/ (x < y)%R

forall x y : R, (x < Rabs y)%R -> (y < - x)%R \/ (x < y)%R
x, y:R

(x < Rabs y)%R -> (y < - x)%R \/ (x < y)%R
x, y:R

(x < (if Rcase_abs y then - y else y))%R -> (y < - x)%R \/ (x < y)%R
x, y:R
Hy:(y < 0)%R
Hxy:(x < - y)%R

(y < - x)%R \/ (x < y)%R
x, y:R
Hy:(y >= 0)%R
Hxy:(x < y)%R
(y < - x)%R \/ (x < y)%R
x, y:R
Hy:(y < 0)%R
Hxy:(x < - y)%R

(y < - x)%R
x, y:R
Hy:(y >= 0)%R
Hxy:(x < y)%R
(y < - x)%R \/ (x < y)%R
x, y:R
Hy:(y < 0)%R
Hxy:(x < - y)%R

(- - x < - y)%R
x, y:R
Hy:(y >= 0)%R
Hxy:(x < y)%R
(y < - x)%R \/ (x < y)%R
x, y:R
Hy:(y >= 0)%R
Hxy:(x < y)%R

(y < - x)%R \/ (x < y)%R
now right. Qed. End Rmissing. Section IZR.

forall m n p : Z, (m <= n < p)%Z -> (IZR m <= IZR n < IZR p)%R

forall m n p : Z, (m <= n < p)%Z -> (IZR m <= IZR n < IZR p)%R
m, n, p:Z
H1:(m <= n)%Z
H2:(n < p)%Z

(IZR m <= IZR n < IZR p)%R
m, n, p:Z
H1:(m <= n)%Z
H2:(n < p)%Z

(IZR m <= IZR n)%R
m, n, p:Z
H1:(m <= n)%Z
H2:(n < p)%Z
(IZR n < IZR p)%R
m, n, p:Z
H1:(m <= n)%Z
H2:(n < p)%Z

(IZR n < IZR p)%R
now apply IZR_lt. Qed.

forall m n p : Z, (IZR m <= IZR n < IZR p)%R -> (m <= n < p)%Z

forall m n p : Z, (IZR m <= IZR n < IZR p)%R -> (m <= n < p)%Z
m, n, p:Z
H1:(IZR m <= IZR n)%R
H2:(IZR n < IZR p)%R

(m <= n < p)%Z
m, n, p:Z
H1:(IZR m <= IZR n)%R
H2:(IZR n < IZR p)%R

(m <= n)%Z
m, n, p:Z
H1:(IZR m <= IZR n)%R
H2:(IZR n < IZR p)%R
(n < p)%Z
m, n, p:Z
H1:(IZR m <= IZR n)%R
H2:(IZR n < IZR p)%R

(n < p)%Z
now apply lt_IZR. Qed.

forall m n : Z, IZR m <> IZR n -> m <> n

forall m n : Z, IZR m <> IZR n -> m <> n
m, n:Z
H:IZR m <> IZR n
H':m = n

False
m, n:Z
H:IZR m <> IZR n
H':m = n

IZR m = IZR n
now apply f_equal. Qed. End IZR.
Decidable comparison on reals
Section Rcompare.

Definition Rcompare x y :=
  match total_order_T x y with
  | inleft (left _) => Lt
  | inleft (right _) => Eq
  | inright _ => Gt
  end.

Inductive Rcompare_prop (x y : R) : comparison -> Prop :=
  | Rcompare_Lt_ : (x < y)%R -> Rcompare_prop x y Lt
  | Rcompare_Eq_ : x = y -> Rcompare_prop x y Eq
  | Rcompare_Gt_ : (y < x)%R -> Rcompare_prop x y Gt.


forall x y : R, Rcompare_prop x y (Rcompare x y)

forall x y : R, Rcompare_prop x y (Rcompare x y)
x, y:R

Rcompare_prop x y (Rcompare x y)
x, y:R

Rcompare_prop x y match total_order_T x y with | inleft (left _) => Lt | inleft (right _) => Eq | inright _ => Gt end
now destruct (total_order_T x y) as [[H|H]|H] ; constructor. Qed. Global Opaque Rcompare.

forall x y : R, (x < y)%R -> Rcompare x y = Lt

forall x y : R, (x < y)%R -> Rcompare x y = Lt
x, y:R
H:(x < y)%R

Rcompare x y = Lt
x, y:R
H, H':(x < y)%R

Lt = Lt
x, y:R
H:(x < y)%R
H':x = y
Eq = Lt
x, y:R
H:(x < y)%R
H':(y < x)%R
Gt = Lt
x, y:R
H:(x < y)%R
H':x = y

Eq = Lt
x, y:R
H:(x < y)%R
H':(y < x)%R
Gt = Lt
x, y:R
H:(y < y)%R
H':x = y

Eq = Lt
x, y:R
H:(x < y)%R
H':(y < x)%R
Gt = Lt
x, y:R
H:(x < y)%R
H':(y < x)%R

Gt = Lt
x, y:R
H:(x < y)%R
H':(y < x)%R

(x < x)%R
now apply Rlt_trans with y. Qed.

forall x y : R, Rcompare x y = Lt -> (x < y)%R

forall x y : R, Rcompare x y = Lt -> (x < y)%R
x, y:R

Rcompare x y = Lt -> (x < y)%R
now case Rcompare_spec. Qed.

forall x y : R, (y <= x)%R -> Rcompare x y <> Lt

forall x y : R, (y <= x)%R -> Rcompare x y <> Lt
x, y:R
H1:(y <= x)%R
H2:Rcompare x y = Lt

False
x, y:R
H1:(y <= x)%R
H2:Rcompare x y = Lt

(x < y)%R
now apply Rcompare_Lt_inv. Qed.

forall x y : R, Rcompare x y <> Lt -> (y <= x)%R

forall x y : R, Rcompare x y <> Lt -> (y <= x)%R
x, y:R
H:Rcompare x y <> Lt

(y <= x)%R
x, y:R
H:Rcompare x y <> Lt

~ (x < y)%R
x, y:R
H:(x < y)%R

Rcompare x y = Lt
now apply Rcompare_Lt. Qed.

forall x y : R, x = y -> Rcompare x y = Eq

forall x y : R, x = y -> Rcompare x y = Eq
x, y:R
H:x = y

Rcompare x y = Eq
x, y:R
H:x = y

Rcompare y y = Eq
now case Rcompare_spec ; intro H' ; try elim (Rlt_irrefl _ H'). Qed.

forall x y : R, Rcompare x y = Eq -> x = y

forall x y : R, Rcompare x y = Eq -> x = y
x, y:R

Rcompare x y = Eq -> x = y
now case Rcompare_spec. Qed.

forall x y : R, (y < x)%R -> Rcompare x y = Gt

forall x y : R, (y < x)%R -> Rcompare x y = Gt
x, y:R
H:(y < x)%R

Rcompare x y = Gt
x, y:R
H:(y < x)%R
H':(x < y)%R

Lt = Gt
x, y:R
H:(y < x)%R
H':x = y
Eq = Gt
x, y:R
H, H':(y < x)%R
Gt = Gt
x, y:R
H:(y < x)%R
H':(x < y)%R

(x < x)%R
x, y:R
H:(y < x)%R
H':x = y
Eq = Gt
x, y:R
H, H':(y < x)%R
Gt = Gt
x, y:R
H:(y < x)%R
H':x = y

Eq = Gt
x, y:R
H, H':(y < x)%R
Gt = Gt
x, y:R
H:(y < y)%R
H':x = y

Eq = Gt
x, y:R
H, H':(y < x)%R
Gt = Gt
x, y:R
H, H':(y < x)%R

Gt = Gt
easy. Qed.

forall x y : R, Rcompare x y = Gt -> (y < x)%R

forall x y : R, Rcompare x y = Gt -> (y < x)%R
x, y:R

Rcompare x y = Gt -> (y < x)%R
now case Rcompare_spec. Qed.

forall x y : R, (x <= y)%R -> Rcompare x y <> Gt

forall x y : R, (x <= y)%R -> Rcompare x y <> Gt
x, y:R
H1:(x <= y)%R
H2:Rcompare x y = Gt

False
x, y:R
H1:(x <= y)%R
H2:Rcompare x y = Gt

(y < x)%R
now apply Rcompare_Gt_inv. Qed.

forall x y : R, Rcompare x y <> Gt -> (x <= y)%R

forall x y : R, Rcompare x y <> Gt -> (x <= y)%R
x, y:R
H:Rcompare x y <> Gt

(x <= y)%R
x, y:R
H:Rcompare x y <> Gt

~ (y < x)%R
x, y:R
H:(y < x)%R

Rcompare x y = Gt
now apply Rcompare_Gt. Qed.

forall x y : Z, Rcompare (IZR x) (IZR y) = (x ?= y)%Z

forall x y : Z, Rcompare (IZR x) (IZR y) = (x ?= y)%Z
x, y:Z

Rcompare (IZR x) (IZR y) = (x ?= y)%Z
x, y:Z
H:(IZR x < IZR y)%R

(x ?= y)%Z = Lt
x, y:Z
H:IZR x = IZR y
(x ?= y)%Z = Eq
x, y:Z
H:(IZR y < IZR x)%R
(x ?= y)%Z = Gt
x, y:Z
H:(IZR x < IZR y)%R

(x < y)%Z
x, y:Z
H:IZR x = IZR y
(x ?= y)%Z = Eq
x, y:Z
H:(IZR y < IZR x)%R
(x ?= y)%Z = Gt
x, y:Z
H:IZR x = IZR y

(x ?= y)%Z = Eq
x, y:Z
H:(IZR y < IZR x)%R
(x ?= y)%Z = Gt
x, y:Z
H:IZR x = IZR y

x = y
x, y:Z
H:(IZR y < IZR x)%R
(x ?= y)%Z = Gt
x, y:Z
H:(IZR y < IZR x)%R

(x ?= y)%Z = Gt
x, y:Z
H:(IZR y < IZR x)%R

(y < x)%Z
now apply lt_IZR. Qed.

forall x y : R, Rcompare x y = CompOpp (Rcompare y x)

forall x y : R, Rcompare x y = CompOpp (Rcompare y x)
x, y:R

Rcompare x y = CompOpp (Rcompare y x)
x, y:R
H:(y < x)%R

Rcompare x y = CompOpp Lt
x, y:R
H:y = x
Rcompare x y = CompOpp Eq
x, y:R
H:(x < y)%R
Rcompare x y = CompOpp Gt
x, y:R
H:y = x

Rcompare x y = CompOpp Eq
x, y:R
H:(x < y)%R
Rcompare x y = CompOpp Gt
x, y:R
H:(x < y)%R

Rcompare x y = CompOpp Gt
now apply Rcompare_Lt. Qed.

forall x y : R, Rcompare (- x) (- y) = Rcompare y x

forall x y : R, Rcompare (- x) (- y) = Rcompare y x
x, y:R

Rcompare (- x) (- y) = Rcompare y x
destruct (Rcompare_spec y x); destruct (Rcompare_spec (- x) (- y)); try reflexivity; exfalso; lra. Qed.

forall z x y : R, Rcompare (x + z) (y + z) = Rcompare x y

forall z x y : R, Rcompare (x + z) (y + z) = Rcompare x y
z, x, y:R

Rcompare (x + z) (y + z) = Rcompare x y
z, x, y:R
H:(x < y)%R

Rcompare (x + z) (y + z) = Lt
z, x, y:R
H:x = y
Rcompare (x + z) (y + z) = Eq
z, x, y:R
H:(y < x)%R
Rcompare (x + z) (y + z) = Gt
z, x, y:R
H:(x < y)%R

(x + z < y + z)%R
z, x, y:R
H:x = y
Rcompare (x + z) (y + z) = Eq
z, x, y:R
H:(y < x)%R
Rcompare (x + z) (y + z) = Gt
z, x, y:R
H:x = y

Rcompare (x + z) (y + z) = Eq
z, x, y:R
H:(y < x)%R
Rcompare (x + z) (y + z) = Gt
z, x, y:R
H:x = y

(x + z)%R = (y + z)%R
z, x, y:R
H:(y < x)%R
Rcompare (x + z) (y + z) = Gt
z, x, y:R
H:(y < x)%R

Rcompare (x + z) (y + z) = Gt
z, x, y:R
H:(y < x)%R

(y + z < x + z)%R
now apply Rplus_lt_compat_r. Qed.

forall z x y : R, Rcompare (z + x) (z + y) = Rcompare x y

forall z x y : R, Rcompare (z + x) (z + y) = Rcompare x y
z, x, y:R

Rcompare (z + x) (z + y) = Rcompare x y
z, x, y:R

Rcompare (x + z) (y + z) = Rcompare x y
apply Rcompare_plus_r. Qed.

forall z x y : R, (0 < z)%R -> Rcompare (x * z) (y * z) = Rcompare x y

forall z x y : R, (0 < z)%R -> Rcompare (x * z) (y * z) = Rcompare x y
z, x, y:R
Hz:(0 < z)%R

Rcompare (x * z) (y * z) = Rcompare x y
z, x, y:R
Hz:(0 < z)%R
H:(x < y)%R

Rcompare (x * z) (y * z) = Lt
z, x, y:R
Hz:(0 < z)%R
H:x = y
Rcompare (x * z) (y * z) = Eq
z, x, y:R
Hz:(0 < z)%R
H:(y < x)%R
Rcompare (x * z) (y * z) = Gt
z, x, y:R
Hz:(0 < z)%R
H:(x < y)%R

(x * z < y * z)%R
z, x, y:R
Hz:(0 < z)%R
H:x = y
Rcompare (x * z) (y * z) = Eq
z, x, y:R
Hz:(0 < z)%R
H:(y < x)%R
Rcompare (x * z) (y * z) = Gt
z, x, y:R
Hz:(0 < z)%R
H:x = y

Rcompare (x * z) (y * z) = Eq
z, x, y:R
Hz:(0 < z)%R
H:(y < x)%R
Rcompare (x * z) (y * z) = Gt
z, x, y:R
Hz:(0 < z)%R
H:x = y

(x * z)%R = (y * z)%R
z, x, y:R
Hz:(0 < z)%R
H:(y < x)%R
Rcompare (x * z) (y * z) = Gt
z, x, y:R
Hz:(0 < z)%R
H:(y < x)%R

Rcompare (x * z) (y * z) = Gt
z, x, y:R
Hz:(0 < z)%R
H:(y < x)%R

(y * z < x * z)%R
now apply Rmult_lt_compat_r. Qed.

forall z x y : R, (0 < z)%R -> Rcompare (z * x) (z * y) = Rcompare x y

forall z x y : R, (0 < z)%R -> Rcompare (z * x) (z * y) = Rcompare x y
z, x, y:R

(0 < z)%R -> Rcompare (z * x) (z * y) = Rcompare x y
z, x, y:R

(0 < z)%R -> Rcompare (x * z) (y * z) = Rcompare x y
apply Rcompare_mult_r. Qed.

forall x d u : R, Rcompare (x - d) (u - x) = Rcompare x ((d + u) / 2)

forall x d u : R, Rcompare (x - d) (u - x) = Rcompare x ((d + u) / 2)
x, d, u:R

Rcompare (x - d) (u - x) = Rcompare x ((d + u) / 2)
x, d, u:R

Rcompare (x - d) (u - x) = Rcompare (x + (- x / 2 - d / 2)) ((d + u) / 2 + (- x / 2 - d / 2))
x, d, u:R

Rcompare ((x - d) * / 2) ((u - x) * / 2) = Rcompare (x + (- x / 2 - d / 2)) ((d + u) / 2 + (- x / 2 - d / 2))
x, d, u:R
(0 < / 2)%R
x, d, u:R

Rcompare ((x - d) * / 2) ((u - x) * / 2) = Rcompare ((x - d) / 2) ((d + u) / 2 + (- x / 2 - d / 2))
x, d, u:R
(0 < / 2)%R
x, d, u:R

(0 < / 2)%R
x, d, u:R

(0 < 2)%R
now apply IZR_lt. Qed.

forall x y : R, Rcompare (x / 2) y = Rcompare x (2 * y)

forall x y : R, Rcompare (x / 2) y = Rcompare x (2 * y)
x, y:R

Rcompare (x / 2) y = Rcompare x (2 * y)
x, y:R

Rcompare (x / 2 * 2) (y * 2) = Rcompare x (2 * y)
x, y:R
(0 < 2)%R
x, y:R

Rcompare (x * / 2 * 2) (y * 2) = Rcompare x (2 * y)
x, y:R
(0 < 2)%R
x, y:R

Rcompare x (y * 2) = Rcompare x (2 * y)
x, y:R
2%R <> 0%R
x, y:R
(0 < 2)%R
x, y:R

2%R <> 0%R
x, y:R
(0 < 2)%R
x, y:R

(0 < 2)%R
now apply IZR_lt. Qed.

forall x y : R, Rcompare x (y / 2) = Rcompare (2 * x) y

forall x y : R, Rcompare x (y / 2) = Rcompare (2 * x) y
x, y:R

Rcompare x (y / 2) = Rcompare (2 * x) y
x, y:R

Rcompare (x * 2) (y / 2 * 2) = Rcompare (2 * x) y
x, y:R
(0 < 2)%R
x, y:R

Rcompare (x * 2) (y * / 2 * 2) = Rcompare (2 * x) y
x, y:R
(0 < 2)%R
x, y:R

Rcompare (x * 2) y = Rcompare (2 * x) y
x, y:R
2%R <> 0%R
x, y:R
(0 < 2)%R
x, y:R

2%R <> 0%R
x, y:R
(0 < 2)%R
x, y:R

(0 < 2)%R
now apply IZR_lt. Qed.

forall x y : R, Rcompare (x * x) (y * y) = Rcompare (Rabs x) (Rabs y)

forall x y : R, Rcompare (x * x) (y * y) = Rcompare (Rabs x) (Rabs y)
x, y:R

Rcompare (x * x) (y * y) = Rcompare (Rabs x) (Rabs y)
x, y:R
H:(Rabs x < Rabs y)%R

Rcompare (x * x) (y * y) = Lt
x, y:R
H:Rabs x = Rabs y
Rcompare (x * x) (y * y) = Eq
x, y:R
H:(Rabs y < Rabs x)%R
Rcompare (x * x) (y * y) = Gt
x, y:R
H:(Rabs x < Rabs y)%R

(x * x < y * y)%R
x, y:R
H:Rabs x = Rabs y
Rcompare (x * x) (y * y) = Eq
x, y:R
H:(Rabs y < Rabs x)%R
Rcompare (x * x) (y * y) = Gt
x, y:R
H:Rabs x = Rabs y

Rcompare (x * x) (y * y) = Eq
x, y:R
H:(Rabs y < Rabs x)%R
Rcompare (x * x) (y * y) = Gt
x, y:R
H:Rabs x = Rabs y

Rcompare x² y² = Eq
x, y:R
H:(Rabs y < Rabs x)%R
Rcompare (x * x) (y * y) = Gt
x, y:R
H:Rabs x = Rabs y

Rcompare (Rabs y)² (Rabs y)² = Eq
x, y:R
H:(Rabs y < Rabs x)%R
Rcompare (x * x) (y * y) = Gt
x, y:R
H:(Rabs y < Rabs x)%R

Rcompare (x * x) (y * y) = Gt
x, y:R
H:(Rabs y < Rabs x)%R

(y * y < x * x)%R
now apply Rsqr_lt_abs_1. Qed.

forall x y : R, Rmin x y = match Rcompare x y with | Gt => y | _ => x end

forall x y : R, Rmin x y = match Rcompare x y with | Gt => y | _ => x end
x, y:R

Rmin x y = match Rcompare x y with | Gt => y | _ => x end
x, y:R

(if Rle_dec x y then x else y) = match Rcompare x y with | Gt => y | _ => x end
x, y:R
Hx:(x < y)%R

x = match Rcompare x y with | Gt => y | _ => x end
x, y:R
Hx:x = y
x = match Rcompare x y with | Gt => y | _ => x end
x, y:R
Hx:~ (x <= y)%R
y = match Rcompare x y with | Gt => y | _ => x end
x, y:R
Hx:x = y

x = match Rcompare x y with | Gt => y | _ => x end
x, y:R
Hx:~ (x <= y)%R
y = match Rcompare x y with | Gt => y | _ => x end
x, y:R
Hx:~ (x <= y)%R

y = match Rcompare x y with | Gt => y | _ => x end
x, y:R
Hx:~ (x <= y)%R

y = y
x, y:R
Hx:~ (x <= y)%R
(y < x)%R
x, y:R
Hx:~ (x <= y)%R

(y < x)%R
now apply Rnot_le_lt. Qed. End Rcompare. Section Rle_bool. Definition Rle_bool x y := match Rcompare x y with | Gt => false | _ => true end. Inductive Rle_bool_prop (x y : R) : bool -> Prop := | Rle_bool_true_ : (x <= y)%R -> Rle_bool_prop x y true | Rle_bool_false_ : (y < x)%R -> Rle_bool_prop x y false.

forall x y : R, Rle_bool_prop x y (Rle_bool x y)

forall x y : R, Rle_bool_prop x y (Rle_bool x y)
x, y:R

Rle_bool_prop x y (Rle_bool x y)
x, y:R

Rle_bool_prop x y match Rcompare x y with | Gt => false | _ => true end
x, y:R
H:(x < y)%R

(x <= y)%R
x, y:R
H:x = y
(x <= y)%R
x, y:R
H:(y < x)%R
(y < x)%R
x, y:R
H:x = y

(x <= y)%R
x, y:R
H:(y < x)%R
(y < x)%R
x, y:R
H:x = y

(y <= y)%R
x, y:R
H:(y < x)%R
(y < x)%R
x, y:R
H:(y < x)%R

(y < x)%R
exact H. Qed.

forall x y : R, (x <= y)%R -> Rle_bool x y = true

forall x y : R, (x <= y)%R -> Rle_bool x y = true
x, y:R
Hxy:(x <= y)%R

Rle_bool x y = true
x, y:R
Hxy, H:(x <= y)%R

true = true
x, y:R
Hxy:(x <= y)%R
H:(y < x)%R
false = true
x, y:R
Hxy:(x <= y)%R
H:(y < x)%R

false = true
x, y:R
Hxy:(x <= y)%R
H:(y < x)%R

(x < x)%R
now apply Rle_lt_trans with y. Qed.

forall x y : R, (y < x)%R -> Rle_bool x y = false

forall x y : R, (y < x)%R -> Rle_bool x y = false
x, y:R
Hxy:(y < x)%R

Rle_bool x y = false
x, y:R
Hxy:(y < x)%R
H:(x <= y)%R

true = false
x, y:R
Hxy, H:(y < x)%R
false = false
x, y:R
Hxy:(y < x)%R
H:(x <= y)%R

(x < x)%R
x, y:R
Hxy, H:(y < x)%R
false = false
x, y:R
Hxy, H:(y < x)%R

false = false
apply refl_equal. Qed. End Rle_bool. Section Rlt_bool. Definition Rlt_bool x y := match Rcompare x y with | Lt => true | _ => false end. Inductive Rlt_bool_prop (x y : R) : bool -> Prop := | Rlt_bool_true_ : (x < y)%R -> Rlt_bool_prop x y true | Rlt_bool_false_ : (y <= x)%R -> Rlt_bool_prop x y false.

forall x y : R, Rlt_bool_prop x y (Rlt_bool x y)

forall x y : R, Rlt_bool_prop x y (Rlt_bool x y)
x, y:R

Rlt_bool_prop x y (Rlt_bool x y)
x, y:R

Rlt_bool_prop x y match Rcompare x y with | Lt => true | _ => false end
x, y:R
H:(x < y)%R

(x < y)%R
x, y:R
H:x = y
(y <= x)%R
x, y:R
H:(y < x)%R
(y <= x)%R
x, y:R
H:x = y

(y <= x)%R
x, y:R
H:(y < x)%R
(y <= x)%R
x, y:R
H:x = y

(y <= y)%R
x, y:R
H:(y < x)%R
(y <= x)%R
x, y:R
H:(y < x)%R

(y <= x)%R
now apply Rlt_le. Qed.

forall x y : R, negb (Rle_bool x y) = Rlt_bool y x

forall x y : R, negb (Rle_bool x y) = Rlt_bool y x
x, y:R

negb (Rle_bool x y) = Rlt_bool y x
x, y:R

negb match Rcompare x y with | Gt => false | _ => true end = match Rcompare y x with | Lt => true | _ => false end
x, y:R

negb match CompOpp (Rcompare y x) with | Gt => false | _ => true end = match Rcompare y x with | Lt => true | _ => false end
now case Rcompare. Qed.

forall x y : R, negb (Rlt_bool x y) = Rle_bool y x

forall x y : R, negb (Rlt_bool x y) = Rle_bool y x
x, y:R

negb (Rlt_bool x y) = Rle_bool y x
x, y:R

negb match Rcompare x y with | Lt => true | _ => false end = match Rcompare y x with | Gt => false | _ => true end
x, y:R

negb match CompOpp (Rcompare y x) with | Lt => true | _ => false end = match Rcompare y x with | Gt => false | _ => true end
now case Rcompare. Qed.

forall x y : R, (x < y)%R -> Rlt_bool x y = true

forall x y : R, (x < y)%R -> Rlt_bool x y = true
x, y:R
Hxy:(x < y)%R

Rlt_bool x y = true
x, y:R
Hxy:(x < y)%R

negb (Rle_bool y x) = true
now rewrite Rle_bool_false. Qed.

forall x y : R, (y <= x)%R -> Rlt_bool x y = false

forall x y : R, (y <= x)%R -> Rlt_bool x y = false
x, y:R
Hxy:(y <= x)%R

Rlt_bool x y = false
x, y:R
Hxy:(y <= x)%R

negb (Rle_bool y x) = false
now rewrite Rle_bool_true. Qed.

forall x y : R, Rlt_bool (- x) (- y) = Rlt_bool y x

forall x y : R, Rlt_bool (- x) (- y) = Rlt_bool y x
x, y:R

Rlt_bool (- x) (- y) = Rlt_bool y x
now unfold Rlt_bool; rewrite Rcompare_opp. Qed. End Rlt_bool. Section Req_bool. Definition Req_bool x y := match Rcompare x y with | Eq => true | _ => false end. Inductive Req_bool_prop (x y : R) : bool -> Prop := | Req_bool_true_ : (x = y)%R -> Req_bool_prop x y true | Req_bool_false_ : (x <> y)%R -> Req_bool_prop x y false.

forall x y : R, Req_bool_prop x y (Req_bool x y)

forall x y : R, Req_bool_prop x y (Req_bool x y)
x, y:R

Req_bool_prop x y (Req_bool x y)
x, y:R

Req_bool_prop x y match Rcompare x y with | Eq => true | _ => false end
x, y:R
H:(x < y)%R

x <> y
x, y:R
H:x = y
x = y
x, y:R
H:(y < x)%R
x <> y
x, y:R
H:x = y

x = y
x, y:R
H:(y < x)%R
x <> y
x, y:R
H:(y < x)%R

x <> y
now apply Rgt_not_eq. Qed.

forall x y : R, x = y -> Req_bool x y = true

forall x y : R, x = y -> Req_bool x y = true
x, y:R
Hxy:x = y

Req_bool x y = true
x, y:R
Hxy, H:x = y

true = true
x, y:R
Hxy:x = y
H:x <> y
false = true
x, y:R
Hxy:x = y
H:x <> y

false = true
x, y:R
Hxy:x = y

x = y
exact Hxy. Qed.

forall x y : R, x <> y -> Req_bool x y = false

forall x y : R, x <> y -> Req_bool x y = false
x, y:R
Hxy:x <> y

Req_bool x y = false
x, y:R
Hxy:x <> y
H:x = y

true = false
x, y:R
Hxy, H:x <> y
false = false
x, y:R
H:x = y

x = y
x, y:R
Hxy, H:x <> y
false = false
x, y:R
Hxy, H:x <> y

false = false
apply refl_equal. Qed. End Req_bool. Section Floor_Ceil.
Zfloor and Zceil
Definition Zfloor (x : R) := (up x - 1)%Z.


forall x : R, (IZR (Zfloor x) <= x)%R

forall x : R, (IZR (Zfloor x) <= x)%R
x:R

(IZR (Zfloor x) <= x)%R
x:R

(IZR (up x - 1) <= x)%R
x:R

(IZR (up x) - 1 <= x)%R
x:R

(IZR (up x) - 1 <= x)%R
x:R

(IZR (up x) - 1 + (1 - x) <= x + (1 - x))%R
x:R

(IZR (up x) - x <= 1)%R
exact (proj2 (archimed x)). Qed.

forall x : R, (x < IZR (Zfloor x) + 1)%R

forall x : R, (x < IZR (Zfloor x) + 1)%R
x:R

(x < IZR (Zfloor x) + 1)%R
x:R

(x < IZR (up x - 1) + 1)%R
x:R

(x < IZR (up x) - 1 + 1)%R
x:R

(x < IZR (up x) + - (1) + 1)%R
x:R

(x < IZR (up x) + (- (1) + 1))%R
x:R

(x < IZR (up x))%R
exact (proj1 (archimed x)). Qed.

forall (n : Z) (x : R), (IZR n <= x)%R -> (n <= Zfloor x)%Z

forall (n : Z) (x : R), (IZR n <= x)%R -> (n <= Zfloor x)%Z
n:Z
x:R
Hnx:(IZR n <= x)%R

(n <= Zfloor x)%Z
n:Z
x:R
Hnx:(IZR n <= x)%R

(n < Z.succ (Zfloor x))%Z
n:Z
x:R
Hnx:(IZR n <= x)%R

(IZR n < IZR (Z.succ (Zfloor x)))%R
n:Z
x:R
Hnx:(IZR n <= x)%R

(x < IZR (Z.succ (Zfloor x)))%R
n:Z
x:R
Hnx:(IZR n <= x)%R

(x < IZR (Zfloor x + 1))%R
n:Z
x:R
Hnx:(IZR n <= x)%R

(x < IZR (Zfloor x) + 1)%R
apply Zfloor_ub. Qed.

forall (n : Z) (x : R), (IZR n <= x < IZR (n + 1))%R -> Zfloor x = n

forall (n : Z) (x : R), (IZR n <= x < IZR (n + 1))%R -> Zfloor x = n
n:Z
x:R
Hnx:(IZR n <= x < IZR (n + 1))%R

Zfloor x = n
n:Z
x:R
Hnx:(IZR n <= x < IZR (n + 1))%R

(Zfloor x <= n)%Z
n:Z
x:R
Hnx:(IZR n <= x < IZR (n + 1))%R
(n <= Zfloor x)%Z
n:Z
x:R
Hnx:(IZR n <= x < IZR (n + 1))%R

(Zfloor x < Z.succ n)%Z
n:Z
x:R
Hnx:(IZR n <= x < IZR (n + 1))%R
(n <= Zfloor x)%Z
n:Z
x:R
Hnx:(IZR n <= x < IZR (n + 1))%R

(IZR (Zfloor x) < IZR (Z.succ n))%R
n:Z
x:R
Hnx:(IZR n <= x < IZR (n + 1))%R
(n <= Zfloor x)%Z
n:Z
x:R
Hnx:(IZR n <= x < IZR (n + 1))%R

(IZR (Zfloor x) <= x)%R
n:Z
x:R
Hnx:(IZR n <= x < IZR (n + 1))%R
(n <= Zfloor x)%Z
n:Z
x:R
Hnx:(IZR n <= x < IZR (n + 1))%R

(n <= Zfloor x)%Z
now apply Zfloor_lub. Qed.

forall n : Z, Zfloor (IZR n) = n

forall n : Z, Zfloor (IZR n) = n
n:Z

Zfloor (IZR n) = n
n:Z

(IZR n <= IZR n < IZR (n + 1))%R
n:Z

(IZR n <= IZR n)%R
n:Z
(IZR n < IZR (n + 1))%R
n:Z

(IZR n < IZR (n + 1))%R
n:Z

(n < n + 1)%Z
apply Zlt_succ. Qed.

forall x y : R, (x <= y)%R -> (Zfloor x <= Zfloor y)%Z

forall x y : R, (x <= y)%R -> (Zfloor x <= Zfloor y)%Z
x, y:R
Hxy:(x <= y)%R

(Zfloor x <= Zfloor y)%Z
x, y:R
Hxy:(x <= y)%R

(IZR (Zfloor x) <= y)%R
x, y:R
Hxy:(x <= y)%R

(IZR (Zfloor x) <= x)%R
apply Zfloor_lb. Qed. Definition Zceil (x : R) := (- Zfloor (- x))%Z.

forall x : R, (x <= IZR (Zceil x))%R

forall x : R, (x <= IZR (Zceil x))%R
x:R

(x <= IZR (Zceil x))%R
x:R

(x <= IZR (- Zfloor (- x)))%R
x:R

(x <= - IZR (Zfloor (- x)))%R
x:R

(- - IZR (Zfloor (- x)) <= - x)%R
x:R

(IZR (Zfloor (- x)) <= - x)%R
apply Zfloor_lb. Qed.

forall x : R, (IZR (Zceil x) < x + 1)%R

forall x : R, (IZR (Zceil x) < x + 1)%R
x:R

(IZR (Zceil x) < x + 1)%R
x:R

(IZR (- Zfloor (- x)) < x + 1)%R
x:R

(- IZR (Zfloor (- x)) < x + 1)%R
x:R

(- IZR (Zfloor (- x)) < - (- x + - (1)))%R
x:R

(- x < IZR (Zfloor (- x)) + 1)%R
apply Zfloor_ub. Qed.

forall (n : Z) (x : R), (x <= IZR n)%R -> (Zceil x <= n)%Z

forall (n : Z) (x : R), (x <= IZR n)%R -> (Zceil x <= n)%Z
n:Z
x:R
Hnx:(x <= IZR n)%R

(Zceil x <= n)%Z
n:Z
x:R
Hnx:(x <= IZR n)%R

(- Zfloor (- x) <= n)%Z
n:Z
x:R
Hnx:(x <= IZR n)%R

(- n <= - - Zfloor (- x))%Z
n:Z
x:R
Hnx:(x <= IZR n)%R

(- n <= Zfloor (- x))%Z
n:Z
x:R
Hnx:(x <= IZR n)%R

(IZR (- n) <= - x)%R
n:Z
x:R
Hnx:(x <= IZR n)%R

(- IZR n <= - x)%R
now apply Ropp_le_contravar. Qed.

forall (n : Z) (x : R), (IZR (n - 1) < x <= IZR n)%R -> Zceil x = n

forall (n : Z) (x : R), (IZR (n - 1) < x <= IZR n)%R -> Zceil x = n
n:Z
x:R
Hnx:(IZR (n - 1) < x <= IZR n)%R

Zceil x = n
n:Z
x:R
Hnx:(IZR (n - 1) < x <= IZR n)%R

(- Zfloor (- x))%Z = n
n:Z
x:R
Hnx:(IZR (n - 1) < x <= IZR n)%R

(- Zfloor (- x))%Z = (- - n)%Z
n:Z
x:R
Hnx:(IZR (n - 1) < x <= IZR n)%R

Zfloor (- x) = (- n)%Z
n:Z
x:R
Hnx:(IZR (n - 1) < x <= IZR n)%R

(IZR (- n) <= - x < IZR (- n + 1))%R
n:Z
x:R
Hnx:(IZR (n - 1) < x <= IZR n)%R

(IZR (- n) <= - x)%R
n:Z
x:R
Hnx:(IZR (n - 1) < x <= IZR n)%R
(- x < IZR (- n + 1))%R
n:Z
x:R
Hnx:(IZR (n - 1) < x <= IZR n)%R

(- IZR n <= - x)%R
n:Z
x:R
Hnx:(IZR (n - 1) < x <= IZR n)%R
(- x < IZR (- n + 1))%R
n:Z
x:R
Hnx:(IZR (n - 1) < x <= IZR n)%R

(- x < IZR (- n + 1))%R
n:Z
x:R
Hnx:(IZR (n - 1) < x <= IZR n)%R

(- x < IZR (- n + - - (1)))%R
n:Z
x:R
Hnx:(IZR (n - 1) < x <= IZR n)%R

(- x < IZR (- (n + - (1))))%R
n:Z
x:R
Hnx:(IZR (n - 1) < x <= IZR n)%R

(- x < - IZR (n + - (1)))%R
now apply Ropp_lt_contravar. Qed.

forall n : Z, Zceil (IZR n) = n

forall n : Z, Zceil (IZR n) = n
n:Z

Zceil (IZR n) = n
n:Z

(- Zfloor (- IZR n))%Z = n
n:Z

(- - n)%Z = n
apply Z.opp_involutive. Qed.

forall x y : R, (x <= y)%R -> (Zceil x <= Zceil y)%Z

forall x y : R, (x <= y)%R -> (Zceil x <= Zceil y)%Z
x, y:R
Hxy:(x <= y)%R

(Zceil x <= Zceil y)%Z
x, y:R
Hxy:(x <= y)%R

(x <= IZR (Zceil y))%R
x, y:R
Hxy:(x <= y)%R

(y <= IZR (Zceil y))%R
apply Zceil_ub. Qed.

forall x : R, IZR (Zfloor x) <> x -> Zceil x = (Zfloor x + 1)%Z

forall x : R, IZR (Zfloor x) <> x -> Zceil x = (Zfloor x + 1)%Z
x:R
Hx:IZR (Zfloor x) <> x

Zceil x = (Zfloor x + 1)%Z
x:R
Hx:IZR (Zfloor x) <> x

(IZR (Zfloor x + 1 - 1) < x <= IZR (Zfloor x + 1))%R
x:R
Hx:IZR (Zfloor x) <> x

(IZR (Zfloor x + 1 - 1) < x)%R
x:R
Hx:IZR (Zfloor x) <> x
(x <= IZR (Zfloor x + 1))%R
x:R
Hx:IZR (Zfloor x) <> x

(IZR (Zfloor x) < x)%R
x:R
Hx:IZR (Zfloor x) <> x
(x <= IZR (Zfloor x + 1))%R
x:R
Hx:IZR (Zfloor x) <> x

~ (x <= IZR (Zfloor x))%R
x:R
Hx:IZR (Zfloor x) <> x
(x <= IZR (Zfloor x + 1))%R
x:R
Hx:IZR (Zfloor x) <> x
H:(x <= IZR (Zfloor x))%R

False
x:R
Hx:IZR (Zfloor x) <> x
(x <= IZR (Zfloor x + 1))%R
x:R
Hx:IZR (Zfloor x) <> x
H:(x <= IZR (Zfloor x))%R

IZR (Zfloor x) = x
x:R
Hx:IZR (Zfloor x) <> x
(x <= IZR (Zfloor x + 1))%R
x:R
Hx:IZR (Zfloor x) <> x
H:(x <= IZR (Zfloor x))%R

(IZR (Zfloor x) <= x)%R
x:R
Hx:IZR (Zfloor x) <> x
H:(x <= IZR (Zfloor x))%R
(x <= IZR (Zfloor x))%R
x:R
Hx:IZR (Zfloor x) <> x
(x <= IZR (Zfloor x + 1))%R
x:R
Hx:IZR (Zfloor x) <> x
H:(x <= IZR (Zfloor x))%R

(x <= IZR (Zfloor x))%R
x:R
Hx:IZR (Zfloor x) <> x
(x <= IZR (Zfloor x + 1))%R
x:R
Hx:IZR (Zfloor x) <> x

(x <= IZR (Zfloor x + 1))%R
x:R
Hx:IZR (Zfloor x) <> x

(x < IZR (Zfloor x + 1))%R
x:R
Hx:IZR (Zfloor x) <> x

(x < IZR (Zfloor x) + 1)%R
apply Zfloor_ub. Qed. Definition Ztrunc x := if Rlt_bool x 0 then Zceil x else Zfloor x.

forall n : Z, Ztrunc (IZR n) = n

forall n : Z, Ztrunc (IZR n) = n
n:Z

Ztrunc (IZR n) = n
n:Z

(if Rlt_bool (IZR n) 0 then Zceil (IZR n) else Zfloor (IZR n)) = n
n:Z
H:(IZR n < 0)%R

Zceil (IZR n) = n
n:Z
H:(0 <= IZR n)%R
Zfloor (IZR n) = n
n:Z
H:(0 <= IZR n)%R

Zfloor (IZR n) = n
apply Zfloor_IZR. Qed.

forall x : R, (0 <= x)%R -> Ztrunc x = Zfloor x

forall x : R, (0 <= x)%R -> Ztrunc x = Zfloor x
x:R
Hx:(0 <= x)%R

Ztrunc x = Zfloor x
x:R
Hx:(0 <= x)%R

(if Rlt_bool x 0 then Zceil x else Zfloor x) = Zfloor x
x:R
Hx:(0 <= x)%R
H:(x < 0)%R

Zceil x = Zfloor x
x:R
Hx, H:(0 <= x)%R
Zfloor x = Zfloor x
x:R
Hx:(0 <= x)%R
H:(x < 0)%R

(x < x)%R
x:R
Hx, H:(0 <= x)%R
Zfloor x = Zfloor x
x:R
Hx, H:(0 <= x)%R

Zfloor x = Zfloor x
apply refl_equal. Qed.

forall x : R, (x <= 0)%R -> Ztrunc x = Zceil x

forall x : R, (x <= 0)%R -> Ztrunc x = Zceil x
x:R
Hx:(x <= 0)%R

Ztrunc x = Zceil x
x:R
Hx:(x <= 0)%R

(if Rlt_bool x 0 then Zceil x else Zfloor x) = Zceil x
x:R
Hx:(x <= 0)%R
H:(x < 0)%R

Zceil x = Zceil x
x:R
Hx:(x <= 0)%R
H:(0 <= x)%R
Zfloor x = Zceil x
x:R
Hx:(x <= 0)%R
H:(0 <= x)%R

Zfloor x = Zceil x
x:R
Hx:(x <= 0)%R
H:(0 <= x)%R

Zfloor 0 = Zceil 0
x:R
Hx:(x <= 0)%R
H:(0 <= x)%R

Zfloor 0 = 0%Z
apply Zfloor_IZR. Qed.

forall x y : R, (x <= y)%R -> (Ztrunc x <= Ztrunc y)%Z

forall x y : R, (x <= y)%R -> (Ztrunc x <= Ztrunc y)%Z
x, y:R
Hxy:(x <= y)%R

(Ztrunc x <= Ztrunc y)%Z
x, y:R
Hxy:(x <= y)%R

((if Rlt_bool x 0 then Zceil x else Zfloor x) <= Ztrunc y)%Z
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R

(Zceil x <= Ztrunc y)%Z
x, y:R
Hxy:(x <= y)%R
Hx:(0 <= x)%R
(Zfloor x <= Ztrunc y)%Z
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R

(Zceil x <= (if Rlt_bool y 0 then Zceil y else Zfloor y))%Z
x, y:R
Hxy:(x <= y)%R
Hx:(0 <= x)%R
(Zfloor x <= Ztrunc y)%Z
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(y < 0)%R

(Zceil x <= Zceil y)%Z
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R
(Zceil x <= Zfloor y)%Z
x, y:R
Hxy:(x <= y)%R
Hx:(0 <= x)%R
(Zfloor x <= Ztrunc y)%Z
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R

(Zceil x <= Zfloor y)%Z
x, y:R
Hxy:(x <= y)%R
Hx:(0 <= x)%R
(Zfloor x <= Ztrunc y)%Z
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R

(Zceil x <= 0)%Z
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R
(0 <= Zfloor y)%Z
x, y:R
Hxy:(x <= y)%R
Hx:(0 <= x)%R
(Zfloor x <= Ztrunc y)%Z
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R

(x <= 0)%R
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R
(0 <= Zfloor y)%Z
x, y:R
Hxy:(x <= y)%R
Hx:(0 <= x)%R
(Zfloor x <= Ztrunc y)%Z
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R

(0 <= Zfloor y)%Z
x, y:R
Hxy:(x <= y)%R
Hx:(0 <= x)%R
(Zfloor x <= Ztrunc y)%Z
x, y:R
Hxy:(x <= y)%R
Hx:(0 <= x)%R

(Zfloor x <= Ztrunc y)%Z
x, y:R
Hxy:(x <= y)%R
Hx:(0 <= x)%R

(Zfloor x <= Zfloor y)%Z
x, y:R
Hxy:(x <= y)%R
Hx:(0 <= x)%R
(0 <= y)%R
x, y:R
Hxy:(x <= y)%R
Hx:(0 <= x)%R

(0 <= y)%R
now apply Rle_trans with x. Qed.

forall x : R, Ztrunc (- x) = (- Ztrunc x)%Z

forall x : R, Ztrunc (- x) = (- Ztrunc x)%Z
x:R

Ztrunc (- x) = (- Ztrunc x)%Z
x:R

Ztrunc (- x) = (- (if Rlt_bool x 0 then Zceil x else Zfloor x))%Z
x:R
Hx:(x < 0)%R

Ztrunc (- x) = (- Zceil x)%Z
x:R
Hx:(0 <= x)%R
Ztrunc (- x) = (- Zfloor x)%Z
x:R
Hx:(x < 0)%R

Zfloor (- x) = (- Zceil x)%Z
x:R
Hx:(x < 0)%R
(0 <= - x)%R
x:R
Hx:(0 <= x)%R
Ztrunc (- x) = (- Zfloor x)%Z
x:R
Hx:(x < 0)%R

(- Zceil x)%Z = Zfloor (- x)
x:R
Hx:(x < 0)%R
(0 <= - x)%R
x:R
Hx:(0 <= x)%R
Ztrunc (- x) = (- Zfloor x)%Z
x:R
Hx:(x < 0)%R

(0 <= - x)%R
x:R
Hx:(0 <= x)%R
Ztrunc (- x) = (- Zfloor x)%Z
x:R
Hx:(x < 0)%R

(- 0 <= - x)%R
x:R
Hx:(0 <= x)%R
Ztrunc (- x) = (- Zfloor x)%Z
x:R
Hx:(x < 0)%R

(x <= 0)%R
x:R
Hx:(0 <= x)%R
Ztrunc (- x) = (- Zfloor x)%Z
x:R
Hx:(0 <= x)%R

Ztrunc (- x) = (- Zfloor x)%Z
x:R
Hx:(0 <= x)%R

Zceil (- x) = (- Zfloor x)%Z
x:R
Hx:(0 <= x)%R
(- x <= 0)%R
x:R
Hx:(0 <= x)%R

(- Zfloor (- - x))%Z = (- Zfloor x)%Z
x:R
Hx:(0 <= x)%R
(- x <= 0)%R
x:R
Hx:(0 <= x)%R

(- x <= 0)%R
x:R
Hx:(0 <= x)%R

(- x <= - 0)%R
now apply Ropp_le_contravar. Qed.

forall x : R, Ztrunc (Rabs x) = Z.abs (Ztrunc x)

forall x : R, Ztrunc (Rabs x) = Z.abs (Ztrunc x)
x:R

Ztrunc (Rabs x) = Z.abs (Ztrunc x)
x:R

Zfloor (Rabs x) = Z.abs (Ztrunc x)
x:R
(0 <= Rabs x)%R
x:R

Zfloor (Rabs x) = Z.abs (Ztrunc x)
x:R

Zfloor (Rabs x) = Z.abs (if Rlt_bool x 0 then Zceil x else Zfloor x)
x:R
H:(x < 0)%R

Zfloor (Rabs x) = Z.abs (Zceil x)
x:R
H:(0 <= x)%R
Zfloor (Rabs x) = Z.abs (Zfloor x)
x:R
H:(x < 0)%R

Zfloor (- x) = Z.abs (Zceil x)
x:R
H:(0 <= x)%R
Zfloor (Rabs x) = Z.abs (Zfloor x)
x:R
H:(x < 0)%R

Zfloor (- x) = (- Zceil x)%Z
x:R
H:(x < 0)%R
(Zceil x <= 0)%Z
x:R
H:(0 <= x)%R
Zfloor (Rabs x) = Z.abs (Zfloor x)
x:R
H:(x < 0)%R

(- Zceil x)%Z = Zfloor (- x)
x:R
H:(x < 0)%R
(Zceil x <= 0)%Z
x:R
H:(0 <= x)%R
Zfloor (Rabs x) = Z.abs (Zfloor x)
x:R
H:(x < 0)%R

(Zceil x <= 0)%Z
x:R
H:(0 <= x)%R
Zfloor (Rabs x) = Z.abs (Zfloor x)
x:R
H:(x < 0)%R

(x <= 0)%R
x:R
H:(0 <= x)%R
Zfloor (Rabs x) = Z.abs (Zfloor x)
x:R
H:(0 <= x)%R

Zfloor (Rabs x) = Z.abs (Zfloor x)
x:R
H:(0 <= x)%R

Zfloor x = Z.abs (Zfloor x)
x:R
H:(0 <= x)%R

Z.abs (Zfloor x) = Zfloor x
x:R
H:(0 <= x)%R

(0 <= Zfloor x)%Z
now apply Zfloor_lub. Qed.

forall (n : Z) (x : R), (IZR n <= Rabs x)%R -> (n <= Z.abs (Ztrunc x))%Z

forall (n : Z) (x : R), (IZR n <= Rabs x)%R -> (n <= Z.abs (Ztrunc x))%Z
n:Z
x:R
H:(IZR n <= Rabs x)%R

(n <= Z.abs (Ztrunc x))%Z
n:Z
x:R
H:(IZR n <= Rabs x)%R

(n <= Ztrunc (Rabs x))%Z
n:Z
x:R
H:(IZR n <= Rabs x)%R

(n <= Zfloor (Rabs x))%Z
n:Z
x:R
H:(IZR n <= Rabs x)%R
(0 <= Rabs x)%R
n:Z
x:R
H:(IZR n <= Rabs x)%R

(n <= Zfloor (Rabs x))%Z
now apply Zfloor_lub. Qed. Definition Zaway x := if Rlt_bool x 0 then Zfloor x else Zceil x.

forall n : Z, Zaway (IZR n) = n

forall n : Z, Zaway (IZR n) = n
n:Z

Zaway (IZR n) = n
n:Z

(if Rlt_bool (IZR n) 0 then Zfloor (IZR n) else Zceil (IZR n)) = n
n:Z
H:(IZR n < 0)%R

Zfloor (IZR n) = n
n:Z
H:(0 <= IZR n)%R
Zceil (IZR n) = n
n:Z
H:(0 <= IZR n)%R

Zceil (IZR n) = n
apply Zceil_IZR. Qed.

forall x : R, (0 <= x)%R -> Zaway x = Zceil x

forall x : R, (0 <= x)%R -> Zaway x = Zceil x
x:R
Hx:(0 <= x)%R

Zaway x = Zceil x
x:R
Hx:(0 <= x)%R

(if Rlt_bool x 0 then Zfloor x else Zceil x) = Zceil x
x:R
Hx:(0 <= x)%R
H:(x < 0)%R

Zfloor x = Zceil x
x:R
Hx, H:(0 <= x)%R
Zceil x = Zceil x
x:R
Hx:(0 <= x)%R
H:(x < 0)%R

(x < x)%R
x:R
Hx, H:(0 <= x)%R
Zceil x = Zceil x
x:R
Hx, H:(0 <= x)%R

Zceil x = Zceil x
apply refl_equal. Qed.

forall x : R, (x <= 0)%R -> Zaway x = Zfloor x

forall x : R, (x <= 0)%R -> Zaway x = Zfloor x
x:R
Hx:(x <= 0)%R

Zaway x = Zfloor x
x:R
Hx:(x <= 0)%R

(if Rlt_bool x 0 then Zfloor x else Zceil x) = Zfloor x
x:R
Hx:(x <= 0)%R
H:(x < 0)%R

Zfloor x = Zfloor x
x:R
Hx:(x <= 0)%R
H:(0 <= x)%R
Zceil x = Zfloor x
x:R
Hx:(x <= 0)%R
H:(0 <= x)%R

Zceil x = Zfloor x
x:R
Hx:(x <= 0)%R
H:(0 <= x)%R

Zceil 0 = Zfloor 0
x:R
Hx:(x <= 0)%R
H:(0 <= x)%R

Zceil 0 = 0%Z
apply Zceil_IZR. Qed.

forall x y : R, (x <= y)%R -> (Zaway x <= Zaway y)%Z

forall x y : R, (x <= y)%R -> (Zaway x <= Zaway y)%Z
x, y:R
Hxy:(x <= y)%R

(Zaway x <= Zaway y)%Z
x, y:R
Hxy:(x <= y)%R

((if Rlt_bool x 0 then Zfloor x else Zceil x) <= Zaway y)%Z
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R

(Zfloor x <= Zaway y)%Z
x, y:R
Hxy:(x <= y)%R
Hx:(0 <= x)%R
(Zceil x <= Zaway y)%Z
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R

(Zfloor x <= (if Rlt_bool y 0 then Zfloor y else Zceil y))%Z
x, y:R
Hxy:(x <= y)%R
Hx:(0 <= x)%R
(Zceil x <= Zaway y)%Z
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(y < 0)%R

(Zfloor x <= Zfloor y)%Z
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R
(Zfloor x <= Zceil y)%Z
x, y:R
Hxy:(x <= y)%R
Hx:(0 <= x)%R
(Zceil x <= Zaway y)%Z
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R

(Zfloor x <= Zceil y)%Z
x, y:R
Hxy:(x <= y)%R
Hx:(0 <= x)%R
(Zceil x <= Zaway y)%Z
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R

(IZR (Zfloor x) <= IZR (Zceil y))%R
x, y:R
Hxy:(x <= y)%R
Hx:(0 <= x)%R
(Zceil x <= Zaway y)%Z
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R

(IZR (Zfloor x) <= 0)%R
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R
(0 <= IZR (Zceil y))%R
x, y:R
Hxy:(x <= y)%R
Hx:(0 <= x)%R
(Zceil x <= Zaway y)%Z
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R

(IZR (Zfloor x) < 0)%R
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R
(0 <= IZR (Zceil y))%R
x, y:R
Hxy:(x <= y)%R
Hx:(0 <= x)%R
(Zceil x <= Zaway y)%Z
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R

(IZR (Zfloor x) <= x)%R
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R
(0 <= IZR (Zceil y))%R
x, y:R
Hxy:(x <= y)%R
Hx:(0 <= x)%R
(Zceil x <= Zaway y)%Z
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R

(0 <= IZR (Zceil y))%R
x, y:R
Hxy:(x <= y)%R
Hx:(0 <= x)%R
(Zceil x <= Zaway y)%Z
x, y:R
Hxy:(x <= y)%R
Hx:(x < 0)%R
Hy:(0 <= y)%R

(y <= IZR (Zceil y))%R
x, y:R
Hxy:(x <= y)%R
Hx:(0 <= x)%R
(Zceil x <= Zaway y)%Z
x, y:R
Hxy:(x <= y)%R
Hx:(0 <= x)%R

(Zceil x <= Zaway y)%Z
x, y:R
Hxy:(x <= y)%R
Hx:(0 <= x)%R

(Zceil x <= Zceil y)%Z
x, y:R
Hxy:(x <= y)%R
Hx:(0 <= x)%R
(0 <= y)%R
x, y:R
Hxy:(x <= y)%R
Hx:(0 <= x)%R

(0 <= y)%R
now apply Rle_trans with x. Qed.

forall x : R, Zaway (- x) = (- Zaway x)%Z

forall x : R, Zaway (- x) = (- Zaway x)%Z
x:R

Zaway (- x) = (- Zaway x)%Z
x:R

Zaway (- x) = (- (if Rlt_bool x 0 then Zfloor x else Zceil x))%Z
x:R
H:(x < 0)%R

Zaway (- x) = (- Zfloor x)%Z
x:R
H:(0 <= x)%R
Zaway (- x) = (- Zceil x)%Z
x:R
H:(x < 0)%R

Zceil (- x) = (- Zfloor x)%Z
x:R
H:(x < 0)%R
(0 <= - x)%R
x:R
H:(0 <= x)%R
Zaway (- x) = (- Zceil x)%Z
x:R
H:(x < 0)%R

(- Zfloor (- - x))%Z = (- Zfloor x)%Z
x:R
H:(x < 0)%R
(0 <= - x)%R
x:R
H:(0 <= x)%R
Zaway (- x) = (- Zceil x)%Z
x:R
H:(x < 0)%R

(0 <= - x)%R
x:R
H:(0 <= x)%R
Zaway (- x) = (- Zceil x)%Z
x:R
H:(x < 0)%R

(0 < - x)%R
x:R
H:(0 <= x)%R
Zaway (- x) = (- Zceil x)%Z
x:R
H:(0 <= x)%R

Zaway (- x) = (- Zceil x)%Z
x:R
H:(0 <= x)%R

Zfloor (- x) = (- Zceil x)%Z
x:R
H:(0 <= x)%R
(- x <= 0)%R
x:R
H:(0 <= x)%R

(- Zceil x)%Z = Zfloor (- x)
x:R
H:(0 <= x)%R
(- x <= 0)%R
x:R
H:(0 <= x)%R

(- x <= 0)%R
x:R
H:(0 <= x)%R

(- x <= - 0)%R
now apply Ropp_le_contravar. Qed.

forall x : R, Zaway (Rabs x) = Z.abs (Zaway x)

forall x : R, Zaway (Rabs x) = Z.abs (Zaway x)
x:R

Zaway (Rabs x) = Z.abs (Zaway x)
x:R

Zceil (Rabs x) = Z.abs (Zaway x)
x:R
(0 <= Rabs x)%R
x:R

Zceil (Rabs x) = Z.abs (Zaway x)
x:R

Zceil (Rabs x) = Z.abs (if Rlt_bool x 0 then Zfloor x else Zceil x)
x:R
H:(x < 0)%R

Zceil (Rabs x) = Z.abs (Zfloor x)
x:R
H:(0 <= x)%R
Zceil (Rabs x) = Z.abs (Zceil x)
x:R
H:(x < 0)%R

Zceil (- x) = Z.abs (Zfloor x)
x:R
H:(0 <= x)%R
Zceil (Rabs x) = Z.abs (Zceil x)
x:R
H:(x < 0)%R

Zceil (- x) = (- Zfloor x)%Z
x:R
H:(x < 0)%R
(Zfloor x <= 0)%Z
x:R
H:(0 <= x)%R
Zceil (Rabs x) = Z.abs (Zceil x)
x:R
H:(x < 0)%R

(- - x)%R = x
x:R
H:(x < 0)%R
(Zfloor x <= 0)%Z
x:R
H:(0 <= x)%R
Zceil (Rabs x) = Z.abs (Zceil x)
x:R
H:(x < 0)%R

(Zfloor x <= 0)%Z
x:R
H:(0 <= x)%R
Zceil (Rabs x) = Z.abs (Zceil x)
x:R
H:(x < 0)%R

(IZR (Zfloor x) <= 0)%R
x:R
H:(0 <= x)%R
Zceil (Rabs x) = Z.abs (Zceil x)
x:R
H:(x < 0)%R

(IZR (Zfloor x) < 0)%R
x:R
H:(0 <= x)%R
Zceil (Rabs x) = Z.abs (Zceil x)
x:R
H:(x < 0)%R

(IZR (Zfloor x) <= x)%R
x:R
H:(0 <= x)%R
Zceil (Rabs x) = Z.abs (Zceil x)
x:R
H:(0 <= x)%R

Zceil (Rabs x) = Z.abs (Zceil x)
x:R
H:(0 <= x)%R

Zceil x = Z.abs (Zceil x)
x:R
H:(0 <= x)%R

Z.abs (Zceil x) = Zceil x
x:R
H:(0 <= x)%R

(0 <= Zceil x)%Z
x:R
H:(0 <= x)%R

(0 <= IZR (Zceil x))%R
x:R
H:(0 <= x)%R

(x <= IZR (Zceil x))%R
apply Zceil_ub. Qed. End Floor_Ceil.

forall x : R, IZR (Zfloor x) <> x -> Rcompare (x - IZR (Zfloor x)) (/ 2) = Rcompare (x - IZR (Zfloor x)) (IZR (Zceil x) - x)

forall x : R, IZR (Zfloor x) <> x -> Rcompare (x - IZR (Zfloor x)) (/ 2) = Rcompare (x - IZR (Zfloor x)) (IZR (Zceil x) - x)
x:R
Hx:IZR (Zfloor x) <> x

Rcompare (x - IZR (Zfloor x)) (/ 2) = Rcompare (x - IZR (Zfloor x)) (IZR (Zceil x) - x)
x:R
Hx:IZR (Zfloor x) <> x

Rcompare (x - IZR (Zfloor x)) (/ 2) = Rcompare (x - IZR (Zfloor x)) (IZR (Zfloor x + 1) - x)
x:R
Hx:IZR (Zfloor x) <> x

Rcompare (x - IZR (Zfloor x)) (/ 2) = Rcompare (x - IZR (Zfloor x)) (IZR (Zfloor x) + 1 - x)
x:R
Hx:IZR (Zfloor x) <> x
H1:(x - IZR (Zfloor x) < / 2)%R

Rcompare (x - IZR (Zfloor x)) (IZR (Zfloor x) + 1 - x) = Lt
x:R
Hx:IZR (Zfloor x) <> x
H1:(x - IZR (Zfloor x))%R = (/ 2)%R
Rcompare (x - IZR (Zfloor x)) (IZR (Zfloor x) + 1 - x) = Eq
x:R
Hx:IZR (Zfloor x) <> x
H1:(/ 2 < x - IZR (Zfloor x))%R
Rcompare (x - IZR (Zfloor x)) (IZR (Zfloor x) + 1 - x) = Gt
(* . *)
x:R
Hx:IZR (Zfloor x) <> x
H1:(x - IZR (Zfloor x) < / 2)%R

(x - IZR (Zfloor x) < IZR (Zfloor x) + 1 - x)%R
x:R
Hx:IZR (Zfloor x) <> x
H1:(x - IZR (Zfloor x))%R = (/ 2)%R
Rcompare (x - IZR (Zfloor x)) (IZR (Zfloor x) + 1 - x) = Eq
x:R
Hx:IZR (Zfloor x) <> x
H1:(/ 2 < x - IZR (Zfloor x))%R
Rcompare (x - IZR (Zfloor x)) (IZR (Zfloor x) + 1 - x) = Gt
x:R
Hx:IZR (Zfloor x) <> x
H1:(x - IZR (Zfloor x) < / 2)%R

(x - IZR (Zfloor x) + (x - IZR (Zfloor x)) < x - IZR (Zfloor x) + (IZR (Zfloor x) + 1 - x))%R
x:R
Hx:IZR (Zfloor x) <> x
H1:(x - IZR (Zfloor x))%R = (/ 2)%R
Rcompare (x - IZR (Zfloor x)) (IZR (Zfloor x) + 1 - x) = Eq
x:R
Hx:IZR (Zfloor x) <> x
H1:(/ 2 < x - IZR (Zfloor x))%R
Rcompare (x - IZR (Zfloor x)) (IZR (Zfloor x) + 1 - x) = Gt
x:R
Hx:IZR (Zfloor x) <> x
H1:(x - IZR (Zfloor x) < / 2)%R

((x - IZR (Zfloor x)) * 2 < x - IZR (Zfloor x) + (IZR (Zfloor x) + 1 - x))%R
x:R
Hx:IZR (Zfloor x) <> x
H1:(x - IZR (Zfloor x))%R = (/ 2)%R
Rcompare (x - IZR (Zfloor x)) (IZR (Zfloor x) + 1 - x) = Eq
x:R
Hx:IZR (Zfloor x) <> x
H1:(/ 2 < x - IZR (Zfloor x))%R
Rcompare (x - IZR (Zfloor x)) (IZR (Zfloor x) + 1 - x) = Gt
x:R
Hx:IZR (Zfloor x) <> x
H1:(x - IZR (Zfloor x) < / 2)%R

((x - IZR (Zfloor x)) * 2 < / 2 * 2)%R
x:R
Hx:IZR (Zfloor x) <> x
H1:(x - IZR (Zfloor x))%R = (/ 2)%R
Rcompare (x - IZR (Zfloor x)) (IZR (Zfloor x) + 1 - x) = Eq
x:R
Hx:IZR (Zfloor x) <> x
H1:(/ 2 < x - IZR (Zfloor x))%R
Rcompare (x - IZR (Zfloor x)) (IZR (Zfloor x) + 1 - x) = Gt
x:R
Hx:IZR (Zfloor x) <> x
H1:(x - IZR (Zfloor x) < / 2)%R

(0 < 2)%R
x:R
Hx:IZR (Zfloor x) <> x
H1:(x - IZR (Zfloor x))%R = (/ 2)%R
Rcompare (x - IZR (Zfloor x)) (IZR (Zfloor x) + 1 - x) = Eq
x:R
Hx:IZR (Zfloor x) <> x
H1:(/ 2 < x - IZR (Zfloor x))%R
Rcompare (x - IZR (Zfloor x)) (IZR (Zfloor x) + 1 - x) = Gt
x:R
Hx:IZR (Zfloor x) <> x
H1:(x - IZR (Zfloor x))%R = (/ 2)%R

Rcompare (x - IZR (Zfloor x)) (IZR (Zfloor x) + 1 - x) = Eq
x:R
Hx:IZR (Zfloor x) <> x
H1:(/ 2 < x - IZR (Zfloor x))%R
Rcompare (x - IZR (Zfloor x)) (IZR (Zfloor x) + 1 - x) = Gt
(* . *)
x:R
Hx:IZR (Zfloor x) <> x
H1:(x - IZR (Zfloor x))%R = (/ 2)%R

(x - IZR (Zfloor x))%R = (IZR (Zfloor x) + 1 - x)%R
x:R
Hx:IZR (Zfloor x) <> x
H1:(/ 2 < x - IZR (Zfloor x))%R
Rcompare (x - IZR (Zfloor x)) (IZR (Zfloor x) + 1 - x) = Gt
x:R
Hx:IZR (Zfloor x) <> x
H1:(x - IZR (Zfloor x))%R = (/ 2)%R

(x - IZR (Zfloor x))%R = (1 - (x - IZR (Zfloor x)))%R
x:R
Hx:IZR (Zfloor x) <> x
H1:(/ 2 < x - IZR (Zfloor x))%R
Rcompare (x - IZR (Zfloor x)) (IZR (Zfloor x) + 1 - x) = Gt
x:R
Hx:IZR (Zfloor x) <> x
H1:(x - IZR (Zfloor x))%R = (/ 2)%R

(/ 2)%R = (1 - / 2)%R
x:R
Hx:IZR (Zfloor x) <> x
H1:(/ 2 < x - IZR (Zfloor x))%R
Rcompare (x - IZR (Zfloor x)) (IZR (Zfloor x) + 1 - x) = Gt
x:R
Hx:IZR (Zfloor x) <> x
H1:(/ 2 < x - IZR (Zfloor x))%R

Rcompare (x - IZR (Zfloor x)) (IZR (Zfloor x) + 1 - x) = Gt
(* . *)
x:R
Hx:IZR (Zfloor x) <> x
H1:(/ 2 < x - IZR (Zfloor x))%R

(IZR (Zfloor x) + 1 - x < x - IZR (Zfloor x))%R
x:R
Hx:IZR (Zfloor x) <> x
H1:(/ 2 < x - IZR (Zfloor x))%R

(x - IZR (Zfloor x) + (IZR (Zfloor x) + 1 - x) < x - IZR (Zfloor x) + (x - IZR (Zfloor x)))%R
x:R
Hx:IZR (Zfloor x) <> x
H1:(/ 2 < x - IZR (Zfloor x))%R

(x - IZR (Zfloor x) + (IZR (Zfloor x) + 1 - x) < (x - IZR (Zfloor x)) * 2)%R
x:R
Hx:IZR (Zfloor x) <> x
H1:(/ 2 < x - IZR (Zfloor x))%R

(/ 2 * 2 < (x - IZR (Zfloor x)) * 2)%R
x:R
Hx:IZR (Zfloor x) <> x
H1:(/ 2 < x - IZR (Zfloor x))%R

(0 < 2)%R
now apply IZR_lt. Qed.

forall x : R, IZR (Zfloor x) <> x -> Rcompare (IZR (Zceil x) - x) (/ 2) = Rcompare (IZR (Zceil x) - x) (x - IZR (Zfloor x))

forall x : R, IZR (Zfloor x) <> x -> Rcompare (IZR (Zceil x) - x) (/ 2) = Rcompare (IZR (Zceil x) - x) (x - IZR (Zfloor x))
x:R
Hx:IZR (Zfloor x) <> x

Rcompare (IZR (Zceil x) - x) (/ 2) = Rcompare (IZR (Zceil x) - x) (x - IZR (Zfloor x))
x:R
Hx:IZR (Zfloor x) <> x

Rcompare (IZR (Zfloor x + 1) - x) (/ 2) = Rcompare (IZR (Zfloor x + 1) - x) (x - IZR (Zfloor x))
x:R
Hx:IZR (Zfloor x) <> x

Rcompare (IZR (Zfloor x) + 1 - x) (/ 2) = Rcompare (IZR (Zfloor x) + 1 - x) (x - IZR (Zfloor x))
x:R
Hx:IZR (Zfloor x) <> x
H1:(IZR (Zfloor x) + 1 - x < / 2)%R

Rcompare (IZR (Zfloor x) + 1 - x) (x - IZR (Zfloor x)) = Lt
x:R
Hx:IZR (Zfloor x) <> x
H1:(IZR (Zfloor x) + 1 - x)%R = (/ 2)%R
Rcompare (IZR (Zfloor x) + 1 - x) (x - IZR (Zfloor x)) = Eq
x:R
Hx:IZR (Zfloor x) <> x
H1:(/ 2 < IZR (Zfloor x) + 1 - x)%R
Rcompare (IZR (Zfloor x) + 1 - x) (x - IZR (Zfloor x)) = Gt
(* . *)
x:R
Hx:IZR (Zfloor x) <> x
H1:(IZR (Zfloor x) + 1 - x < / 2)%R

(IZR (Zfloor x) + 1 - x < x - IZR (Zfloor x))%R
x:R
Hx:IZR (Zfloor x) <> x
H1:(IZR (Zfloor x) + 1 - x)%R = (/ 2)%R
Rcompare (IZR (Zfloor x) + 1 - x) (x - IZR (Zfloor x)) = Eq
x:R
Hx:IZR (Zfloor x) <> x
H1:(/ 2 < IZR (Zfloor x) + 1 - x)%R
Rcompare (IZR (Zfloor x) + 1 - x) (x - IZR (Zfloor x)) = Gt
x:R
Hx:IZR (Zfloor x) <> x
H1:(IZR (Zfloor x) + 1 - x < / 2)%R

(IZR (Zfloor x) + 1 - x + (IZR (Zfloor x) + 1 - x) < IZR (Zfloor x) + 1 - x + (x - IZR (Zfloor x)))%R
x:R
Hx:IZR (Zfloor x) <> x
H1:(IZR (Zfloor x) + 1 - x)%R = (/ 2)%R
Rcompare (IZR (Zfloor x) + 1 - x) (x - IZR (Zfloor x)) = Eq
x:R
Hx:IZR (Zfloor x) <> x
H1:(/ 2 < IZR (Zfloor x) + 1 - x)%R
Rcompare (IZR (Zfloor x) + 1 - x) (x - IZR (Zfloor x)) = Gt
x:R
Hx:IZR (Zfloor x) <> x
H1:(IZR (Zfloor x) + 1 - x < / 2)%R

((IZR (Zfloor x) + 1 - x) * 2 < IZR (Zfloor x) + 1 - x + (x - IZR (Zfloor x)))%R
x:R
Hx:IZR (Zfloor x) <> x
H1:(IZR (Zfloor x) + 1 - x)%R = (/ 2)%R
Rcompare (IZR (Zfloor x) + 1 - x) (x - IZR (Zfloor x)) = Eq
x:R
Hx:IZR (Zfloor x) <> x
H1:(/ 2 < IZR (Zfloor x) + 1 - x)%R
Rcompare (IZR (Zfloor x) + 1 - x) (x - IZR (Zfloor x)) = Gt
x:R
Hx:IZR (Zfloor x) <> x
H1:(IZR (Zfloor x) + 1 - x < / 2)%R

((IZR (Zfloor x) + 1 - x) * 2 < / 2 * 2)%R
x:R
Hx:IZR (Zfloor x) <> x
H1:(IZR (Zfloor x) + 1 - x)%R = (/ 2)%R
Rcompare (IZR (Zfloor x) + 1 - x) (x - IZR (Zfloor x)) = Eq
x:R
Hx:IZR (Zfloor x) <> x
H1:(/ 2 < IZR (Zfloor x) + 1 - x)%R
Rcompare (IZR (Zfloor x) + 1 - x) (x - IZR (Zfloor x)) = Gt
x:R
Hx:IZR (Zfloor x) <> x
H1:(IZR (Zfloor x) + 1 - x < / 2)%R

(0 < 2)%R
x:R
Hx:IZR (Zfloor x) <> x
H1:(IZR (Zfloor x) + 1 - x)%R = (/ 2)%R
Rcompare (IZR (Zfloor x) + 1 - x) (x - IZR (Zfloor x)) = Eq
x:R
Hx:IZR (Zfloor x) <> x
H1:(/ 2 < IZR (Zfloor x) + 1 - x)%R
Rcompare (IZR (Zfloor x) + 1 - x) (x - IZR (Zfloor x)) = Gt
x:R
Hx:IZR (Zfloor x) <> x
H1:(IZR (Zfloor x) + 1 - x)%R = (/ 2)%R

Rcompare (IZR (Zfloor x) + 1 - x) (x - IZR (Zfloor x)) = Eq
x:R
Hx:IZR (Zfloor x) <> x
H1:(/ 2 < IZR (Zfloor x) + 1 - x)%R
Rcompare (IZR (Zfloor x) + 1 - x) (x - IZR (Zfloor x)) = Gt
(* . *)
x:R
Hx:IZR (Zfloor x) <> x
H1:(IZR (Zfloor x) + 1 - x)%R = (/ 2)%R

(IZR (Zfloor x) + 1 - x)%R = (x - IZR (Zfloor x))%R
x:R
Hx:IZR (Zfloor x) <> x
H1:(/ 2 < IZR (Zfloor x) + 1 - x)%R
Rcompare (IZR (Zfloor x) + 1 - x) (x - IZR (Zfloor x)) = Gt
x:R
Hx:IZR (Zfloor x) <> x
H1:(IZR (Zfloor x) + 1 - x)%R = (/ 2)%R

(IZR (Zfloor x) + 1 - x)%R = (1 - (IZR (Zfloor x) + 1 - x))%R
x:R
Hx:IZR (Zfloor x) <> x
H1:(/ 2 < IZR (Zfloor x) + 1 - x)%R
Rcompare (IZR (Zfloor x) + 1 - x) (x - IZR (Zfloor x)) = Gt
x:R
Hx:IZR (Zfloor x) <> x
H1:(IZR (Zfloor x) + 1 - x)%R = (/ 2)%R

(/ 2)%R = (1 - / 2)%R
x:R
Hx:IZR (Zfloor x) <> x
H1:(/ 2 < IZR (Zfloor x) + 1 - x)%R
Rcompare (IZR (Zfloor x) + 1 - x) (x - IZR (Zfloor x)) = Gt
x:R
Hx:IZR (Zfloor x) <> x
H1:(/ 2 < IZR (Zfloor x) + 1 - x)%R

Rcompare (IZR (Zfloor x) + 1 - x) (x - IZR (Zfloor x)) = Gt
(* . *)
x:R
Hx:IZR (Zfloor x) <> x
H1:(/ 2 < IZR (Zfloor x) + 1 - x)%R

(x - IZR (Zfloor x) < IZR (Zfloor x) + 1 - x)%R
x:R
Hx:IZR (Zfloor x) <> x
H1:(/ 2 < IZR (Zfloor x) + 1 - x)%R

(IZR (Zfloor x) + 1 - x + (x - IZR (Zfloor x)) < IZR (Zfloor x) + 1 - x + (IZR (Zfloor x) + 1 - x))%R
x:R
Hx:IZR (Zfloor x) <> x
H1:(/ 2 < IZR (Zfloor x) + 1 - x)%R

(IZR (Zfloor x) + 1 - x + (x - IZR (Zfloor x)) < (IZR (Zfloor x) + 1 - x) * 2)%R
x:R
Hx:IZR (Zfloor x) <> x
H1:(/ 2 < IZR (Zfloor x) + 1 - x)%R

(/ 2 * 2 < (IZR (Zfloor x) + 1 - x) * 2)%R
x:R
Hx:IZR (Zfloor x) <> x
H1:(/ 2 < IZR (Zfloor x) + 1 - x)%R

(0 < 2)%R
now apply IZR_lt. Qed. Section Zdiv_Rdiv.

forall x y : Z, y <> 0%Z -> Zfloor (IZR x / IZR y) = (x / y)%Z

forall x y : Z, y <> 0%Z -> Zfloor (IZR x / IZR y) = (x / y)%Z
x, y:Z
Zy:y <> 0%Z

Zfloor (IZR x / IZR y) = (x / y)%Z
x, y:Z
Zy:y <> 0%Z

x = (y * (x / y) + x mod y)%Z -> Zfloor (IZR x / IZR y) = (x / y)%Z
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z

Zfloor (IZR x / IZR y) = (x / y)%Z
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z

Zfloor (IZR (y * (x / y) + x mod y) / IZR y) = (x / y)%Z
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z

IZR y <> 0%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
Zfloor (IZR (y * (x / y) + x mod y) / IZR y) = (x / y)%Z
x, y:Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy:IZR y = 0%R

y = 0%Z
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
Zfloor (IZR (y * (x / y) + x mod y) / IZR y) = (x / y)%Z
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R

Zfloor (IZR (y * (x / y) + x mod y) / IZR y) = (x / y)%Z
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R

Zfloor (IZR (y * (x / y) + x mod y) * / IZR y) = (x / y)%Z
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R

Zfloor (IZR y * IZR (x / y) * / IZR y + IZR (x mod y) * / IZR y) = (x / y)%Z
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R

Zfloor (IZR (x / y) + IZR (x mod y) * / IZR y) = (x / y)%Z
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R

(IZR (x / y) <= IZR (x / y) + IZR (x mod y) * / IZR y < IZR (x / y + 1))%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R

(IZR (x / y) <= IZR (x / y) + IZR (x mod y) * / IZR y < IZR (x / y) + 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R

(0 <= IZR (x mod y) * / IZR y < 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:(0 <= IZR (x mod y) * / IZR y < 1)%R
(IZR (x / y) <= IZR (x / y) + IZR (x mod y) * / IZR y < IZR (x / y) + 1)%R
(* *)
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R

forall x' y' : Z, (0 < y')%Z -> (0 <= IZR (x' mod y') * / IZR y' < 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:forall x' y' : Z, (0 < y')%Z -> (0 <= IZR (x' mod y') * / IZR y' < 1)%R
(0 <= IZR (x mod y) * / IZR y < 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:(0 <= IZR (x mod y) * / IZR y < 1)%R
(IZR (x / y) <= IZR (x / y) + IZR (x mod y) * / IZR y < IZR (x / y) + 1)%R
(* . *)

forall x' y' : Z, (0 < y')%Z -> (0 <= IZR (x' mod y') * / IZR y' < 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:forall x' y' : Z, (0 < y')%Z -> (0 <= IZR (x' mod y') * / IZR y' < 1)%R
(0 <= IZR (x mod y) * / IZR y < 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:(0 <= IZR (x mod y) * / IZR y < 1)%R
(IZR (x / y) <= IZR (x / y) + IZR (x mod y) * / IZR y < IZR (x / y) + 1)%R
x, y:Z
Hy:(0 < y)%Z

(0 <= IZR (x mod y) * / IZR y < 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:forall x' y' : Z, (0 < y')%Z -> (0 <= IZR (x' mod y') * / IZR y' < 1)%R
(0 <= IZR (x mod y) * / IZR y < 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:(0 <= IZR (x mod y) * / IZR y < 1)%R
(IZR (x / y) <= IZR (x / y) + IZR (x mod y) * / IZR y < IZR (x / y) + 1)%R
x, y:Z
Hy:(0 < y)%Z

(0 <= IZR (x mod y) * / IZR y)%R
x, y:Z
Hy:(0 < y)%Z
(IZR (x mod y) * / IZR y < 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:forall x' y' : Z, (0 < y')%Z -> (0 <= IZR (x' mod y') * / IZR y' < 1)%R
(0 <= IZR (x mod y) * / IZR y < 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:(0 <= IZR (x mod y) * / IZR y < 1)%R
(IZR (x / y) <= IZR (x / y) + IZR (x mod y) * / IZR y < IZR (x / y) + 1)%R
x, y:Z
Hy:(0 < y)%Z

(0 <= IZR (x mod y))%R
x, y:Z
Hy:(0 < y)%Z
(0 <= / IZR y)%R
x, y:Z
Hy:(0 < y)%Z
(IZR (x mod y) * / IZR y < 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:forall x' y' : Z, (0 < y')%Z -> (0 <= IZR (x' mod y') * / IZR y' < 1)%R
(0 <= IZR (x mod y) * / IZR y < 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:(0 <= IZR (x mod y) * / IZR y < 1)%R
(IZR (x / y) <= IZR (x / y) + IZR (x mod y) * / IZR y < IZR (x / y) + 1)%R
x, y:Z
Hy:(0 < y)%Z

(0 <= x mod y)%Z
x, y:Z
Hy:(0 < y)%Z
(0 <= / IZR y)%R
x, y:Z
Hy:(0 < y)%Z
(IZR (x mod y) * / IZR y < 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:forall x' y' : Z, (0 < y')%Z -> (0 <= IZR (x' mod y') * / IZR y' < 1)%R
(0 <= IZR (x mod y) * / IZR y < 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:(0 <= IZR (x mod y) * / IZR y < 1)%R
(IZR (x / y) <= IZR (x / y) + IZR (x mod y) * / IZR y < IZR (x / y) + 1)%R
x, y:Z
Hy:(0 < y)%Z

(y > 0)%Z
x, y:Z
Hy:(0 < y)%Z
(0 <= / IZR y)%R
x, y:Z
Hy:(0 < y)%Z
(IZR (x mod y) * / IZR y < 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:forall x' y' : Z, (0 < y')%Z -> (0 <= IZR (x' mod y') * / IZR y' < 1)%R
(0 <= IZR (x mod y) * / IZR y < 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:(0 <= IZR (x mod y) * / IZR y < 1)%R
(IZR (x / y) <= IZR (x / y) + IZR (x mod y) * / IZR y < IZR (x / y) + 1)%R
x, y:Z
Hy:(0 < y)%Z

(0 <= / IZR y)%R
x, y:Z
Hy:(0 < y)%Z
(IZR (x mod y) * / IZR y < 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:forall x' y' : Z, (0 < y')%Z -> (0 <= IZR (x' mod y') * / IZR y' < 1)%R
(0 <= IZR (x mod y) * / IZR y < 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:(0 <= IZR (x mod y) * / IZR y < 1)%R
(IZR (x / y) <= IZR (x / y) + IZR (x mod y) * / IZR y < IZR (x / y) + 1)%R
x, y:Z
Hy:(0 < y)%Z

(0 < / IZR y)%R
x, y:Z
Hy:(0 < y)%Z
(IZR (x mod y) * / IZR y < 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:forall x' y' : Z, (0 < y')%Z -> (0 <= IZR (x' mod y') * / IZR y' < 1)%R
(0 <= IZR (x mod y) * / IZR y < 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:(0 <= IZR (x mod y) * / IZR y < 1)%R
(IZR (x / y) <= IZR (x / y) + IZR (x mod y) * / IZR y < IZR (x / y) + 1)%R
x, y:Z
Hy:(0 < y)%Z

(0 < IZR y)%R
x, y:Z
Hy:(0 < y)%Z
(IZR (x mod y) * / IZR y < 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:forall x' y' : Z, (0 < y')%Z -> (0 <= IZR (x' mod y') * / IZR y' < 1)%R
(0 <= IZR (x mod y) * / IZR y < 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:(0 <= IZR (x mod y) * / IZR y < 1)%R
(IZR (x / y) <= IZR (x / y) + IZR (x mod y) * / IZR y < IZR (x / y) + 1)%R
x, y:Z
Hy:(0 < y)%Z

(IZR (x mod y) * / IZR y < 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:forall x' y' : Z, (0 < y')%Z -> (0 <= IZR (x' mod y') * / IZR y' < 1)%R
(0 <= IZR (x mod y) * / IZR y < 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:(0 <= IZR (x mod y) * / IZR y < 1)%R
(IZR (x / y) <= IZR (x / y) + IZR (x mod y) * / IZR y < IZR (x / y) + 1)%R
x, y:Z
Hy:(0 < y)%Z

(0 < IZR y)%R
x, y:Z
Hy:(0 < y)%Z
(IZR (x mod y) * / IZR y * IZR y < 1 * IZR y)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:forall x' y' : Z, (0 < y')%Z -> (0 <= IZR (x' mod y') * / IZR y' < 1)%R
(0 <= IZR (x mod y) * / IZR y < 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:(0 <= IZR (x mod y) * / IZR y < 1)%R
(IZR (x / y) <= IZR (x / y) + IZR (x mod y) * / IZR y < IZR (x / y) + 1)%R
x, y:Z
Hy:(0 < y)%Z

(IZR (x mod y) * / IZR y * IZR y < 1 * IZR y)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:forall x' y' : Z, (0 < y')%Z -> (0 <= IZR (x' mod y') * / IZR y' < 1)%R
(0 <= IZR (x mod y) * / IZR y < 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:(0 <= IZR (x mod y) * / IZR y < 1)%R
(IZR (x / y) <= IZR (x / y) + IZR (x mod y) * / IZR y < IZR (x / y) + 1)%R
x, y:Z
Hy:(0 < y)%Z

(IZR (x mod y) < IZR y)%R
x, y:Z
Hy:(0 < y)%Z
IZR y <> 0%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:forall x' y' : Z, (0 < y')%Z -> (0 <= IZR (x' mod y') * / IZR y' < 1)%R
(0 <= IZR (x mod y) * / IZR y < 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:(0 <= IZR (x mod y) * / IZR y < 1)%R
(IZR (x / y) <= IZR (x / y) + IZR (x mod y) * / IZR y < IZR (x / y) + 1)%R
x, y:Z
Hy:(0 < y)%Z

(x mod y < y)%Z
x, y:Z
Hy:(0 < y)%Z
IZR y <> 0%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:forall x' y' : Z, (0 < y')%Z -> (0 <= IZR (x' mod y') * / IZR y' < 1)%R
(0 <= IZR (x mod y) * / IZR y < 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:(0 <= IZR (x mod y) * / IZR y < 1)%R
(IZR (x / y) <= IZR (x / y) + IZR (x mod y) * / IZR y < IZR (x / y) + 1)%R
x, y:Z
Hy:(0 < y)%Z

(y > 0)%Z
x, y:Z
Hy:(0 < y)%Z
IZR y <> 0%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:forall x' y' : Z, (0 < y')%Z -> (0 <= IZR (x' mod y') * / IZR y' < 1)%R
(0 <= IZR (x mod y) * / IZR y < 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:(0 <= IZR (x mod y) * / IZR y < 1)%R
(IZR (x / y) <= IZR (x / y) + IZR (x mod y) * / IZR y < IZR (x / y) + 1)%R
x, y:Z
Hy:(0 < y)%Z

IZR y <> 0%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:forall x' y' : Z, (0 < y')%Z -> (0 <= IZR (x' mod y') * / IZR y' < 1)%R
(0 <= IZR (x mod y) * / IZR y < 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:(0 <= IZR (x mod y) * / IZR y < 1)%R
(IZR (x / y) <= IZR (x / y) + IZR (x mod y) * / IZR y < IZR (x / y) + 1)%R
x, y:Z
Hy:(0 < y)%Z

(IZR y > 0)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:forall x' y' : Z, (0 < y')%Z -> (0 <= IZR (x' mod y') * / IZR y' < 1)%R
(0 <= IZR (x mod y) * / IZR y < 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:(0 <= IZR (x mod y) * / IZR y < 1)%R
(IZR (x / y) <= IZR (x / y) + IZR (x mod y) * / IZR y < IZR (x / y) + 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:forall x' y' : Z, (0 < y')%Z -> (0 <= IZR (x' mod y') * / IZR y' < 1)%R

(0 <= IZR (x mod y) * / IZR y < 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:(0 <= IZR (x mod y) * / IZR y < 1)%R
(IZR (x / y) <= IZR (x / y) + IZR (x mod y) * / IZR y < IZR (x / y) + 1)%R
(* . *)
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:forall x' y' : Z, (0 < y')%Z -> (0 <= IZR (x' mod y') * / IZR y' < 1)%R
Hy:(y < 0)%Z

(0 <= IZR (x mod y) * / IZR y < 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:forall x' y' : Z, (0 < y')%Z -> (0 <= IZR (x' mod y') * / IZR y' < 1)%R
Hy:(0 <= y)%Z
(0 <= IZR (x mod y) * / IZR y < 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:(0 <= IZR (x mod y) * / IZR y < 1)%R
(IZR (x / y) <= IZR (x / y) + IZR (x mod y) * / IZR y < IZR (x / y) + 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:forall x' y' : Z, (0 < y')%Z -> (0 <= IZR (x' mod y') * / IZR y' < 1)%R
Hy:(y < 0)%Z

(0 <= - IZR (x mod y) * - / IZR y < 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:forall x' y' : Z, (0 < y')%Z -> (0 <= IZR (x' mod y') * / IZR y' < 1)%R
Hy:(0 <= y)%Z
(0 <= IZR (x mod y) * / IZR y < 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:(0 <= IZR (x mod y) * / IZR y < 1)%R
(IZR (x / y) <= IZR (x / y) + IZR (x mod y) * / IZR y < IZR (x / y) + 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:forall x' y' : Z, (0 < y')%Z -> (0 <= IZR (x' mod y') * / IZR y' < 1)%R
Hy:(y < 0)%Z

(0 <= - IZR (x mod y) * / - IZR y < 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:forall x' y' : Z, (0 < y')%Z -> (0 <= IZR (x' mod y') * / IZR y' < 1)%R
Hy:(0 <= y)%Z
(0 <= IZR (x mod y) * / IZR y < 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:(0 <= IZR (x mod y) * / IZR y < 1)%R
(IZR (x / y) <= IZR (x / y) + IZR (x mod y) * / IZR y < IZR (x / y) + 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:forall x' y' : Z, (0 < y')%Z -> (0 <= IZR (x' mod y') * / IZR y' < 1)%R
Hy:(y < 0)%Z

(0 <= IZR (- (x mod y)) * / IZR (- y) < 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:forall x' y' : Z, (0 < y')%Z -> (0 <= IZR (x' mod y') * / IZR y' < 1)%R
Hy:(0 <= y)%Z
(0 <= IZR (x mod y) * / IZR y < 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:(0 <= IZR (x mod y) * / IZR y < 1)%R
(IZR (x / y) <= IZR (x / y) + IZR (x mod y) * / IZR y < IZR (x / y) + 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:forall x' y' : Z, (0 < y')%Z -> (0 <= IZR (x' mod y') * / IZR y' < 1)%R
Hy:(y < 0)%Z

(0 <= IZR (- x mod - y) * / IZR (- y) < 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:forall x' y' : Z, (0 < y')%Z -> (0 <= IZR (x' mod y') * / IZR y' < 1)%R
Hy:(0 <= y)%Z
(0 <= IZR (x mod y) * / IZR y < 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:(0 <= IZR (x mod y) * / IZR y < 1)%R
(IZR (x / y) <= IZR (x / y) + IZR (x mod y) * / IZR y < IZR (x / y) + 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:forall x' y' : Z, (0 < y')%Z -> (0 <= IZR (x' mod y') * / IZR y' < 1)%R
Hy:(y < 0)%Z

(0 < - y)%Z
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:forall x' y' : Z, (0 < y')%Z -> (0 <= IZR (x' mod y') * / IZR y' < 1)%R
Hy:(0 <= y)%Z
(0 <= IZR (x mod y) * / IZR y < 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:(0 <= IZR (x mod y) * / IZR y < 1)%R
(IZR (x / y) <= IZR (x / y) + IZR (x mod y) * / IZR y < IZR (x / y) + 1)%R
y:Z
Hy:(y < 0)%Z

(0 < - y)%Z
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:forall x' y' : Z, (0 < y')%Z -> (0 <= IZR (x' mod y') * / IZR y' < 1)%R
Hy:(0 <= y)%Z
(0 <= IZR (x mod y) * / IZR y < 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:(0 <= IZR (x mod y) * / IZR y < 1)%R
(IZR (x / y) <= IZR (x / y) + IZR (x mod y) * / IZR y < IZR (x / y) + 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:forall x' y' : Z, (0 < y')%Z -> (0 <= IZR (x' mod y') * / IZR y' < 1)%R
Hy:(0 <= y)%Z

(0 <= IZR (x mod y) * / IZR y < 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:(0 <= IZR (x mod y) * / IZR y < 1)%R
(IZR (x / y) <= IZR (x / y) + IZR (x mod y) * / IZR y < IZR (x / y) + 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:forall x' y' : Z, (0 < y')%Z -> (0 <= IZR (x' mod y') * / IZR y' < 1)%R
Hy:(0 <= y)%Z

(0 < y)%Z
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:(0 <= IZR (x mod y) * / IZR y < 1)%R
(IZR (x / y) <= IZR (x / y) + IZR (x mod y) * / IZR y < IZR (x / y) + 1)%R
y:Z
Zy:y <> 0%Z
Hy:(0 <= y)%Z

(0 < y)%Z
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:(0 <= IZR (x mod y) * / IZR y < 1)%R
(IZR (x / y) <= IZR (x / y) + IZR (x mod y) * / IZR y < IZR (x / y) + 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:(0 <= IZR (x mod y) * / IZR y < 1)%R

(IZR (x / y) <= IZR (x / y) + IZR (x mod y) * / IZR y < IZR (x / y) + 1)%R
(* *)
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:(0 <= IZR (x mod y) * / IZR y < 1)%R

(IZR (x / y) <= IZR (x / y) + IZR (x mod y) * / IZR y)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:(0 <= IZR (x mod y) * / IZR y < 1)%R
(IZR (x / y) + IZR (x mod y) * / IZR y < IZR (x / y) + 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:(0 <= IZR (x mod y) * / IZR y < 1)%R

(IZR (x / y) + 0 <= IZR (x / y) + IZR (x mod y) * / IZR y)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:(0 <= IZR (x mod y) * / IZR y < 1)%R
(IZR (x / y) + IZR (x mod y) * / IZR y < IZR (x / y) + 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:(0 <= IZR (x mod y) * / IZR y < 1)%R

(0 <= IZR (x mod y) * / IZR y)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:(0 <= IZR (x mod y) * / IZR y < 1)%R
(IZR (x / y) + IZR (x mod y) * / IZR y < IZR (x / y) + 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:(0 <= IZR (x mod y) * / IZR y < 1)%R

(IZR (x / y) + IZR (x mod y) * / IZR y < IZR (x / y) + 1)%R
x, y:Z
Zy:y <> 0%Z
Hx:x = (y * (x / y) + x mod y)%Z
Zy':IZR y <> 0%R
H:(0 <= IZR (x mod y) * / IZR y < 1)%R

(IZR (x mod y) * / IZR y < 1)%R
apply H. Qed. End Zdiv_Rdiv. Section pow. Variable r : radix.
r:radix

(0 < IZR r)%R
r:radix

(0 < IZR r)%R
r:radix
v:Z
Hr:(2 <=? v)%Z = true

(0 < IZR {| radix_val := v; radix_prop := Hr |})%R
r:radix
v:Z
Hr:(2 <=? v)%Z = true

(0 < IZR v)%R
r:radix
v:Z
Hr:(2 <=? v)%Z = true

(0 < v)%Z
r:radix
v:Z
Hr:(2 <=? v)%Z = true

(0 < 2)%Z
r:radix
v:Z
Hr:(2 <=? v)%Z = true
(2 <= v)%Z
r:radix
v:Z
Hr:(2 <=? v)%Z = true

(2 <= v)%Z
now apply Zle_bool_imp_le. Qed.
Well-used function called bpow for computing the power function β^e
Definition bpow e :=
  match e with
  | Zpos p => IZR (Zpower_pos r p)
  | Zneg p => Rinv (IZR (Zpower_pos r p))
  | Z0 => 1%R
  end.

r:radix

forall (n : Z) (m : positive), IZR (Z.pow_pos n m) = powerRZ (IZR n) (Z.pos m)
r:radix

forall (n : Z) (m : positive), IZR (Z.pow_pos n m) = powerRZ (IZR n) (Z.pos m)
r:radix
n:Z
m:positive

IZR (Z.pow_pos n m) = powerRZ (IZR n) (Z.pos m)
r:radix
n:Z
m:positive

IZR (Zpower_nat n (Pos.to_nat m)) = powerRZ (IZR n) (Z.pos m)
r:radix
n:Z
m:positive

IZR (Zpower_nat n (Pos.to_nat m)) = (IZR n ^ Pos.to_nat m)%R
r:radix
n:Z
m:positive

IZR (Zpower_nat n 0) = (IZR n ^ 0)%R
r:radix
n:Z
m:positive
n0:nat
IHn0:IZR (Zpower_nat n n0) = (IZR n ^ n0)%R
IZR (Zpower_nat n (S n0)) = (IZR n ^ S n0)%R
r:radix
n:Z
m:positive
n0:nat
IHn0:IZR (Zpower_nat n n0) = (IZR n ^ n0)%R

IZR (Zpower_nat n (S n0)) = (IZR n ^ S n0)%R
r:radix
n:Z
m:positive
n0:nat
IHn0:IZR (Zpower_nat n n0) = (IZR n ^ n0)%R

IZR (nat_rect (fun _ : nat => Z) 1%Z (fun _ : nat => Z.mul n) (S n0)) = (IZR n ^ S n0)%R
r:radix
n:Z
m:positive
n0:nat
IHn0:IZR (Zpower_nat n n0) = (IZR n ^ n0)%R

IZR (n * nat_rect (fun _ : nat => Z) 1%Z (fun _ : nat => Z.mul n) n0) = (IZR n * IZR n ^ n0)%R
r:radix
n:Z
m:positive
n0:nat
IHn0:IZR (Zpower_nat n n0) = (IZR n ^ n0)%R

(IZR n * IZR (nat_rect (fun _ : nat => Z) 1%Z (fun _ : nat => Z.mul n) n0))%R = (IZR n * IZR n ^ n0)%R
r:radix
n:Z
m:positive
n0:nat
IHn0:IZR (Zpower_nat n n0) = (IZR n ^ n0)%R

IZR (nat_rect (fun _ : nat => Z) 1%Z (fun _ : nat => Z.mul n) n0) = (IZR n ^ n0)%R
exact IHn0. Qed.
r:radix

forall e : Z, bpow e = powerRZ (IZR r) e
r:radix

forall e : Z, bpow e = powerRZ (IZR r) e
r:radix

1%R = powerRZ (IZR r) 0
r:radix
p:positive
IZR (Z.pow_pos r p) = powerRZ (IZR r) (Z.pos p)
r:radix
p:positive
(/ IZR (Z.pow_pos r p))%R = powerRZ (IZR r) (Z.neg p)
r:radix
p:positive

IZR (Z.pow_pos r p) = powerRZ (IZR r) (Z.pos p)
r:radix
p:positive
(/ IZR (Z.pow_pos r p))%R = powerRZ (IZR r) (Z.neg p)
r:radix
p:positive

(/ IZR (Z.pow_pos r p))%R = powerRZ (IZR r) (Z.neg p)
now rewrite IZR_Zpower_pos. Qed.
r:radix

forall e : Z, (0 <= bpow e)%R
r:radix

forall e : Z, (0 <= bpow e)%R
r:radix
e:Z

(0 <= bpow e)%R
r:radix
e:Z

(0 <= powerRZ (IZR r) e)%R
r:radix
e:Z

(0 < IZR r)%R
apply radix_pos. Qed.
r:radix

forall e : Z, (0 < bpow e)%R
r:radix

forall e : Z, (0 < bpow e)%R
r:radix
e:Z

(0 < bpow e)%R
r:radix
e:Z

(0 < powerRZ (IZR r) e)%R
r:radix
e:Z

(0 < IZR r)%R
apply radix_pos. Qed.
r:radix

forall e1 e2 : Z, bpow (e1 + e2) = (bpow e1 * bpow e2)%R
r:radix

forall e1 e2 : Z, bpow (e1 + e2) = (bpow e1 * bpow e2)%R
r:radix
e1, e2:Z

bpow (e1 + e2) = (bpow e1 * bpow e2)%R
r:radix
e1, e2:Z

powerRZ (IZR r) (e1 + e2) = (powerRZ (IZR r) e1 * powerRZ (IZR r) e2)%R
r:radix
e1, e2:Z

IZR r <> 0%R
r:radix
e1, e2:Z

(IZR r > 0)%R
apply radix_pos. Qed.
r:radix

bpow 1 = IZR r
r:radix

bpow 1 = IZR r
r:radix

IZR (Pos.iter (Z.mul r) 1%Z 1) = IZR r
r:radix

IZR (r * 1) = IZR r
now rewrite Zmult_1_r. Qed.
r:radix

forall e : Z, bpow (e + 1) = (IZR r * bpow e)%R
r:radix

forall e : Z, bpow (e + 1) = (IZR r * bpow e)%R
r:radix
e:Z

bpow (e + 1) = (IZR r * bpow e)%R
r:radix
e:Z

bpow (e + 1) = (bpow 1 * bpow e)%R
r:radix
e:Z

bpow (e + 1) = bpow (1 + e)
now rewrite Zplus_comm. Qed.
r:radix

forall e : Z, bpow (- e) = (/ bpow e)%R
r:radix

forall e : Z, bpow (- e) = (/ bpow e)%R
r:radix

bpow (- 0) = (/ bpow 0)%R
r:radix
p:positive
bpow (- Z.pos p) = (/ bpow (Z.pos p))%R
r:radix
p:positive
bpow (- Z.neg p) = (/ bpow (Z.neg p))%R
r:radix
p:positive

bpow (- Z.pos p) = (/ bpow (Z.pos p))%R
r:radix
p:positive
bpow (- Z.neg p) = (/ bpow (Z.neg p))%R
r:radix
p:positive

bpow (- Z.neg p) = (/ bpow (Z.neg p))%R
r:radix
p:positive

bpow (Z.pos p) = (/ bpow (Z.neg p))%R
r:radix
p:positive

IZR (Z.pow_pos r p) <> 0%R
r:radix
p:positive

(IZR (Z.pow_pos r p) > 0)%R
apply (bpow_gt_0 (Zpos p)). Qed.
r:radix

forall e : nat, IZR (Zpower_nat r e) = bpow (Z.of_nat e)
r:radix

forall e : nat, IZR (Zpower_nat r e) = bpow (Z.of_nat e)
r:radix

IZR (Zpower_nat r 0) = bpow (Z.of_nat 0)
r:radix
e:nat
IZR (Zpower_nat r (S e)) = bpow (Z.of_nat (S e))
r:radix
e:nat

IZR (Zpower_nat r (S e)) = bpow (Z.of_nat (S e))
r:radix
e:nat

IZR (Zpower_nat r (Pos.to_nat (Pos.of_succ_nat e))) = bpow (Z.of_nat (Pos.to_nat (Pos.of_succ_nat e)))
r:radix
e:nat

IZR (Z.pow_pos r (Pos.of_succ_nat e)) = bpow (Z.of_nat (Pos.to_nat (Pos.of_succ_nat e)))
now rewrite <- Zpos_eq_Z_of_nat_o_nat_of_P. Qed.
r:radix

forall e : Z, (0 <= e)%Z -> IZR (r ^ e) = bpow e
r:radix

forall e : Z, (0 <= e)%Z -> IZR (r ^ e) = bpow e
r:radix
H:(0 <= 0)%Z

IZR (r ^ 0) = bpow 0
r:radix
e:positive
H:(0 <= Z.pos e)%Z
IZR (r ^ Z.pos e) = bpow (Z.pos e)
r:radix
e:positive
H:(0 <= Z.neg e)%Z
IZR (r ^ Z.neg e) = bpow (Z.neg e)
r:radix
e:positive
H:(0 <= Z.pos e)%Z

IZR (r ^ Z.pos e) = bpow (Z.pos e)
r:radix
e:positive
H:(0 <= Z.neg e)%Z
IZR (r ^ Z.neg e) = bpow (Z.neg e)
r:radix
e:positive
H:(0 <= Z.neg e)%Z

IZR (r ^ Z.neg e) = bpow (Z.neg e)
now elim H. Qed.
r:radix

forall e1 e2 : Z, (e1 < e2)%Z -> (bpow e1 < bpow e2)%R
r:radix

forall e1 e2 : Z, (e1 < e2)%Z -> (bpow e1 < bpow e2)%R
r:radix
e1, e2:Z
H:(e1 < e2)%Z

(bpow e1 < bpow e2)%R
r:radix
e1, e2:Z
H:(e1 < e2)%Z

(bpow e1 < bpow (e1 + (e2 - e1)))%R
r:radix
e1, e2:Z
H:(e1 < e2)%Z

(bpow e1 * 1 < bpow (e1 + (e2 - e1)))%R
r:radix
e1, e2:Z
H:(e1 < e2)%Z

(bpow e1 * 1 < bpow e1 * bpow (e2 - e1))%R
r:radix
e1, e2:Z
H:(e1 < e2)%Z

(0 < bpow e1)%R
r:radix
e1, e2:Z
H:(e1 < e2)%Z
(1 < bpow (e2 - e1))%R
r:radix
e1, e2:Z
H:(e1 < e2)%Z

(1 < bpow (e2 - e1))%R
r:radix
e1, e2:Z
H:(e1 < e2)%Z
H0:(0 < e2 - e1)%Z

(1 < bpow (e2 - e1))%R
r:radix
e1, e2:Z
H:(e1 < e2)%Z
p:positive
H0:(0 < Z.pos p)%Z

(1 < bpow (Z.pos p))%R
r:radix
p:positive

(1 < bpow (Z.pos p))%R
r:radix
p:positive

(1 < IZR (r ^ Z.pos p))%R
r:radix
p:positive

(1 < r ^ Z.pos p)%Z
now apply Zpower_gt_1. Qed.
r:radix

forall e1 e2 : Z, (bpow e1 < bpow e2)%R -> (e1 < e2)%Z
r:radix

forall e1 e2 : Z, (bpow e1 < bpow e2)%R -> (e1 < e2)%Z
r:radix
e1, e2:Z
H:(bpow e1 < bpow e2)%R

(e1 < e2)%Z
r:radix
e1, e2:Z
H:(bpow e1 < bpow e2)%R

(e2 > e1)%Z
r:radix
e1, e2:Z
H:(bpow e1 < bpow e2)%R

~ (e2 <= e1)%Z
r:radix
e1, e2:Z
H:(bpow e1 < bpow e2)%R
H':(e2 <= e1)%Z

False
r:radix
e1, e2:Z
H:(bpow e1 < bpow e2)%R
H':(e2 <= e1)%Z

(bpow e2 <= bpow e1)%R
r:radix
e1, e2:Z
H:(bpow e1 < bpow e2)%R
H':(e2 <= e1)%Z
H0:(e2 < e1)%Z

(bpow e2 <= bpow e1)%R
r:radix
e1, e2:Z
H:(bpow e1 < bpow e2)%R
H':(e2 <= e1)%Z
H0:e2 = e1
(bpow e2 <= bpow e1)%R
r:radix
e1, e2:Z
H:(bpow e1 < bpow e2)%R
H':(e2 <= e1)%Z
H0:(e2 < e1)%Z

(bpow e2 < bpow e1)%R
r:radix
e1, e2:Z
H:(bpow e1 < bpow e2)%R
H':(e2 <= e1)%Z
H0:e2 = e1
(bpow e2 <= bpow e1)%R
r:radix
e1, e2:Z
H:(bpow e1 < bpow e2)%R
H':(e2 <= e1)%Z
H0:e2 = e1

(bpow e2 <= bpow e1)%R
r:radix
e1, e2:Z
H:(bpow e1 < bpow e2)%R
H':(e2 <= e1)%Z
H0:e2 = e1

(bpow e1 <= bpow e1)%R
apply Rle_refl. Qed.
r:radix

forall e1 e2 : Z, (e1 <= e2)%Z -> (bpow e1 <= bpow e2)%R
r:radix

forall e1 e2 : Z, (e1 <= e2)%Z -> (bpow e1 <= bpow e2)%R
r:radix
e1, e2:Z
H:(e1 <= e2)%Z

(bpow e1 <= bpow e2)%R
r:radix
e1, e2:Z
H:(e1 <= e2)%Z

~ (bpow e2 < bpow e1)%R
r:radix
e1, e2:Z
H:(e1 <= e2)%Z
H':(bpow e2 < bpow e1)%R

False
r:radix
e1, e2:Z
H:(e1 <= e2)%Z
H':(bpow e2 < bpow e1)%R

(e1 > e2)%Z
r:radix
e1, e2:Z
H:(e1 <= e2)%Z
H':(bpow e2 < bpow e1)%R

(e2 < e1)%Z
now apply lt_bpow. Qed.
r:radix

forall e1 e2 : Z, (bpow e1 <= bpow e2)%R -> (e1 <= e2)%Z
r:radix

forall e1 e2 : Z, (bpow e1 <= bpow e2)%R -> (e1 <= e2)%Z
r:radix
e1, e2:Z
H:(bpow e1 <= bpow e2)%R

(e1 <= e2)%Z
r:radix
e1, e2:Z
H:(bpow e1 <= bpow e2)%R

~ (e1 > e2)%Z
r:radix
e1, e2:Z
H:(bpow e1 <= bpow e2)%R
H':(e1 > e2)%Z

False
r:radix
e1, e2:Z
H:(bpow e1 <= bpow e2)%R
H':(e1 > e2)%Z

(bpow e2 < bpow e1)%R
r:radix
e1, e2:Z
H:(bpow e1 <= bpow e2)%R
H':(e1 > e2)%Z

(e2 < e1)%Z
now apply Z.gt_lt. Qed.
r:radix

forall e1 e2 : Z, bpow e1 = bpow e2 -> e1 = e2
r:radix

forall e1 e2 : Z, bpow e1 = bpow e2 -> e1 = e2
r:radix
e1, e2:Z
H:bpow e1 = bpow e2

e1 = e2
r:radix
e1, e2:Z
H:bpow e1 = bpow e2

(e1 <= e2)%Z
r:radix
e1, e2:Z
H:bpow e1 = bpow e2
(e2 <= e1)%Z
r:radix
e1, e2:Z
H:bpow e1 = bpow e2

(bpow e1 <= bpow e2)%R
r:radix
e1, e2:Z
H:bpow e1 = bpow e2
(e2 <= e1)%Z
r:radix
e1, e2:Z
H:bpow e1 = bpow e2

(e2 <= e1)%Z
r:radix
e1, e2:Z
H:bpow e1 = bpow e2

(bpow e2 <= bpow e1)%R
now apply Req_le. Qed.
r:radix

forall e : Z, bpow e = exp (IZR e * ln (IZR r))
r:radix

forall e : Z, bpow e = exp (IZR e * ln (IZR r))
(* positive case *)
r:radix

forall e : positive, bpow (Z.pos e) = exp (IZR (Z.pos e) * ln (IZR r))
r:radix
H:forall e : positive, bpow (Z.pos e) = exp (IZR (Z.pos e) * ln (IZR r))
forall e : Z, bpow e = exp (IZR e * ln (IZR r))
r:radix
e:positive

bpow (Z.pos e) = exp (IZR (Z.pos e) * ln (IZR r))
r:radix
H:forall e : positive, bpow (Z.pos e) = exp (IZR (Z.pos e) * ln (IZR r))
forall e : Z, bpow e = exp (IZR e * ln (IZR r))
r:radix
e:positive

IZR (Z.pow_pos r e) = exp (IZR (Z.pos e) * ln (IZR r))
r:radix
H:forall e : positive, bpow (Z.pos e) = exp (IZR (Z.pos e) * ln (IZR r))
forall e : Z, bpow e = exp (IZR e * ln (IZR r))
r:radix
e:positive

IZR (Zpower_nat r (Pos.to_nat e)) = exp (IZR (Z.pos e) * ln (IZR r))
r:radix
H:forall e : positive, bpow (Z.pos e) = exp (IZR (Z.pos e) * ln (IZR r))
forall e : Z, bpow e = exp (IZR e * ln (IZR r))
r:radix
e:positive

IZR (Zpower_nat r (Pos.to_nat e)) = exp (IZR (Z.of_nat (Pos.to_nat e)) * ln (IZR r))
r:radix
H:forall e : positive, bpow (Z.pos e) = exp (IZR (Z.pos e) * ln (IZR r))
forall e : Z, bpow e = exp (IZR e * ln (IZR r))
r:radix
e:positive

IZR (Zpower_nat r (Pos.to_nat e)) = exp (INR (Pos.to_nat e) * ln (IZR r))
r:radix
H:forall e : positive, bpow (Z.pos e) = exp (IZR (Z.pos e) * ln (IZR r))
forall e : Z, bpow e = exp (IZR e * ln (IZR r))
r:radix
e:positive

IZR (Zpower_nat r 0) = exp (INR 0 * ln (IZR r))
r:radix
e:positive
n:nat
IHn:IZR (Zpower_nat r n) = exp (INR n * ln (IZR r))
IZR (Zpower_nat r (S n)) = exp (INR (S n) * ln (IZR r))
r:radix
H:forall e : positive, bpow (Z.pos e) = exp (IZR (Z.pos e) * ln (IZR r))
forall e : Z, bpow e = exp (IZR e * ln (IZR r))
r:radix
e:positive

IZR (Zpower_nat r 0) = exp 0
r:radix
e:positive
n:nat
IHn:IZR (Zpower_nat r n) = exp (INR n * ln (IZR r))
IZR (Zpower_nat r (S n)) = exp (INR (S n) * ln (IZR r))
r:radix
H:forall e : positive, bpow (Z.pos e) = exp (IZR (Z.pos e) * ln (IZR r))
forall e : Z, bpow e = exp (IZR e * ln (IZR r))
r:radix
e:positive
n:nat
IHn:IZR (Zpower_nat r n) = exp (INR n * ln (IZR r))

IZR (Zpower_nat r (S n)) = exp (INR (S n) * ln (IZR r))
r:radix
H:forall e : positive, bpow (Z.pos e) = exp (IZR (Z.pos e) * ln (IZR r))
forall e : Z, bpow e = exp (IZR e * ln (IZR r))
r:radix
e:positive
n:nat
IHn:IZR (Zpower_nat r n) = exp (INR n * ln (IZR r))

IZR (r * Zpower_nat r n) = exp (INR (S n) * ln (IZR r))
r:radix
H:forall e : positive, bpow (Z.pos e) = exp (IZR (Z.pos e) * ln (IZR r))
forall e : Z, bpow e = exp (IZR e * ln (IZR r))
r:radix
e:positive
n:nat
IHn:IZR (Zpower_nat r n) = exp (INR n * ln (IZR r))

IZR (r * Zpower_nat r n) = exp ((INR n + 1) * ln (IZR r))
r:radix
H:forall e : positive, bpow (Z.pos e) = exp (IZR (Z.pos e) * ln (IZR r))
forall e : Z, bpow e = exp (IZR e * ln (IZR r))
r:radix
e:positive
n:nat
IHn:IZR (Zpower_nat r n) = exp (INR n * ln (IZR r))

IZR (r * Zpower_nat r n) = exp (INR n * ln (IZR r) + 1 * ln (IZR r))
r:radix
H:forall e : positive, bpow (Z.pos e) = exp (IZR (Z.pos e) * ln (IZR r))
forall e : Z, bpow e = exp (IZR e * ln (IZR r))
r:radix
e:positive
n:nat
IHn:IZR (Zpower_nat r n) = exp (INR n * ln (IZR r))

IZR (r * Zpower_nat r n) = (exp (INR n * ln (IZR r)) * exp (1 * ln (IZR r)))%R
r:radix
H:forall e : positive, bpow (Z.pos e) = exp (IZR (Z.pos e) * ln (IZR r))
forall e : Z, bpow e = exp (IZR e * ln (IZR r))
r:radix
e:positive
n:nat
IHn:IZR (Zpower_nat r n) = exp (INR n * ln (IZR r))

IZR (r * Zpower_nat r n) = (exp (INR n * ln (IZR r)) * exp (ln (IZR r)))%R
r:radix
H:forall e : positive, bpow (Z.pos e) = exp (IZR (Z.pos e) * ln (IZR r))
forall e : Z, bpow e = exp (IZR e * ln (IZR r))
r:radix
e:positive
n:nat
IHn:IZR (Zpower_nat r n) = exp (INR n * ln (IZR r))

IZR (r * Zpower_nat r n) = (exp (INR n * ln (IZR r)) * IZR r)%R
r:radix
e:positive
n:nat
IHn:IZR (Zpower_nat r n) = exp (INR n * ln (IZR r))
(0 < IZR r)%R
r:radix
H:forall e : positive, bpow (Z.pos e) = exp (IZR (Z.pos e) * ln (IZR r))
forall e : Z, bpow e = exp (IZR e * ln (IZR r))
r:radix
e:positive
n:nat
IHn:IZR (Zpower_nat r n) = exp (INR n * ln (IZR r))

IZR (r * Zpower_nat r n) = (IZR (Zpower_nat r n) * IZR r)%R
r:radix
e:positive
n:nat
IHn:IZR (Zpower_nat r n) = exp (INR n * ln (IZR r))
(0 < IZR r)%R
r:radix
H:forall e : positive, bpow (Z.pos e) = exp (IZR (Z.pos e) * ln (IZR r))
forall e : Z, bpow e = exp (IZR e * ln (IZR r))
r:radix
e:positive
n:nat
IHn:IZR (Zpower_nat r n) = exp (INR n * ln (IZR r))

IZR (r * Zpower_nat r n) = IZR (Zpower_nat r n * r)
r:radix
e:positive
n:nat
IHn:IZR (Zpower_nat r n) = exp (INR n * ln (IZR r))
(0 < IZR r)%R
r:radix
H:forall e : positive, bpow (Z.pos e) = exp (IZR (Z.pos e) * ln (IZR r))
forall e : Z, bpow e = exp (IZR e * ln (IZR r))
r:radix
e:positive
n:nat
IHn:IZR (Zpower_nat r n) = exp (INR n * ln (IZR r))

(0 < IZR r)%R
r:radix
H:forall e : positive, bpow (Z.pos e) = exp (IZR (Z.pos e) * ln (IZR r))
forall e : Z, bpow e = exp (IZR e * ln (IZR r))
r:radix
H:forall e : positive, bpow (Z.pos e) = exp (IZR (Z.pos e) * ln (IZR r))

forall e : Z, bpow e = exp (IZR e * ln (IZR r))
(* general case *)
r:radix
H:forall e : positive, bpow (Z.pos e) = exp (IZR (Z.pos e) * ln (IZR r))

bpow 0 = exp (0 * ln (IZR r))
r:radix
H:forall e0 : positive, bpow (Z.pos e0) = exp (IZR (Z.pos e0) * ln (IZR r))
e:positive
bpow (Z.pos e) = exp (IZR (Z.pos e) * ln (IZR r))
r:radix
H:forall e0 : positive, bpow (Z.pos e0) = exp (IZR (Z.pos e0) * ln (IZR r))
e:positive
bpow (Z.neg e) = exp (IZR (Z.neg e) * ln (IZR r))
r:radix
H:forall e : positive, bpow (Z.pos e) = exp (IZR (Z.pos e) * ln (IZR r))

bpow 0 = exp 0
r:radix
H:forall e0 : positive, bpow (Z.pos e0) = exp (IZR (Z.pos e0) * ln (IZR r))
e:positive
bpow (Z.pos e) = exp (IZR (Z.pos e) * ln (IZR r))
r:radix
H:forall e0 : positive, bpow (Z.pos e0) = exp (IZR (Z.pos e0) * ln (IZR r))
e:positive
bpow (Z.neg e) = exp (IZR (Z.neg e) * ln (IZR r))
r:radix
H:forall e0 : positive, bpow (Z.pos e0) = exp (IZR (Z.pos e0) * ln (IZR r))
e:positive

bpow (Z.pos e) = exp (IZR (Z.pos e) * ln (IZR r))
r:radix
H:forall e0 : positive, bpow (Z.pos e0) = exp (IZR (Z.pos e0) * ln (IZR r))
e:positive
bpow (Z.neg e) = exp (IZR (Z.neg e) * ln (IZR r))
r:radix
H:forall e0 : positive, bpow (Z.pos e0) = exp (IZR (Z.pos e0) * ln (IZR r))
e:positive

bpow (Z.neg e) = exp (IZR (Z.neg e) * ln (IZR r))
r:radix
H:forall e0 : positive, bpow (Z.pos e0) = exp (IZR (Z.pos e0) * ln (IZR r))
e:positive

(/ IZR (Z.pow_pos r e))%R = exp (IZR (Z.neg e) * ln (IZR r))
r:radix
H:forall e0 : positive, bpow (Z.pos e0) = exp (IZR (Z.pos e0) * ln (IZR r))
e:positive

(/ bpow (Z.pos e))%R = exp (IZR (Z.neg e) * ln (IZR r))
r:radix
H:forall e0 : positive, bpow (Z.pos e0) = exp (IZR (Z.pos e0) * ln (IZR r))
e:positive

(/ exp (IZR (Z.pos e) * ln (IZR r)))%R = exp (IZR (Z.neg e) * ln (IZR r))
r:radix
H:forall e0 : positive, bpow (Z.pos e0) = exp (IZR (Z.pos e0) * ln (IZR r))
e:positive

exp (- (IZR (Z.pos e) * ln (IZR r))) = exp (IZR (Z.neg e) * ln (IZR r))
r:radix
H:forall e0 : positive, bpow (Z.pos e0) = exp (IZR (Z.pos e0) * ln (IZR r))
e:positive

exp (- IZR (Z.pos e) * ln (IZR r)) = exp (IZR (Z.neg e) * ln (IZR r))
now rewrite <- opp_IZR. Qed.
r:radix

forall e : Z, sqrt (bpow (2 * e)) = bpow e
r:radix

forall e : Z, sqrt (bpow (2 * e)) = bpow e
r:radix
e:Z

sqrt (bpow (2 * e)) = bpow e
r:radix
e:Z

sqrt (bpow e * bpow e) = bpow e
apply sqrt_square, bpow_ge_0. Qed.
r:radix

forall e : Z, (bpow (e / 2) <= sqrt (bpow e))%R
r:radix

forall e : Z, (bpow (e / 2) <= sqrt (bpow e))%R
r:radix
e:Z

(bpow (e / 2) <= sqrt (bpow e))%R
r:radix
e:Z

(sqrt (bpow (e / 2) * bpow (e / 2)) <= sqrt (bpow e))%R
r:radix
e:Z

(e / 2 + e / 2 <= e)%Z
now replace (_ + _)%Z with (2 * (e / 2))%Z by ring; apply Z_mult_div_ge. Qed.
Another well-used function for having the magnitude of a real number x to the base β
Record mag_prop x := {
  mag_val :> Z ;
  _ : (x <> 0)%R -> (bpow (mag_val - 1)%Z <= Rabs x < bpow mag_val)%R
}.

r:radix

forall x : R, mag_prop x
r:radix

forall x : R, mag_prop x
r:radix
x:R

mag_prop x
r:radix
x:R
fact:=ln (IZR r):R

mag_prop x
(* . *)
r:radix
x:R
fact:=ln (IZR r):R

(0 < fact)%R
r:radix
x:R
fact:=ln (IZR r):R
H:(0 < fact)%R
mag_prop x
r:radix
x:R
fact:=ln (IZR r):R

(exp 0 < exp fact)%R
r:radix
x:R
fact:=ln (IZR r):R
H:(0 < fact)%R
mag_prop x
r:radix
x:R
fact:=ln (IZR r):R

(1 < exp fact)%R
r:radix
x:R
fact:=ln (IZR r):R
H:(0 < fact)%R
mag_prop x
r:radix
x:R
fact:=ln (IZR r):R

(1 < exp (ln (IZR r)))%R
r:radix
x:R
fact:=ln (IZR r):R
H:(0 < fact)%R
mag_prop x
r:radix
x:R
fact:=ln (IZR r):R

(1 < IZR r)%R
r:radix
x:R
fact:=ln (IZR r):R
(0 < IZR r)%R
r:radix
x:R
fact:=ln (IZR r):R
H:(0 < fact)%R
mag_prop x
r:radix
x:R
fact:=ln (IZR r):R

(1 < r)%Z
r:radix
x:R
fact:=ln (IZR r):R
(0 < IZR r)%R
r:radix
x:R
fact:=ln (IZR r):R
H:(0 < fact)%R
mag_prop x
r:radix
x:R
fact:=ln (IZR r):R

(0 < IZR r)%R
r:radix
x:R
fact:=ln (IZR r):R
H:(0 < fact)%R
mag_prop x
r:radix
x:R
fact:=ln (IZR r):R
H:(0 < fact)%R

mag_prop x
(* . *)
r:radix
x:R
fact:=ln (IZR r):R
H:(0 < fact)%R

x <> 0%R -> (bpow (Zfloor (ln (Rabs x) / fact) + 1 - 1) <= Rabs x < bpow (Zfloor (ln (Rabs x) / fact) + 1))%R
r:radix
x:R
fact:=ln (IZR r):R
H:(0 < fact)%R
Hx':x <> 0%R

(bpow (Zfloor (ln (Rabs x) / fact) + 1 - 1) <= Rabs x < bpow (Zfloor (ln (Rabs x) / fact) + 1))%R
r:radix
x:R
fact:=ln (IZR r):R
H:(0 < fact)%R
Hx':x <> 0%R

(0 < Rabs x)%R -> (bpow (Zfloor (ln (Rabs x) / fact) + 1 - 1) <= Rabs x < bpow (Zfloor (ln (Rabs x) / fact) + 1))%R
r:radix
x:R
fact:=ln (IZR r):R
H:(0 < fact)%R

(0 < Rabs x)%R -> (bpow (Zfloor (ln (Rabs x) / fact) + 1 - 1) <= Rabs x < bpow (Zfloor (ln (Rabs x) / fact) + 1))%R
r:radix
x:R
fact:=ln (IZR r):R
H:(0 < fact)%R

forall r0 : R, (0 < r0)%R -> (bpow (Zfloor (ln r0 / fact) + 1 - 1) <= r0 < bpow (Zfloor (ln r0 / fact) + 1))%R
r:radix
fact:=ln (IZR r):R
H:(0 < fact)%R

forall r0 : R, (0 < r0)%R -> (bpow (Zfloor (ln r0 / fact) + 1 - 1) <= r0 < bpow (Zfloor (ln r0 / fact) + 1))%R
r:radix
fact:=ln (IZR r):R
H:(0 < fact)%R
x:R
Hx:(0 < x)%R

(bpow (Zfloor (ln x / fact) + 1 - 1) <= x < bpow (Zfloor (ln x / fact) + 1))%R
r:radix
fact:=ln (IZR r):R
H:(0 < fact)%R
x:R
Hx:(0 < x)%R

(exp (IZR (Zfloor (ln x / fact) + 1 - 1) * ln (IZR r)) <= x < exp (IZR (Zfloor (ln x / fact) + 1) * ln (IZR r)))%R
r:radix
fact:=ln (IZR r):R
H:(0 < fact)%R
x:R
Hx:(0 < x)%R

(exp (IZR (Zfloor (ln x / fact) + 1 - 1) * fact) <= x < exp (IZR (Zfloor (ln x / fact) + 1) * fact))%R
r:radix
fact:=ln (IZR r):R
H:(0 < fact)%R
x:R
Hx:(0 < x)%R

(exp (IZR (Zfloor (ln x / fact) + 1 - 1) * fact) <= exp (ln x * / fact * fact) < exp (IZR (Zfloor (ln x / fact) + 1) * fact))%R
r:radix
fact:=ln (IZR r):R
H:(0 < fact)%R
x:R
Hx:(0 < x)%R
exp (ln x * / fact * fact) = x
r:radix
fact:=ln (IZR r):R
H:(0 < fact)%R
x:R
Hx:(0 < x)%R

(exp (IZR (Zfloor (ln x / fact) + 1 - 1) * fact) <= exp (ln x * / fact * fact))%R
r:radix
fact:=ln (IZR r):R
H:(0 < fact)%R
x:R
Hx:(0 < x)%R
(exp (ln x * / fact * fact) < exp (IZR (Zfloor (ln x / fact) + 1) * fact))%R
r:radix
fact:=ln (IZR r):R
H:(0 < fact)%R
x:R
Hx:(0 < x)%R
exp (ln x * / fact * fact) = x
r:radix
fact:=ln (IZR r):R
H:(0 < fact)%R
x:R
Hx:(0 < x)%R

(exp ((IZR (Zfloor (ln x / fact) + 1) - 1) * fact) <= exp (ln x * / fact * fact))%R
r:radix
fact:=ln (IZR r):R
H:(0 < fact)%R
x:R
Hx:(0 < x)%R
(exp (ln x * / fact * fact) < exp (IZR (Zfloor (ln x / fact) + 1) * fact))%R
r:radix
fact:=ln (IZR r):R
H:(0 < fact)%R
x:R
Hx:(0 < x)%R
exp (ln x * / fact * fact) = x
r:radix
fact:=ln (IZR r):R
H:(0 < fact)%R
x:R
Hx:(0 < x)%R

((IZR (Zfloor (ln x / fact) + 1) - 1) * fact <= ln x * / fact * fact)%R
r:radix
fact:=ln (IZR r):R
H:(0 < fact)%R
x:R
Hx:(0 < x)%R
(exp (ln x * / fact * fact) < exp (IZR (Zfloor (ln x / fact) + 1) * fact))%R
r:radix
fact:=ln (IZR r):R
H:(0 < fact)%R
x:R
Hx:(0 < x)%R
exp (ln x * / fact * fact) = x
r:radix
fact:=ln (IZR r):R
H:(0 < fact)%R
x:R
Hx:(0 < x)%R

(0 <= fact)%R
r:radix
fact:=ln (IZR r):R
H:(0 < fact)%R
x:R
Hx:(0 < x)%R
(IZR (Zfloor (ln x / fact) + 1) - 1 <= ln x * / fact)%R
r:radix
fact:=ln (IZR r):R
H:(0 < fact)%R
x:R
Hx:(0 < x)%R
(exp (ln x * / fact * fact) < exp (IZR (Zfloor (ln x / fact) + 1) * fact))%R
r:radix
fact:=ln (IZR r):R
H:(0 < fact)%R
x:R
Hx:(0 < x)%R
exp (ln x * / fact * fact) = x
r:radix
fact:=ln (IZR r):R
H:(0 < fact)%R
x:R
Hx:(0 < x)%R

(IZR (Zfloor (ln x / fact) + 1) - 1 <= ln x * / fact)%R
r:radix
fact:=ln (IZR r):R
H:(0 < fact)%R
x:R
Hx:(0 < x)%R
(exp (ln x * / fact * fact) < exp (IZR (Zfloor (ln x / fact) + 1) * fact))%R
r:radix
fact:=ln (IZR r):R
H:(0 < fact)%R
x:R
Hx:(0 < x)%R
exp (ln x * / fact * fact) = x
r:radix
fact:=ln (IZR r):R
H:(0 < fact)%R
x:R
Hx:(0 < x)%R

(IZR (Zfloor (ln x / fact) + 1) + - (1) <= ln x * / fact)%R
r:radix
fact:=ln (IZR r):R
H:(0 < fact)%R
x:R
Hx:(0 < x)%R
(exp (ln x * / fact * fact) < exp (IZR (Zfloor (ln x / fact) + 1) * fact))%R
r:radix
fact:=ln (IZR r):R
H:(0 < fact)%R
x:R
Hx:(0 < x)%R
exp (ln x * / fact * fact) = x
r:radix
fact:=ln (IZR r):R
H:(0 < fact)%R
x:R
Hx:(0 < x)%R

(IZR (Zfloor (ln x / fact)) + 1 + - (1) <= ln x * / fact)%R
r:radix
fact:=ln (IZR r):R
H:(0 < fact)%R
x:R
Hx:(0 < x)%R
(exp (ln x * / fact * fact) < exp (IZR (Zfloor (ln x / fact) + 1) * fact))%R
r:radix
fact:=ln (IZR r):R
H:(0 < fact)%R
x:R
Hx:(0 < x)%R
exp (ln x * / fact * fact) = x
r:radix
fact:=ln (IZR r):R
H:(0 < fact)%R
x:R
Hx:(0 < x)%R

(IZR (Zfloor (ln x / fact)) + (1 + - (1)) <= ln x * / fact)%R
r:radix
fact:=ln (IZR r):R
H:(0 < fact)%R
x:R
Hx:(0 < x)%R
(exp (ln x * / fact * fact) < exp (IZR (Zfloor (ln x / fact) + 1) * fact))%R
r:radix
fact:=ln (IZR r):R
H:(0 < fact)%R
x:R
Hx:(0 < x)%R
exp (ln x * / fact * fact) = x
r:radix
fact:=ln (IZR r):R
H:(0 < fact)%R
x:R
Hx:(0 < x)%R

(IZR (Zfloor (ln x / fact)) <= ln x * / fact)%R
r:radix
fact:=ln (IZR r):R
H:(0 < fact)%R
x:R
Hx:(0 < x)%R
(exp (ln x * / fact * fact) < exp (IZR (Zfloor (ln x / fact) + 1) * fact))%R
r:radix
fact:=ln (IZR r):R
H:(0 < fact)%R
x:R
Hx:(0 < x)%R
exp (ln x * / fact * fact) = x
r:radix
fact:=ln (IZR r):R
H:(0 < fact)%R
x:R
Hx:(0 < x)%R

(exp (ln x * / fact * fact) < exp (IZR (Zfloor (ln x / fact) + 1) * fact))%R
r:radix
fact:=ln (IZR r):R
H:(0 < fact)%R
x:R
Hx:(0 < x)%R
exp (ln x * / fact * fact) = x
r:radix
fact:=ln (IZR r):R
H:(0 < fact)%R
x:R
Hx:(0 < x)%R

(ln x * / fact * fact < IZR (Zfloor (ln x / fact) + 1) * fact)%R
r:radix
fact:=ln (IZR r):R
H:(0 < fact)%R
x:R
Hx:(0 < x)%R
exp (ln x * / fact * fact) = x
r:radix
fact:=ln (IZR r):R
H:(0 < fact)%R
x:R
Hx:(0 < x)%R

(0 < fact)%R
r:radix
fact:=ln (IZR r):R
H:(0 < fact)%R
x:R
Hx:(0 < x)%R
(ln x * / fact < IZR (Zfloor (ln x / fact) + 1))%R
r:radix
fact:=ln (IZR r):R
H:(0 < fact)%R
x:R
Hx:(0 < x)%R
exp (ln x * / fact * fact) = x
r:radix
fact:=ln (IZR r):R
H:(0 < fact)%R
x:R
Hx:(0 < x)%R

(ln x * / fact < IZR (Zfloor (ln x / fact) + 1))%R
r:radix
fact:=ln (IZR r):R
H:(0 < fact)%R
x:R
Hx:(0 < x)%R
exp (ln x * / fact * fact) = x
r:radix
fact:=ln (IZR r):R
H:(0 < fact)%R
x:R
Hx:(0 < x)%R

(ln x * / fact < IZR (Zfloor (ln x / fact)) + 1)%R
r:radix
fact:=ln (IZR r):R
H:(0 < fact)%R
x:R
Hx:(0 < x)%R
exp (ln x * / fact * fact) = x
r:radix
fact:=ln (IZR r):R
H:(0 < fact)%R
x:R
Hx:(0 < x)%R

exp (ln x * / fact * fact) = x
r:radix
fact:=ln (IZR r):R
H:(0 < fact)%R
x:R
Hx:(0 < x)%R

exp (ln x * (/ fact * fact)) = x
r:radix
fact:=ln (IZR r):R
H:(0 < fact)%R
x:R
Hx:(0 < x)%R

exp (ln x * 1) = x
r:radix
fact:=ln (IZR r):R
H:(0 < fact)%R
x:R
Hx:(0 < x)%R
fact <> 0%R
r:radix
fact:=ln (IZR r):R
H:(0 < fact)%R
x:R
Hx:(0 < x)%R

exp (ln x) = x
r:radix
fact:=ln (IZR r):R
H:(0 < fact)%R
x:R
Hx:(0 < x)%R
fact <> 0%R
r:radix
fact:=ln (IZR r):R
H:(0 < fact)%R
x:R
Hx:(0 < x)%R

fact <> 0%R
now apply Rgt_not_eq. Qed.
r:radix

forall e1 e2 : Z, (bpow (e1 - 1) < bpow e2)%R -> (e1 <= e2)%Z
r:radix

forall e1 e2 : Z, (bpow (e1 - 1) < bpow e2)%R -> (e1 <= e2)%Z
r:radix
e1, e2:Z
He:(bpow (e1 - 1) < bpow e2)%R

(e1 <= e2)%Z
r:radix
e1, e2:Z
He:(bpow (e1 - 1) < bpow e2)%R

(Z.succ (Z.pred e1) <= e2)%Z
r:radix
e1, e2:Z
He:(bpow (e1 - 1) < bpow e2)%R

(Z.pred e1 < e2)%Z
now apply lt_bpow. Qed.
r:radix

forall (x : R) (e1 e2 : Z), (bpow (e1 - 1) <= x < bpow e1)%R -> (bpow (e2 - 1) <= x < bpow e2)%R -> e1 = e2
r:radix

forall (x : R) (e1 e2 : Z), (bpow (e1 - 1) <= x < bpow e1)%R -> (bpow (e2 - 1) <= x < bpow e2)%R -> e1 = e2
r:radix
x:R
e1, e2:Z
H1a:(bpow (e1 - 1) <= x)%R
H1b:(x < bpow e1)%R
H2a:(bpow (e2 - 1) <= x)%R
H2b:(x < bpow e2)%R

e1 = e2
apply Zle_antisym ; apply bpow_lt_bpow ; apply Rle_lt_trans with x ; assumption. Qed.
r:radix

forall (x : R) (e : Z), (bpow (e - 1) <= Rabs x < bpow e)%R -> mag x = e
r:radix

forall (x : R) (e : Z), (bpow (e - 1) <= Rabs x < bpow e)%R -> mag x = e
r:radix
x:R
e1:Z
He:(bpow (e1 - 1) <= Rabs x < bpow e1)%R

mag x = e1
r:radix
x:R
e1:Z
He:(bpow (e1 - 1) <= Rabs x < bpow e1)%R
Hx:x = 0%R

mag x = e1
r:radix
x:R
e1:Z
He:(bpow (e1 - 1) <= Rabs x < bpow e1)%R
Hx:x <> 0%R
mag x = e1
r:radix
x:R
e1:Z
He:(bpow (e1 - 1) <= Rabs x < bpow e1)%R
Hx:x = 0%R

(Rabs x < bpow (e1 - 1))%R
r:radix
x:R
e1:Z
He:(bpow (e1 - 1) <= Rabs x < bpow e1)%R
Hx:x <> 0%R
mag x = e1
r:radix
x:R
e1:Z
He:(bpow (e1 - 1) <= Rabs x < bpow e1)%R
Hx:x = 0%R

(0 < bpow (e1 - 1))%R
r:radix
x:R
e1:Z
He:(bpow (e1 - 1) <= Rabs x < bpow e1)%R
Hx:x <> 0%R
mag x = e1
r:radix
x:R
e1:Z
He:(bpow (e1 - 1) <= Rabs x < bpow e1)%R
Hx:x <> 0%R

mag x = e1
r:radix
x:R
e1:Z
He:(bpow (e1 - 1) <= Rabs x < bpow e1)%R
Hx:x <> 0%R
e2:Z
Hx2:x <> 0%R -> (bpow (e2 - 1) <= Rabs x < bpow e2)%R

Build_mag_prop x e2 Hx2 = e1
r:radix
x:R
e1:Z
He:(bpow (e1 - 1) <= Rabs x < bpow e1)%R
Hx:x <> 0%R
e2:Z
Hx2:x <> 0%R -> (bpow (e2 - 1) <= Rabs x < bpow e2)%R

e2 = e1
r:radix
x:R
e1:Z
He:(bpow (e1 - 1) <= Rabs x < bpow e1)%R
Hx:x <> 0%R
e2:Z
Hx2:x <> 0%R -> (bpow (e2 - 1) <= Rabs x < bpow e2)%R

(bpow (e2 - 1) <= Rabs x < bpow e2)%R
now apply Hx2. Qed.
r:radix

forall x : R, mag (- x) = mag x
r:radix

forall x : R, mag (- x) = mag x
r:radix
x:R

mag (- x) = mag x
r:radix
x:R
Hx:x = 0%R

mag (- x) = mag x
r:radix
x:R
Hx:x <> 0%R
mag (- x) = mag x
r:radix
x:R
Hx:x <> 0%R

mag (- x) = mag x
r:radix
x:R
Hx:x <> 0%R
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R

mag (- x) = Build_mag_prop x e He
r:radix
x:R
Hx:x <> 0%R
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R

mag (- x) = e
r:radix
x:R
Hx:x <> 0%R
e:Z
He:(bpow (e - 1) <= Rabs x < bpow e)%R

mag (- x) = e
r:radix
x:R
Hx:x <> 0%R
e:Z
He:(bpow (e - 1) <= Rabs x < bpow e)%R

(bpow (e - 1) <= Rabs (- x) < bpow e)%R
now rewrite Rabs_Ropp. Qed.
r:radix

forall x : R, mag (Rabs x) = mag x
r:radix

forall x : R, mag (Rabs x) = mag x
r:radix
x:R

mag (Rabs x) = mag x
r:radix
x:R

mag (if Rcase_abs x then (- x)%R else x) = mag x
r:radix
x:R

mag (- x) = mag x
r:radix
x:R
mag x = mag x
r:radix
x:R

mag x = mag x
apply refl_equal. Qed.
r:radix

forall (x : R) (e : Z), (bpow (e - 1) <= x < bpow e)%R -> mag x = e
r:radix

forall (x : R) (e : Z), (bpow (e - 1) <= x < bpow e)%R -> mag x = e
r:radix
x:R
e1:Z
He1:(bpow (e1 - 1) <= x < bpow e1)%R

mag x = e1
r:radix
x:R
e1:Z
He1:(bpow (e1 - 1) <= x < bpow e1)%R

mag (Rabs x) = e1
r:radix
x:R
e1:Z
He1:(bpow (e1 - 1) <= x < bpow e1)%R

(bpow (e1 - 1) <= Rabs (Rabs x) < bpow e1)%R
r:radix
x:R
e1:Z
He1:(bpow (e1 - 1) <= x < bpow e1)%R

(bpow (e1 - 1) <= x < bpow e1)%R
r:radix
x:R
e1:Z
He1:(bpow (e1 - 1) <= x < bpow e1)%R
(0 <= x)%R
r:radix
x:R
e1:Z
He1:(bpow (e1 - 1) <= x < bpow e1)%R
(0 <= Rabs x)%R
r:radix
x:R
e1:Z
He1:(bpow (e1 - 1) <= x < bpow e1)%R

(0 <= x)%R
r:radix
x:R
e1:Z
He1:(bpow (e1 - 1) <= x < bpow e1)%R
(0 <= Rabs x)%R
r:radix
x:R
e1:Z
He1:(bpow (e1 - 1) <= x < bpow e1)%R

(0 <= bpow (e1 - 1))%R
r:radix
x:R
e1:Z
He1:(bpow (e1 - 1) <= x < bpow e1)%R
(0 <= Rabs x)%R
r:radix
x:R
e1:Z
He1:(bpow (e1 - 1) <= x < bpow e1)%R

(0 <= Rabs x)%R
apply Rabs_pos. Qed.
r:radix

forall x y : R, x <> 0%R -> (Rabs x <= Rabs y)%R -> (mag x <= mag y)%Z
r:radix

forall x y : R, x <> 0%R -> (Rabs x <= Rabs y)%R -> (mag x <= mag y)%Z
r:radix
x, y:R
H0x:x <> 0%R
Hxy:(Rabs x <= Rabs y)%R

(mag x <= mag y)%Z
r:radix
x, y:R
H0x:x <> 0%R
Hxy:(Rabs x <= Rabs y)%R
ex:Z
Hx:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R

(Build_mag_prop x ex Hx <= mag y)%Z
r:radix
x, y:R
H0x:x <> 0%R
Hxy:(Rabs x <= Rabs y)%R
ex:Z
Hx:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hy:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R

(Build_mag_prop x ex Hx <= Build_mag_prop y ey Hy)%Z
r:radix
x, y:R
H0x:x <> 0%R
Hxy:(Rabs x <= Rabs y)%R
ex:Z
Hx:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hy:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R

(ex <= ey)%Z
r:radix
x, y:R
H0x:x <> 0%R
Hxy:(Rabs x <= Rabs y)%R
ex:Z
Hx:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hy:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R

(bpow (ex - 1) < bpow ey)%R
r:radix
x, y:R
H0x:x <> 0%R
Hxy:(Rabs x <= Rabs y)%R
ex:Z
Hx:(bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hy:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R

(bpow (ex - 1) < bpow ey)%R
r:radix
x, y:R
H0x:x <> 0%R
Hxy:(Rabs x <= Rabs y)%R
ex:Z
Hx:(bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hy:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R

(Rabs x < bpow ey)%R
r:radix
x, y:R
H0x:x <> 0%R
Hxy:(Rabs x <= Rabs y)%R
ex:Z
Hx:(bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hy:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R

(Rabs y < bpow ey)%R
r:radix
x, y:R
H0x:x <> 0%R
Hxy:(Rabs x <= Rabs y)%R
ex:Z
Hx:(bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hy:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R

y <> 0%R
r:radix
x, y:R
H0x:x <> 0%R
Hxy:(Rabs x <= Rabs y)%R
ex:Z
Hx:(bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hy:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy':y = 0%R

False
r:radix
x, y:R
H0x:x <> 0%R
Hxy:(Rabs x <= Rabs y)%R
ex:Z
Hx:(bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hy:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy':y = 0%R

(Rabs x <= 0)%R
r:radix
x, y:R
H0x:x <> 0%R
Hxy:(Rabs x <= Rabs y)%R
ex:Z
Hx:(bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hy:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy':y = 0%R

(Rabs y <= 0)%R
r:radix
x, y:R
H0x:x <> 0%R
Hxy:(Rabs x <= Rabs y)%R
ex:Z
Hx:(bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hy:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hy':y = 0%R

(0 <= 0)%R
apply Rle_refl. Qed.
r:radix

forall x y : R, (0 < x)%R -> (x <= y)%R -> (mag x <= mag y)%Z
r:radix

forall x y : R, (0 < x)%R -> (x <= y)%R -> (mag x <= mag y)%Z
r:radix
x, y:R
H0x:(0 < x)%R
Hxy:(x <= y)%R

(mag x <= mag y)%Z
r:radix
x, y:R
H0x:(0 < x)%R
Hxy:(x <= y)%R

x <> 0%R
r:radix
x, y:R
H0x:(0 < x)%R
Hxy:(x <= y)%R
(Rabs x <= Rabs y)%R
r:radix
x, y:R
H0x:(0 < x)%R
Hxy:(x <= y)%R

(Rabs x <= Rabs y)%R
r:radix
x, y:R
H0x:(0 < x)%R
Hxy:(x <= y)%R

(x <= y)%R
r:radix
x, y:R
H0x:(0 < x)%R
Hxy:(x <= y)%R
(0 <= y)%R
r:radix
x, y:R
H0x:(0 < x)%R
Hxy:(x <= y)%R
(0 <= x)%R
r:radix
x, y:R
H0x:(0 < x)%R
Hxy:(x <= y)%R

(0 <= y)%R
r:radix
x, y:R
H0x:(0 < x)%R
Hxy:(x <= y)%R
(0 <= x)%R
r:radix
x, y:R
H0x:(0 < x)%R
Hxy:(x <= y)%R

(0 <= x)%R
r:radix
x, y:R
H0x:(0 < x)%R
Hxy:(x <= y)%R
(0 <= x)%R
r:radix
x, y:R
H0x:(0 < x)%R
Hxy:(x <= y)%R

(0 <= x)%R
now apply Rlt_le. Qed.
r:radix

forall x y : R, (0 < y)%R -> (mag x < mag y)%Z -> (x < y)%R
r:radix

forall x y : R, (0 < y)%R -> (mag x < mag y)%Z -> (x < y)%R
r:radix
x, y:R
Py:(0 < y)%R

(mag x < mag y)%Z -> (x < y)%R
r:radix
x, y:R
Py:(0 < y)%R
Px:(x <= 0)%R

(mag x < mag y)%Z -> (x < y)%R
r:radix
x, y:R
Py:(0 < y)%R
Px:(0 < x)%R
(mag x < mag y)%Z -> (x < y)%R
r:radix
x, y:R
Py:(0 < y)%R
Px:(x <= 0)%R
H:(mag x < mag y)%Z

(x < y)%R
r:radix
x, y:R
Py:(0 < y)%R
Px:(0 < x)%R
(mag x < mag y)%Z -> (x < y)%R
r:radix
x, y:R
Py:(0 < y)%R
Px:(0 < x)%R

(mag x < mag y)%Z -> (x < y)%R
r:radix
x, y:R
Py:(0 < y)%R
Px:(0 < x)%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R

(Build_mag_prop x ex Hex < mag y)%Z -> (x < y)%R
r:radix
x, y:R
Py:(0 < y)%R
Px:(0 < x)%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R

(Build_mag_prop x ex Hex < Build_mag_prop y ey Hey)%Z -> (x < y)%R
r:radix
x, y:R
Py:(0 < y)%R
Px:(0 < x)%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R

(ex < ey)%Z -> (x < y)%R
r:radix
x, y:R
Py:(0 < y)%R
Px:(0 < x)%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
H:(ex < ey)%Z

(x < y)%R
r:radix
x, y:R
Py:(0 < y)%R
Px:(0 < x)%R
ex, ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
H:(ex < ey)%Z
Hex:(Rabs x < bpow ex)%R

(x < y)%R
r:radix
x, y:R
Py:(0 < y)%R
Px:(0 < x)%R
ex, ey:Z
H:(ex < ey)%Z
Hex:(Rabs x < bpow ex)%R
Hey:(bpow (ey - 1) <= Rabs y)%R

(x < y)%R
r:radix
x, y:R
Py:(0 < y)%R
Px:(0 < x)%R
ex, ey:Z
H:(ex < ey)%Z
Hex:(x < bpow ex)%R
Hey:(bpow (ey - 1) <= Rabs y)%R

(x < y)%R
r:radix
x, y:R
Py:(0 < y)%R
Px:(0 < x)%R
ex, ey:Z
H:(ex < ey)%Z
Hex:(x < bpow ex)%R
Hey:(bpow (ey - 1) <= y)%R

(x < y)%R
r:radix
x, y:R
Py:(0 < y)%R
Px:(0 < x)%R
ex, ey:Z
H:(ex < ey)%Z
Hex:(x < bpow ex)%R
Hey:(bpow (ey - 1) <= y)%R

(bpow ex <= y)%R
r:radix
x, y:R
Py:(0 < y)%R
Px:(0 < x)%R
ex, ey:Z
H:(ex < ey)%Z
Hex:(x < bpow ex)%R
Hey:(bpow (ey - 1) <= y)%R

(bpow ex <= bpow (ey - 1))%R
now apply bpow_le; omega. Qed.
r:radix

forall e : Z, mag (bpow e) = (e + 1)%Z
r:radix

forall e : Z, mag (bpow e) = (e + 1)%Z
r:radix
e:Z

mag (bpow e) = (e + 1)%Z
r:radix
e:Z

(bpow (e + 1 - 1) <= Rabs (bpow e) < bpow (e + 1))%R
r:radix
e:Z

(bpow (e + 1 - 1) <= bpow e < bpow (e + 1))%R
r:radix
e:Z
(bpow e >= 0)%R
r:radix
e:Z

(bpow e <= bpow e < bpow (e + 1))%R
r:radix
e:Z
(bpow e >= 0)%R
r:radix
e:Z

(bpow e <= bpow e)%R
r:radix
e:Z
(bpow e < bpow (e + 1))%R
r:radix
e:Z
(bpow e >= 0)%R
r:radix
e:Z

(bpow e < bpow (e + 1))%R
r:radix
e:Z
(bpow e >= 0)%R
r:radix
e:Z

(e < e + 1)%Z
r:radix
e:Z
(bpow e >= 0)%R
r:radix
e:Z

(bpow e >= 0)%R
r:radix
e:Z

(0 <= bpow e)%R
apply bpow_ge_0. Qed.
r:radix

forall (x : R) (e : Z), x <> 0%R -> mag (x * bpow e) = (mag x + e)%Z
r:radix

forall (x : R) (e : Z), x <> 0%R -> mag (x * bpow e) = (mag x + e)%Z
r:radix
x:R
e:Z
Zx:x <> 0%R

mag (x * bpow e) = (mag x + e)%Z
r:radix
x:R
e:Z
Zx:x <> 0%R
ex:Z
Ex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R

mag (x * bpow e) = (ex + e)%Z
r:radix
x:R
e:Z
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R

mag (x * bpow e) = (ex + e)%Z
r:radix
x:R
e:Z
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R

(bpow (ex + e - 1) <= Rabs (x * bpow e) < bpow (ex + e))%R
r:radix
x:R
e:Z
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R

(bpow (ex + e - 1) <= Rabs x * Rabs (bpow e) < bpow (ex + e))%R
r:radix
x:R
e:Z
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R

(bpow (ex + e - 1) <= Rabs x * bpow e < bpow (ex + e))%R
r:radix
x:R
e:Z
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R

(bpow (ex + e - 1) <= Rabs x * bpow e)%R
r:radix
x:R
e:Z
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(Rabs x * bpow e < bpow (ex + e))%R
r:radix
x:R
e:Z
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R

(bpow (ex - 1 + e) <= Rabs x * bpow e)%R
r:radix
x:R
e:Z
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(Rabs x * bpow e < bpow (ex + e))%R
r:radix
x:R
e:Z
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R

(bpow (ex - 1) * bpow e <= Rabs x * bpow e)%R
r:radix
x:R
e:Z
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(Rabs x * bpow e < bpow (ex + e))%R
r:radix
x:R
e:Z
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R

(0 <= bpow e)%R
r:radix
x:R
e:Z
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(bpow (ex - 1) <= Rabs x)%R
r:radix
x:R
e:Z
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(Rabs x * bpow e < bpow (ex + e))%R
r:radix
x:R
e:Z
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R

(bpow (ex - 1) <= Rabs x)%R
r:radix
x:R
e:Z
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(Rabs x * bpow e < bpow (ex + e))%R
r:radix
x:R
e:Z
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R

(Rabs x * bpow e < bpow (ex + e))%R
r:radix
x:R
e:Z
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R

(Rabs x * bpow e < bpow ex * bpow e)%R
r:radix
x:R
e:Z
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R

(0 < bpow e)%R
r:radix
x:R
e:Z
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R
(Rabs x < bpow ex)%R
r:radix
x:R
e:Z
Zx:x <> 0%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R

(Rabs x < bpow ex)%R
apply Ex. Qed.
r:radix

forall (x : R) (e : Z), x <> 0%R -> (Rabs x < bpow e)%R -> (mag x <= e)%Z
r:radix

forall (x : R) (e : Z), x <> 0%R -> (Rabs x < bpow e)%R -> (mag x <= e)%Z
r:radix
x:R
e:Z
Zx:x <> 0%R
Hx:(Rabs x < bpow e)%R

(mag x <= e)%Z
r:radix
x:R
e:Z
Zx:x <> 0%R
Hx:(Rabs x < bpow e)%R
ex:Z
Ex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R

(ex <= e)%Z
r:radix
x:R
e:Z
Zx:x <> 0%R
Hx:(Rabs x < bpow e)%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R

(ex <= e)%Z
r:radix
x:R
e:Z
Zx:x <> 0%R
Hx:(Rabs x < bpow e)%R
ex:Z
Ex:(bpow (ex - 1) <= Rabs x < bpow ex)%R

(bpow (ex - 1) < bpow e)%R
now apply Rle_lt_trans with (Rabs x). Qed.
r:radix

forall (x : R) (e : Z), (bpow e <= Rabs x)%R -> (e < mag x)%Z
r:radix

forall (x : R) (e : Z), (bpow e <= Rabs x)%R -> (e < mag x)%Z
r:radix
x:R
e:Z
Hx:(bpow e <= Rabs x)%R

(e < mag x)%Z
r:radix
x:R
e:Z
Hx:(bpow e <= Rabs x)%R
ex:Z
Ex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R

(e < ex)%Z
r:radix
x:R
e:Z
Hx:(bpow e <= Rabs x)%R
ex:Z
Ex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R

(bpow e < bpow ex)%R
r:radix
x:R
e:Z
Hx:(bpow e <= Rabs x)%R
ex:Z
Ex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R

(Rabs x < bpow ex)%R
r:radix
x:R
e:Z
Hx:(bpow e <= Rabs x)%R
ex:Z
Ex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R

x <> 0%R
r:radix
x:R
e:Z
Hx:(bpow e <= Rabs x)%R
ex:Z
Ex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
Zx:x = 0%R

False
r:radix
x:R
e:Z
Hx:(bpow e <= Rabs x)%R
ex:Z
Ex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
Zx:x = 0%R

(Rabs x < bpow e)%R
r:radix
x:R
e:Z
Hx:(bpow e <= Rabs x)%R
ex:Z
Ex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
Zx:x = 0%R

(0 < bpow e)%R
apply bpow_gt_0. Qed.
r:radix

forall (x : R) (e : Z), (bpow (e - 1) <= Rabs x)%R -> (e <= mag x)%Z
r:radix

forall (x : R) (e : Z), (bpow (e - 1) <= Rabs x)%R -> (e <= mag x)%Z
r:radix
x:R
e:Z
H:(bpow (e - 1) <= Rabs x)%R

(e <= mag x)%Z
r:radix
x:R
e:Z
H:(bpow (e - 1) <= Rabs x)%R
Hxe:(Rabs x < bpow e)%R

(e <= mag x)%Z
r:radix
x:R
e:Z
H:(bpow (e - 1) <= Rabs x)%R
Hxe:(bpow e <= Rabs x)%R
(e <= mag x)%Z
r:radix
x:R
e:Z
H:(bpow (e - 1) <= Rabs x)%R
Hxe:(Rabs x < bpow e)%R

(e <= mag x)%Z
r:radix
x:R
e:Z
H:(bpow (e - 1) <= Rabs x)%R
Hxe:(Rabs x < bpow e)%R

mag x = e
r:radix
x:R
e:Z
H:(bpow (e - 1) <= Rabs x)%R
Hxe:(Rabs x < bpow e)%R
Hln:mag x = e
(e <= mag x)%Z
r:radix
x:R
e:Z
H:(bpow (e - 1) <= Rabs x)%R
Hxe:(Rabs x < bpow e)%R
Hln:mag x = e

(e <= mag x)%Z
r:radix
x:R
e:Z
H:(bpow (e - 1) <= Rabs x)%R
Hxe:(Rabs x < bpow e)%R
Hln:mag x = e

(e <= e)%Z
now apply Z.le_refl.
r:radix
x:R
e:Z
H:(bpow (e - 1) <= Rabs x)%R
Hxe:(bpow e <= Rabs x)%R

(e <= mag x)%Z
r:radix
x:R
e:Z
H:(bpow (e - 1) <= Rabs x)%R
Hxe:(bpow e <= Rabs x)%R

(e < mag x)%Z
now apply mag_gt_bpow. Qed.
r:radix

forall x : R, (Rabs x < bpow (mag x))%R
r:radix

forall x : R, (Rabs x < bpow (mag x))%R
r:radix
x:R

(Rabs x < bpow (mag x))%R
r:radix
x:R
Zx:x = 0%R

(Rabs x < bpow (mag x))%R
r:radix
x:R
Zx:x <> 0%R
(Rabs x < bpow (mag x))%R
r:radix
x:R
Zx:x = 0%R

(0 < bpow (mag 0))%R
r:radix
x:R
Zx:x <> 0%R
(Rabs x < bpow (mag x))%R
r:radix
x:R
Zx:x <> 0%R

(Rabs x < bpow (mag x))%R
r:radix
x:R
Zx:x <> 0%R
ex:Z
Ex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R

(Rabs x < bpow ex)%R
now apply Ex. Qed.
r:radix

forall x : R, x <> 0%R -> (bpow (mag x - 1) <= Rabs x)%R
r:radix

forall x : R, x <> 0%R -> (bpow (mag x - 1) <= Rabs x)%R
r:radix
x:R
Hx:x <> 0%R

(bpow (mag x - 1) <= Rabs x)%R
r:radix
x:R
Hx:x <> 0%R
ex:Z
Ex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R

(bpow (ex - 1) <= Rabs x)%R
now apply Ex. Qed.
r:radix

forall m e : Z, m <> 0%Z -> (Z.abs m < r ^ e)%Z -> (mag (IZR m) <= e)%Z
r:radix

forall m e : Z, m <> 0%Z -> (Z.abs m < r ^ e)%Z -> (mag (IZR m) <= e)%Z
r:radix
m, e:Z
Zm:m <> 0%Z
Hm:(Z.abs m < r ^ e)%Z

(mag (IZR m) <= e)%Z
r:radix
m, e:Z
Zm:m <> 0%Z
Hm:(Z.abs m < r ^ e)%Z

IZR m <> 0%R
r:radix
m, e:Z
Zm:m <> 0%Z
Hm:(Z.abs m < r ^ e)%Z
(Rabs (IZR m) < bpow e)%R
r:radix
m, e:Z
Zm:m <> 0%Z
Hm:(Z.abs m < r ^ e)%Z

(Rabs (IZR m) < bpow e)%R
r:radix
m, e:Z
Zm:m <> 0%Z
Hm:(Z.abs m < r ^ e)%Z
H:(0 <= e)%Z

(Rabs (IZR m) < bpow e)%R
r:radix
m, e:Z
Zm:m <> 0%Z
Hm:(Z.abs m < r ^ e)%Z
H:(e < 0)%Z
(Rabs (IZR m) < bpow e)%R
r:radix
m, e:Z
Zm:m <> 0%Z
Hm:(Z.abs m < r ^ e)%Z
H:(0 <= e)%Z

(IZR (Z.abs m) < IZR (r ^ e))%R
r:radix
m, e:Z
Zm:m <> 0%Z
Hm:(Z.abs m < r ^ e)%Z
H:(e < 0)%Z
(Rabs (IZR m) < bpow e)%R
r:radix
m, e:Z
Zm:m <> 0%Z
Hm:(Z.abs m < r ^ e)%Z
H:(e < 0)%Z

(Rabs (IZR m) < bpow e)%R
r:radix
m, e:Z
Zm:m <> 0%Z
Hm:(Z.abs m < r ^ e)%Z
H:(e < 0)%Z

m = 0%Z
r:radix
m, e:Z
Zm:m <> 0%Z
Hm:(Z.abs m < r ^ e)%Z
H:(e < 0)%Z

(Z.abs m < 0)%Z -> m = 0%Z
r:radix
m, e:Z
Zm:m <> 0%Z
Hm:(Z.abs m < r ^ e)%Z
H:(e < 0)%Z
(Z.abs m < 0)%Z
r:radix
m, e:Z
Zm:m <> 0%Z
Hm:(Z.abs m < r ^ e)%Z
H:(e < 0)%Z

(Z.abs m < 0)%Z
r:radix
m, e:Z
Hm:(Z.abs m < r ^ e)%Z
H:(e < 0)%Z

(Z.abs m < 0)%Z
now destruct e. Qed.
r:radix

forall m e : Z, m <> 0%Z -> (r ^ e <= Z.abs m)%Z -> (e < mag (IZR m))%Z
r:radix

forall m e : Z, m <> 0%Z -> (r ^ e <= Z.abs m)%Z -> (e < mag (IZR m))%Z
r:radix
m, e:Z
Zm:m <> 0%Z
Hm:(r ^ e <= Z.abs m)%Z

(e < mag (IZR m))%Z
r:radix
m, e:Z
Zm:m <> 0%Z
Hm:(r ^ e <= Z.abs m)%Z

(bpow e <= Rabs (IZR m))%R
r:radix
m, e:Z
Zm:m <> 0%Z
Hm:(r ^ e <= Z.abs m)%Z

(bpow e <= IZR (Z.abs m))%R
r:radix
m, e:Z
Zm:m <> 0%Z
Hm:(r ^ e <= Z.abs m)%Z
H:(0 <= e)%Z

(bpow e <= IZR (Z.abs m))%R
r:radix
m, e:Z
Zm:m <> 0%Z
Hm:(r ^ e <= Z.abs m)%Z
H:(e < 0)%Z
(bpow e <= IZR (Z.abs m))%R
r:radix
m, e:Z
Zm:m <> 0%Z
Hm:(r ^ e <= Z.abs m)%Z
H:(0 <= e)%Z

(IZR (r ^ e) <= IZR (Z.abs m))%R
r:radix
m, e:Z
Zm:m <> 0%Z
Hm:(r ^ e <= Z.abs m)%Z
H:(e < 0)%Z
(bpow e <= IZR (Z.abs m))%R
r:radix
m, e:Z
Zm:m <> 0%Z
Hm:(r ^ e <= Z.abs m)%Z
H:(e < 0)%Z

(bpow e <= IZR (Z.abs m))%R
r:radix
m, e:Z
Zm:m <> 0%Z
Hm:(r ^ e <= Z.abs m)%Z
H:(e < 0)%Z

(bpow e <= bpow 0)%R
r:radix
m, e:Z
Zm:m <> 0%Z
Hm:(r ^ e <= Z.abs m)%Z
H:(e < 0)%Z
(bpow 0 <= IZR (Z.abs m))%R
r:radix
m, e:Z
Zm:m <> 0%Z
Hm:(r ^ e <= Z.abs m)%Z
H:(e < 0)%Z

(e <= 0)%Z
r:radix
m, e:Z
Zm:m <> 0%Z
Hm:(r ^ e <= Z.abs m)%Z
H:(e < 0)%Z
(bpow 0 <= IZR (Z.abs m))%R
r:radix
m, e:Z
Zm:m <> 0%Z
Hm:(r ^ e <= Z.abs m)%Z
H:(e < 0)%Z

(bpow 0 <= IZR (Z.abs m))%R
r:radix
m, e:Z
Zm:m <> 0%Z
Hm:(r ^ e <= Z.abs m)%Z
H:(e < 0)%Z

(1 <= Z.abs m)%Z
m:Z
Zm:m <> 0%Z

(1 <= Z.abs m)%Z
zify ; omega. Qed.
r:radix

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

forall x y : R, x <> 0%R -> y <> 0%R -> (mag x + mag y - 1 <= mag (x * y) <= mag x + mag y)%Z
r:radix
x, y:R
Hx:x <> 0%R
Hy:y <> 0%R

(mag x + mag y - 1 <= mag (x * y) <= mag x + mag y)%Z
r:radix
x, y:R
Hx:x <> 0%R
Hy:y <> 0%R
ex:Z
Hx2:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R

(Build_mag_prop x ex Hx2 + mag y - 1 <= mag (x * y) <= Build_mag_prop x ex Hx2 + mag y)%Z
r:radix
x, y:R
Hx:x <> 0%R
Hy:y <> 0%R
ex:Z
Hx2:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hy2:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R

(Build_mag_prop x ex Hx2 + Build_mag_prop y ey Hy2 - 1 <= mag (x * y) <= Build_mag_prop x ex Hx2 + Build_mag_prop y ey Hy2)%Z
r:radix
x, y:R
Hx:x <> 0%R
Hy:y <> 0%R
ex:Z
Hx2:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hy2:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R

(ex + ey - 1 <= mag (x * y) <= ex + ey)%Z
r:radix
x, y:R
Hx:x <> 0%R
Hy:y <> 0%R
ex, ey:Z
Hy2:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hx21:(bpow (ex - 1) <= Rabs x)%R
Hx22:(Rabs x < bpow ex)%R

(ex + ey - 1 <= mag (x * y) <= ex + ey)%Z
r:radix
x, y:R
Hx:x <> 0%R
Hy:y <> 0%R
ex, ey:Z
Hx21:(bpow (ex - 1) <= Rabs x)%R
Hx22:(Rabs x < bpow ex)%R
Hy21:(bpow (ey - 1) <= Rabs y)%R
Hy22:(Rabs y < bpow ey)%R

(ex + ey - 1 <= mag (x * y) <= ex + ey)%Z
r:radix
x, y:R
Hx:x <> 0%R
Hy:y <> 0%R
ex, ey:Z
Hx21:(bpow (ex - 1) <= Rabs x)%R
Hx22:(Rabs x < bpow ex)%R
Hy21:(bpow (ey - 1) <= Rabs y)%R
Hy22:(Rabs y < bpow ey)%R

(bpow (ex + ey - 1 - 1) <= Rabs (x * y))%R
r:radix
x, y:R
Hx:x <> 0%R
Hy:y <> 0%R
ex, ey:Z
Hx21:(bpow (ex - 1) <= Rabs x)%R
Hx22:(Rabs x < bpow ex)%R
Hy21:(bpow (ey - 1) <= Rabs y)%R
Hy22:(Rabs y < bpow ey)%R
Hxy:(bpow (ex + ey - 1 - 1) <= Rabs (x * y))%R
(ex + ey - 1 <= mag (x * y) <= ex + ey)%Z
r:radix
x, y:R
Hx:x <> 0%R
Hy:y <> 0%R
ex, ey:Z
Hx21:(bpow (ex - 1) <= Rabs x)%R
Hx22:(Rabs x < bpow ex)%R
Hy21:(bpow (ey - 1) <= Rabs y)%R
Hy22:(Rabs y < bpow ey)%R

(bpow (ex + ey - 1 - 1) <= Rabs (x * y))%R
r:radix
x, y:R
Hx:x <> 0%R
Hy:y <> 0%R
ex, ey:Z
Hx21:(bpow (ex - 1) <= Rabs x)%R
Hx22:(Rabs x < bpow ex)%R
Hy21:(bpow (ey - 1) <= Rabs y)%R
Hy22:(Rabs y < bpow ey)%R

(bpow (ex - 1 + (ey - 1)) <= Rabs (x * y))%R
r:radix
x, y:R
Hx:x <> 0%R
Hy:y <> 0%R
ex, ey:Z
Hx21:(bpow (ex - 1) <= Rabs x)%R
Hx22:(Rabs x < bpow ex)%R
Hy21:(bpow (ey - 1) <= Rabs y)%R
Hy22:(Rabs y < bpow ey)%R

(bpow (ex - 1) * bpow (ey - 1) <= Rabs (x * y))%R
r:radix
x, y:R
Hx:x <> 0%R
Hy:y <> 0%R
ex, ey:Z
Hx21:(bpow (ex - 1) <= Rabs x)%R
Hx22:(Rabs x < bpow ex)%R
Hy21:(bpow (ey - 1) <= Rabs y)%R
Hy22:(Rabs y < bpow ey)%R

(bpow (ex - 1) * bpow (ey - 1) <= Rabs x * Rabs y)%R
now apply Rmult_le_compat; try apply bpow_ge_0.
r:radix
x, y:R
Hx:x <> 0%R
Hy:y <> 0%R
ex, ey:Z
Hx21:(bpow (ex - 1) <= Rabs x)%R
Hx22:(Rabs x < bpow ex)%R
Hy21:(bpow (ey - 1) <= Rabs y)%R
Hy22:(Rabs y < bpow ey)%R
Hxy:(bpow (ex + ey - 1 - 1) <= Rabs (x * y))%R

(ex + ey - 1 <= mag (x * y) <= ex + ey)%Z
r:radix
x, y:R
Hx:x <> 0%R
Hy:y <> 0%R
ex, ey:Z
Hx21:(bpow (ex - 1) <= Rabs x)%R
Hx22:(Rabs x < bpow ex)%R
Hy21:(bpow (ey - 1) <= Rabs y)%R
Hy22:(Rabs y < bpow ey)%R
Hxy:(bpow (ex + ey - 1 - 1) <= Rabs (x * y))%R

(Rabs (x * y) < bpow (ex + ey))%R
r:radix
x, y:R
Hx:x <> 0%R
Hy:y <> 0%R
ex, ey:Z
Hx21:(bpow (ex - 1) <= Rabs x)%R
Hx22:(Rabs x < bpow ex)%R
Hy21:(bpow (ey - 1) <= Rabs y)%R
Hy22:(Rabs y < bpow ey)%R
Hxy:(bpow (ex + ey - 1 - 1) <= Rabs (x * y))%R
Hxy2:(Rabs (x * y) < bpow (ex + ey))%R
(ex + ey - 1 <= mag (x * y) <= ex + ey)%Z
r:radix
x, y:R
Hx:x <> 0%R
Hy:y <> 0%R
ex, ey:Z
Hx21:(bpow (ex - 1) <= Rabs x)%R
Hx22:(Rabs x < bpow ex)%R
Hy21:(bpow (ey - 1) <= Rabs y)%R
Hy22:(Rabs y < bpow ey)%R
Hxy:(bpow (ex + ey - 1 - 1) <= Rabs (x * y))%R

(Rabs (x * y) < bpow (ex + ey))%R
r:radix
x, y:R
Hx:x <> 0%R
Hy:y <> 0%R
ex, ey:Z
Hx21:(bpow (ex - 1) <= Rabs x)%R
Hx22:(Rabs x < bpow ex)%R
Hy21:(bpow (ey - 1) <= Rabs y)%R
Hy22:(Rabs y < bpow ey)%R
Hxy:(bpow (ex + ey - 1 - 1) <= Rabs (x * y))%R

(Rabs x * Rabs y < bpow (ex + ey))%R
r:radix
x, y:R
Hx:x <> 0%R
Hy:y <> 0%R
ex, ey:Z
Hx21:(bpow (ex - 1) <= Rabs x)%R
Hx22:(Rabs x < bpow ex)%R
Hy21:(bpow (ey - 1) <= Rabs y)%R
Hy22:(Rabs y < bpow ey)%R
Hxy:(bpow (ex + ey - 1 - 1) <= Rabs (x * y))%R

(Rabs x * Rabs y < bpow ex * bpow ey)%R
r:radix
x, y:R
Hx:x <> 0%R
Hy:y <> 0%R
ex, ey:Z
Hx21:(bpow (ex - 1) <= Rabs x)%R
Hx22:(Rabs x < bpow ex)%R
Hy21:(bpow (ey - 1) <= Rabs y)%R
Hy22:(Rabs y < bpow ey)%R
Hxy:(bpow (ex + ey - 1 - 1) <= Rabs (x * y))%R

(0 <= Rabs x)%R
r:radix
x, y:R
Hx:x <> 0%R
Hy:y <> 0%R
ex, ey:Z
Hx21:(bpow (ex - 1) <= Rabs x)%R
Hx22:(Rabs x < bpow ex)%R
Hy21:(bpow (ey - 1) <= Rabs y)%R
Hy22:(Rabs y < bpow ey)%R
Hxy:(bpow (ex + ey - 1 - 1) <= Rabs (x * y))%R
(0 <= Rabs y)%R
r:radix
x, y:R
Hx:x <> 0%R
Hy:y <> 0%R
ex, ey:Z
Hx21:(bpow (ex - 1) <= Rabs x)%R
Hx22:(Rabs x < bpow ex)%R
Hy21:(bpow (ey - 1) <= Rabs y)%R
Hy22:(Rabs y < bpow ey)%R
Hxy:(bpow (ex + ey - 1 - 1) <= Rabs (x * y))%R

(0 <= Rabs y)%R
now apply Rle_trans with (bpow (ey - 1)); try apply bpow_ge_0.
r:radix
x, y:R
Hx:x <> 0%R
Hy:y <> 0%R
ex, ey:Z
Hx21:(bpow (ex - 1) <= Rabs x)%R
Hx22:(Rabs x < bpow ex)%R
Hy21:(bpow (ey - 1) <= Rabs y)%R
Hy22:(Rabs y < bpow ey)%R
Hxy:(bpow (ex + ey - 1 - 1) <= Rabs (x * y))%R
Hxy2:(Rabs (x * y) < bpow (ex + ey))%R

(ex + ey - 1 <= mag (x * y) <= ex + ey)%Z
r:radix
x, y:R
Hx:x <> 0%R
Hy:y <> 0%R
ex, ey:Z
Hx21:(bpow (ex - 1) <= Rabs x)%R
Hx22:(Rabs x < bpow ex)%R
Hy21:(bpow (ey - 1) <= Rabs y)%R
Hy22:(Rabs y < bpow ey)%R
Hxy:(bpow (ex + ey - 1 - 1) <= Rabs (x * y))%R
Hxy2:(Rabs (x * y) < bpow (ex + ey))%R

(ex + ey - 1 <= mag (x * y))%Z
r:radix
x, y:R
Hx:x <> 0%R
Hy:y <> 0%R
ex, ey:Z
Hx21:(bpow (ex - 1) <= Rabs x)%R
Hx22:(Rabs x < bpow ex)%R
Hy21:(bpow (ey - 1) <= Rabs y)%R
Hy22:(Rabs y < bpow ey)%R
Hxy:(bpow (ex + ey - 1 - 1) <= Rabs (x * y))%R
Hxy2:(Rabs (x * y) < bpow (ex + ey))%R
(mag (x * y) <= ex + ey)%Z
r:radix
x, y:R
Hx:x <> 0%R
Hy:y <> 0%R
ex, ey:Z
Hx21:(bpow (ex - 1) <= Rabs x)%R
Hx22:(Rabs x < bpow ex)%R
Hy21:(bpow (ey - 1) <= Rabs y)%R
Hy22:(Rabs y < bpow ey)%R
Hxy:(bpow (ex + ey - 1 - 1) <= Rabs (x * y))%R
Hxy2:(Rabs (x * y) < bpow (ex + ey))%R

(ex + ey - 1 <= mag (x * y))%Z
now apply mag_ge_bpow.
r:radix
x, y:R
Hx:x <> 0%R
Hy:y <> 0%R
ex, ey:Z
Hx21:(bpow (ex - 1) <= Rabs x)%R
Hx22:(Rabs x < bpow ex)%R
Hy21:(bpow (ey - 1) <= Rabs y)%R
Hy22:(Rabs y < bpow ey)%R
Hxy:(bpow (ex + ey - 1 - 1) <= Rabs (x * y))%R
Hxy2:(Rabs (x * y) < bpow (ex + ey))%R

(mag (x * y) <= ex + ey)%Z
r:radix
x, y:R
Hx:x <> 0%R
Hy:y <> 0%R
ex, ey:Z
Hx21:(bpow (ex - 1) <= Rabs x)%R
Hx22:(Rabs x < bpow ex)%R
Hy21:(bpow (ey - 1) <= Rabs y)%R
Hy22:(Rabs y < bpow ey)%R
Hxy:(bpow (ex + ey - 1 - 1) <= Rabs (x * y))%R
Hxy2:(Rabs (x * y) < bpow (ex + ey))%R

(x * y)%R <> 0%R
r:radix
x, y:R
Hx:x <> 0%R
Hy:y <> 0%R
ex, ey:Z
Hx21:(bpow (ex - 1) <= Rabs x)%R
Hx22:(Rabs x < bpow ex)%R
Hy21:(bpow (ey - 1) <= Rabs y)%R
Hy22:(Rabs y < bpow ey)%R
Hxy:(bpow (ex + ey - 1 - 1) <= Rabs (x * y))%R
Hxy2:(Rabs (x * y) < bpow (ex + ey))%R
(Rabs (x * y) < bpow (ex + ey))%R
r:radix
x, y:R
Hx:x <> 0%R
Hy:y <> 0%R
ex, ey:Z
Hx21:(bpow (ex - 1) <= Rabs x)%R
Hx22:(Rabs x < bpow ex)%R
Hy21:(bpow (ey - 1) <= Rabs y)%R
Hy22:(Rabs y < bpow ey)%R
Hxy:(bpow (ex + ey - 1 - 1) <= Rabs (x * y))%R
Hxy2:(Rabs (x * y) < bpow (ex + ey))%R

(x * y)%R <> 0%R
now apply Rmult_integral_contrapositive_currified.
r:radix
x, y:R
Hx:x <> 0%R
Hy:y <> 0%R
ex, ey:Z
Hx21:(bpow (ex - 1) <= Rabs x)%R
Hx22:(Rabs x < bpow ex)%R
Hy21:(bpow (ey - 1) <= Rabs y)%R
Hy22:(Rabs y < bpow ey)%R
Hxy:(bpow (ex + ey - 1 - 1) <= Rabs (x * y))%R
Hxy2:(Rabs (x * y) < bpow (ex + ey))%R

(Rabs (x * y) < bpow (ex + ey))%R
assumption. Qed.
r:radix

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

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

(2 <= r)%Z
r:radix
Hr:(2 <= r)%Z
forall x y : R, (0 < y)%R -> (y <= x)%R -> (mag x <= mag (x + y) <= mag x + 1)%Z
r:radix

(2 <= r)%Z
r: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.
r:radix
Hr:(2 <= r)%Z

forall x y : R, (0 < y)%R -> (y <= x)%R -> (mag x <= mag (x + y) <= mag x + 1)%Z
r:radix
Hr:(2 <= r)%Z
x, y:R
Hy:(0 < y)%R
Hxy:(y <= x)%R

(mag x <= mag (x + y) <= mag x + 1)%Z
r:radix
Hr:(2 <= r)%Z
x, y:R
Hy:(0 < y)%R
Hxy:(y <= x)%R
Hx:(0 < x)%R

(mag x <= mag (x + y) <= mag x + 1)%Z
r:radix
Hr:(2 <= r)%Z
x, y:R
Hy:(0 < y)%R
Hxy:(y <= x)%R
Hx:(0 < x)%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R

(ex <= mag (x + y) <= ex + 1)%Z
r:radix
Hr:(2 <= r)%Z
x, y:R
Hy:(0 < y)%R
Hxy:(y <= x)%R
Hx:(0 < x)%R
ex:Z
Hex0:(bpow (ex - 1) <= Rabs x)%R
Hex1:(Rabs x < bpow ex)%R

(ex <= mag (x + y) <= ex + 1)%Z
r:radix
Hr:(2 <= r)%Z
x, y:R
Hy:(0 < y)%R
Hxy:(y <= x)%R
Hx:(0 < x)%R
ex:Z
Hex0:(bpow (ex - 1) <= Rabs x)%R
Hex1:(Rabs x < bpow ex)%R

(Rabs (x + y) < bpow (ex + 1))%R
r:radix
Hr:(2 <= r)%Z
x, y:R
Hy:(0 < y)%R
Hxy:(y <= x)%R
Hx:(0 < x)%R
ex:Z
Hex0:(bpow (ex - 1) <= Rabs x)%R
Hex1:(Rabs x < bpow ex)%R
Haxy:(Rabs (x + y) < bpow (ex + 1))%R
(ex <= mag (x + y) <= ex + 1)%Z
r:radix
Hr:(2 <= r)%Z
x, y:R
Hy:(0 < y)%R
Hxy:(y <= x)%R
Hx:(0 < x)%R
ex:Z
Hex0:(bpow (ex - 1) <= Rabs x)%R
Hex1:(Rabs x < bpow ex)%R

(Rabs (x + y) < bpow (ex + 1))%R
r:radix
Hr:(2 <= r)%Z
x, y:R
Hy:(0 < y)%R
Hxy:(y <= x)%R
Hx:(0 < x)%R
ex:Z
Hex0:(bpow (ex - 1) <= Rabs x)%R
Hex1:(Rabs x < bpow ex)%R

(Rabs (x + y) < IZR r * bpow ex)%R
r:radix
Hr:(2 <= r)%Z
x, y:R
Hy:(0 < y)%R
Hxy:(y <= x)%R
Hx:(0 < x)%R
ex:Z
Hex0:(bpow (ex - 1) <= Rabs x)%R
Hex1:(Rabs x < bpow ex)%R

(Rabs (x + y) < 2 * bpow ex)%R
r:radix
Hr:(2 <= r)%Z
x, y:R
Hy:(0 < y)%R
Hxy:(y <= x)%R
Hx:(0 < x)%R
ex:Z
Hex0:(bpow (ex - 1) <= Rabs x)%R
Hex1:(Rabs x < bpow ex)%R
(2 * bpow ex <= IZR r * bpow ex)%R
r:radix
Hr:(2 <= r)%Z
x, y:R
Hy:(0 < y)%R
Hxy:(y <= x)%R
Hx:(0 < x)%R
ex:Z
Hex0:(bpow (ex - 1) <= Rabs x)%R
Hex1:(Rabs x < bpow ex)%R

(Rabs (x + y) < 2 * bpow ex)%R
r:radix
Hr:(2 <= r)%Z
x, y:R
Hy:(0 < y)%R
Hxy:(y <= x)%R
Hx:(0 < x)%R
ex:Z
Hex0:(bpow (ex - 1) <= Rabs x)%R
Hex1:(Rabs x < bpow ex)%R

(x + y < 2 * bpow ex)%R
r:radix
Hr:(2 <= r)%Z
x, y:R
Hy:(0 < y)%R
Hxy:(y <= x)%R
Hx:(0 < x)%R
ex:Z
Hex0:(bpow (ex - 1) <= Rabs x)%R
Hex1:(Rabs x < bpow ex)%R
(0 <= x + y)%R
r:radix
Hr:(2 <= r)%Z
x, y:R
Hy:(0 < y)%R
Hxy:(y <= x)%R
Hx:(0 < x)%R
ex:Z
Hex0:(bpow (ex - 1) <= Rabs x)%R
Hex1:(Rabs x < bpow ex)%R

(x + y <= 2 * Rabs x)%R
r:radix
Hr:(2 <= r)%Z
x, y:R
Hy:(0 < y)%R
Hxy:(y <= x)%R
Hx:(0 < x)%R
ex:Z
Hex0:(bpow (ex - 1) <= Rabs x)%R
Hex1:(Rabs x < bpow ex)%R
(2 * Rabs x < 2 * bpow ex)%R
r:radix
Hr:(2 <= r)%Z
x, y:R
Hy:(0 < y)%R
Hxy:(y <= x)%R
Hx:(0 < x)%R
ex:Z
Hex0:(bpow (ex - 1) <= Rabs x)%R
Hex1:(Rabs x < bpow ex)%R
(0 <= x + y)%R
r:radix
Hr:(2 <= r)%Z
x, y:R
Hy:(0 < y)%R
Hxy:(y <= x)%R
Hx:(0 < x)%R
ex:Z
Hex0:(bpow (ex - 1) <= Rabs x)%R
Hex1:(Rabs x < bpow ex)%R

(x + y <= 2 * Rabs x)%R
r:radix
Hr:(2 <= r)%Z
x, y:R
Hy:(0 < y)%R
Hxy:(y <= x)%R
Hx:(0 < x)%R
ex:Z
Hex0:(bpow (ex - 1) <= Rabs x)%R
Hex1:(Rabs x < bpow ex)%R

(x + y <= 2 * x)%R
r:radix
Hr:(2 <= r)%Z
x, y:R
Hy:(0 < y)%R
Hxy:(y <= x)%R
Hx:(0 < x)%R
ex:Z
Hex0:(bpow (ex - 1) <= Rabs x)%R
Hex1:(Rabs x < bpow ex)%R
(0 <= x)%R
r:radix
Hr:(2 <= r)%Z
x, y:R
Hy:(0 < y)%R
Hxy:(y <= x)%R
Hx:(0 < x)%R
ex:Z
Hex0:(bpow (ex - 1) <= Rabs x)%R
Hex1:(Rabs x < bpow ex)%R

(x + y <= x + x)%R
r:radix
Hr:(2 <= r)%Z
x, y:R
Hy:(0 < y)%R
Hxy:(y <= x)%R
Hx:(0 < x)%R
ex:Z
Hex0:(bpow (ex - 1) <= Rabs x)%R
Hex1:(Rabs x < bpow ex)%R
(0 <= x)%R
r:radix
Hr:(2 <= r)%Z
x, y:R
Hy:(0 < y)%R
Hxy:(y <= x)%R
Hx:(0 < x)%R
ex:Z
Hex0:(bpow (ex - 1) <= Rabs x)%R
Hex1:(Rabs x < bpow ex)%R

(0 <= x)%R
now apply Rlt_le.
r:radix
Hr:(2 <= r)%Z
x, y:R
Hy:(0 < y)%R
Hxy:(y <= x)%R
Hx:(0 < x)%R
ex:Z
Hex0:(bpow (ex - 1) <= Rabs x)%R
Hex1:(Rabs x < bpow ex)%R

(2 * Rabs x < 2 * bpow ex)%R
r:radix
Hr:(2 <= r)%Z
x, y:R
Hy:(0 < y)%R
Hxy:(y <= x)%R
Hx:(0 < x)%R
ex:Z
Hex0:(bpow (ex - 1) <= Rabs x)%R
Hex1:(Rabs x < bpow ex)%R

(0 < 2)%R
exact Rlt_0_2.
r:radix
Hr:(2 <= r)%Z
x, y:R
Hy:(0 < y)%R
Hxy:(y <= x)%R
Hx:(0 < x)%R
ex:Z
Hex0:(bpow (ex - 1) <= Rabs x)%R
Hex1:(Rabs x < bpow ex)%R

(0 <= x + y)%R
r:radix
Hr:(2 <= r)%Z
x, y:R
Hy:(0 < y)%R
Hxy:(y <= x)%R
Hx:(0 < x)%R
ex:Z
Hex0:(bpow (ex - 1) <= Rabs x)%R
Hex1:(Rabs x < bpow ex)%R

(0 + 0 <= x + y)%R
now apply Rlt_le, Rplus_lt_compat.
r:radix
Hr:(2 <= r)%Z
x, y:R
Hy:(0 < y)%R
Hxy:(y <= x)%R
Hx:(0 < x)%R
ex:Z
Hex0:(bpow (ex - 1) <= Rabs x)%R
Hex1:(Rabs x < bpow ex)%R

(2 * bpow ex <= IZR r * bpow ex)%R
r:radix
Hr:(2 <= r)%Z
x, y:R
Hy:(0 < y)%R
Hxy:(y <= x)%R
Hx:(0 < x)%R
ex:Z
Hex0:(bpow (ex - 1) <= Rabs x)%R
Hex1:(Rabs x < bpow ex)%R

(0 <= bpow ex)%R
r:radix
Hr:(2 <= r)%Z
x, y:R
Hy:(0 < y)%R
Hxy:(y <= x)%R
Hx:(0 < x)%R
ex:Z
Hex0:(bpow (ex - 1) <= Rabs x)%R
Hex1:(Rabs x < bpow ex)%R
(2 <= IZR r)%R
r:radix
Hr:(2 <= r)%Z
x, y:R
Hy:(0 < y)%R
Hxy:(y <= x)%R
Hx:(0 < x)%R
ex:Z
Hex0:(bpow (ex - 1) <= Rabs x)%R
Hex1:(Rabs x < bpow ex)%R

(2 <= IZR r)%R
now apply IZR_le.
r:radix
Hr:(2 <= r)%Z
x, y:R
Hy:(0 < y)%R
Hxy:(y <= x)%R
Hx:(0 < x)%R
ex:Z
Hex0:(bpow (ex - 1) <= Rabs x)%R
Hex1:(Rabs x < bpow ex)%R
Haxy:(Rabs (x + y) < bpow (ex + 1))%R

(ex <= mag (x + y) <= ex + 1)%Z
r:radix
Hr:(2 <= r)%Z
x, y:R
Hy:(0 < y)%R
Hxy:(y <= x)%R
Hx:(0 < x)%R
ex:Z
Hex0:(bpow (ex - 1) <= Rabs x)%R
Hex1:(Rabs x < bpow ex)%R
Haxy:(Rabs (x + y) < bpow (ex + 1))%R

(bpow (ex - 1) <= Rabs (x + y))%R
r:radix
Hr:(2 <= r)%Z
x, y:R
Hy:(0 < y)%R
Hxy:(y <= x)%R
Hx:(0 < x)%R
ex:Z
Hex0:(bpow (ex - 1) <= Rabs x)%R
Hex1:(Rabs x < bpow ex)%R
Haxy:(Rabs (x + y) < bpow (ex + 1))%R
Haxy2:(bpow (ex - 1) <= Rabs (x + y))%R
(ex <= mag (x + y) <= ex + 1)%Z
r:radix
Hr:(2 <= r)%Z
x, y:R
Hy:(0 < y)%R
Hxy:(y <= x)%R
Hx:(0 < x)%R
ex:Z
Hex0:(bpow (ex - 1) <= Rabs x)%R
Hex1:(Rabs x < bpow ex)%R
Haxy:(Rabs (x + y) < bpow (ex + 1))%R

(bpow (ex - 1) <= Rabs (x + y))%R
r:radix
Hr:(2 <= r)%Z
x, y:R
Hy:(0 < y)%R
Hxy:(y <= x)%R
Hx:(0 < x)%R
ex:Z
Hex0:(bpow (ex - 1) <= Rabs x)%R
Hex1:(Rabs x < bpow ex)%R
Haxy:(Rabs (x + y) < bpow (ex + 1))%R

(Rabs x <= Rabs (x + y))%R
r:radix
Hr:(2 <= r)%Z
x, y:R
Hy:(0 < y)%R
Hxy:(y <= x)%R
Hx:(0 < x)%R
ex:Z
Hex0:(bpow (ex - 1) <= Rabs x)%R
Hex1:(Rabs x < bpow ex)%R
Haxy:(Rabs (x + y) < bpow (ex + 1))%R

(x <= Rabs (x + y))%R
r:radix
Hr:(2 <= r)%Z
x, y:R
Hy:(0 < y)%R
Hxy:(y <= x)%R
Hx:(0 < x)%R
ex:Z
Hex0:(bpow (ex - 1) <= Rabs x)%R
Hex1:(Rabs x < bpow ex)%R
Haxy:(Rabs (x + y) < bpow (ex + 1))%R

(x <= x + y)%R
r:radix
Hr:(2 <= r)%Z
x, y:R
Hy:(0 < y)%R
Hxy:(y <= x)%R
Hx:(0 < x)%R
ex:Z
Hex0:(bpow (ex - 1) <= Rabs x)%R
Hex1:(Rabs x < bpow ex)%R
Haxy:(Rabs (x + y) < bpow (ex + 1))%R

(x + 0 <= x + y)%R
r:radix
Hr:(2 <= r)%Z
x, y:R
Hy:(0 < y)%R
Hxy:(y <= x)%R
Hx:(0 < x)%R
ex:Z
Hex0:(bpow (ex - 1) <= Rabs x)%R
Hex1:(Rabs x < bpow ex)%R
Haxy:(Rabs (x + y) < bpow (ex + 1))%R

(0 <= y)%R
now apply Rlt_le.
r:radix
Hr:(2 <= r)%Z
x, y:R
Hy:(0 < y)%R
Hxy:(y <= x)%R
Hx:(0 < x)%R
ex:Z
Hex0:(bpow (ex - 1) <= Rabs x)%R
Hex1:(Rabs x < bpow ex)%R
Haxy:(Rabs (x + y) < bpow (ex + 1))%R
Haxy2:(bpow (ex - 1) <= Rabs (x + y))%R

(ex <= mag (x + y) <= ex + 1)%Z
r:radix
Hr:(2 <= r)%Z
x, y:R
Hy:(0 < y)%R
Hxy:(y <= x)%R
Hx:(0 < x)%R
ex:Z
Hex0:(bpow (ex - 1) <= Rabs x)%R
Hex1:(Rabs x < bpow ex)%R
Haxy:(Rabs (x + y) < bpow (ex + 1))%R
Haxy2:(bpow (ex - 1) <= Rabs (x + y))%R

(ex <= mag (x + y))%Z
r:radix
Hr:(2 <= r)%Z
x, y:R
Hy:(0 < y)%R
Hxy:(y <= x)%R
Hx:(0 < x)%R
ex:Z
Hex0:(bpow (ex - 1) <= Rabs x)%R
Hex1:(Rabs x < bpow ex)%R
Haxy:(Rabs (x + y) < bpow (ex + 1))%R
Haxy2:(bpow (ex - 1) <= Rabs (x + y))%R
(mag (x + y) <= ex + 1)%Z
r:radix
Hr:(2 <= r)%Z
x, y:R
Hy:(0 < y)%R
Hxy:(y <= x)%R
Hx:(0 < x)%R
ex:Z
Hex0:(bpow (ex - 1) <= Rabs x)%R
Hex1:(Rabs x < bpow ex)%R
Haxy:(Rabs (x + y) < bpow (ex + 1))%R
Haxy2:(bpow (ex - 1) <= Rabs (x + y))%R

(ex <= mag (x + y))%Z
now apply mag_ge_bpow.
r:radix
Hr:(2 <= r)%Z
x, y:R
Hy:(0 < y)%R
Hxy:(y <= x)%R
Hx:(0 < x)%R
ex:Z
Hex0:(bpow (ex - 1) <= Rabs x)%R
Hex1:(Rabs x < bpow ex)%R
Haxy:(Rabs (x + y) < bpow (ex + 1))%R
Haxy2:(bpow (ex - 1) <= Rabs (x + y))%R

(mag (x + y) <= ex + 1)%Z
r:radix
Hr:(2 <= r)%Z
x, y:R
Hy:(0 < y)%R
Hxy:(y <= x)%R
Hx:(0 < x)%R
ex:Z
Hex0:(bpow (ex - 1) <= Rabs x)%R
Hex1:(Rabs x < bpow ex)%R
Haxy:(Rabs (x + y) < bpow (ex + 1))%R
Haxy2:(bpow (ex - 1) <= Rabs (x + y))%R

(x + y)%R <> 0%R
r:radix
Hr:(2 <= r)%Z
x, y:R
Hy:(0 < y)%R
Hxy:(y <= x)%R
Hx:(0 < x)%R
ex:Z
Hex0:(bpow (ex - 1) <= Rabs x)%R
Hex1:(Rabs x < bpow ex)%R
Haxy:(Rabs (x + y) < bpow (ex + 1))%R
Haxy2:(bpow (ex - 1) <= Rabs (x + y))%R
(Rabs (x + y) < bpow (ex + 1))%R
r:radix
Hr:(2 <= r)%Z
x, y:R
Hy:(0 < y)%R
Hxy:(y <= x)%R
Hx:(0 < x)%R
ex:Z
Hex0:(bpow (ex - 1) <= Rabs x)%R
Hex1:(Rabs x < bpow ex)%R
Haxy:(Rabs (x + y) < bpow (ex + 1))%R
Haxy2:(bpow (ex - 1) <= Rabs (x + y))%R

(x + y)%R <> 0%R
now apply tech_Rplus; [apply Rlt_le|].
r:radix
Hr:(2 <= r)%Z
x, y:R
Hy:(0 < y)%R
Hxy:(y <= x)%R
Hx:(0 < x)%R
ex:Z
Hex0:(bpow (ex - 1) <= Rabs x)%R
Hex1:(Rabs x < bpow ex)%R
Haxy:(Rabs (x + y) < bpow (ex + 1))%R
Haxy2:(bpow (ex - 1) <= Rabs (x + y))%R

(Rabs (x + y) < bpow (ex + 1))%R
exact Haxy. Qed.
r:radix

forall x y : R, (0 < y)%R -> (y < x)%R -> (mag (x - y) <= mag x)%Z
r:radix

forall x y : R, (0 < y)%R -> (y < x)%R -> (mag (x - y) <= mag x)%Z
r:radix
x, y:R
Py:(0 < y)%R
Hxy:(y < x)%R

(mag (x - y) <= mag x)%Z
r:radix
x, y:R
Py:(0 < y)%R
Hxy:(y < x)%R
Px:(0 < x)%R

(mag (x - y) <= mag x)%Z
r:radix
x, y:R
Py:(0 < y)%R
Hxy:(y < x)%R
Px:(0 < x)%R

(0 < x - y)%R
r:radix
x, y:R
Py:(0 < y)%R
Hxy:(y < x)%R
Px:(0 < x)%R
(x - y <= x)%R
r:radix
x, y:R
Py:(0 < y)%R
Hxy:(y < x)%R
Px:(0 < x)%R

(0 < x - y)%R
now apply Rlt_Rminus.
r:radix
x, y:R
Py:(0 < y)%R
Hxy:(y < x)%R
Px:(0 < x)%R

(x - y <= x)%R
r:radix
x, y:R
Py:(0 < y)%R
Hxy:(y < x)%R
Px:(0 < x)%R

(x - y <= x + 0)%R
r:radix
x, y:R
Py:(0 < y)%R
Hxy:(y < x)%R
Px:(0 < x)%R

(- y <= 0)%R
r:radix
x, y:R
Py:(0 < y)%R
Hxy:(y < x)%R
Px:(0 < x)%R

(- y <= - 0)%R
now apply Ropp_le_contravar; apply Rlt_le. Qed.
r:radix

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

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

(2 <= r)%Z
r:radix
Hbeta:(2 <= r)%Z
forall x y : R, (0 < x)%R -> (0 < y)%R -> (mag y <= mag x - 2)%Z -> (mag x - 1 <= mag (x - y))%Z
r:radix

(2 <= r)%Z
r: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.
r:radix
Hbeta:(2 <= r)%Z

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

(mag x - 1 <= mag (x - y))%Z
r:radix
Hbeta:(2 <= r)%Z
x, y:R
Px:(0 < x)%R
Py:(0 < y)%R
Hln:(mag y <= mag x - 2)%Z
Oxy:(y < x)%R

(mag x - 1 <= mag (x - y))%Z
r:radix
Hbeta:(2 <= r)%Z
x, y:R
Px:(0 < x)%R
Py:(0 < y)%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
Hln:(mag y <= Build_mag_prop x ex Hex - 2)%Z
Oxy:(y < x)%R

(Build_mag_prop x ex Hex - 1 <= mag (x - y))%Z
r:radix
Hbeta:(2 <= r)%Z
x, y:R
Px:(0 < x)%R
Py:(0 < y)%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hln:(Build_mag_prop y ey Hey <= Build_mag_prop x ex Hex - 2)%Z
Oxy:(y < x)%R

(Build_mag_prop x ex Hex - 1 <= mag (x - y))%Z
r:radix
Hbeta:(2 <= r)%Z
x, y:R
Px:(0 < x)%R
Py:(0 < y)%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hln:(ey <= ex - 2)%Z
Oxy:(y < x)%R

(ex - 1 <= mag (x - y))%Z
r:radix
Hbeta:(2 <= r)%Z
x, y:R
Px:(0 < x)%R
Py:(0 < y)%R
ex, ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Hln:(ey <= ex - 2)%Z
Oxy:(y < x)%R
Hex:(bpow (ex - 1) <= Rabs x)%R

(ex - 1 <= mag (x - y))%Z
r:radix
Hbeta:(2 <= r)%Z
x, y:R
Px:(0 < x)%R
Py:(0 < y)%R
ex, ey:Z
Hln:(ey <= ex - 2)%Z
Oxy:(y < x)%R
Hex:(bpow (ex - 1) <= Rabs x)%R
Hey:(Rabs y < bpow ey)%R

(ex - 1 <= mag (x - y))%Z
r:radix
Hbeta:(2 <= r)%Z
x, y:R
Px:(0 < x)%R
Py:(0 < y)%R
ex, ey:Z
Hln:(ey <= ex - 2)%Z
Oxy:(y < x)%R
Hex:(bpow (ex - 1) <= Rabs x)%R
Hey:(Rabs y < bpow ey)%R

(bpow (ex - 2) + bpow (ex - 2) <= x)%R
r:radix
Hbeta:(2 <= r)%Z
x, y:R
Px:(0 < x)%R
Py:(0 < y)%R
ex, ey:Z
Hln:(ey <= ex - 2)%Z
Oxy:(y < x)%R
Hex:(bpow (ex - 1) <= Rabs x)%R
Hey:(Rabs y < bpow ey)%R
Hbx:(bpow (ex - 2) + bpow (ex - 2) <= x)%R
(ex - 1 <= mag (x - y))%Z
r:radix
Hbeta:(2 <= r)%Z
x, y:R
Px:(0 < x)%R
Py:(0 < y)%R
ex, ey:Z
Hln:(ey <= ex - 2)%Z
Oxy:(y < x)%R
Hex:(bpow (ex - 1) <= Rabs x)%R
Hey:(Rabs y < bpow ey)%R

(bpow (ex - 2) + bpow (ex - 2) <= x)%R
r:radix
Hbeta:(2 <= r)%Z
x, y:R
Px:(0 < x)%R
Py:(0 < y)%R
ex, ey:Z
Hln:(ey <= ex - 2)%Z
Oxy:(y < x)%R
Hex:(bpow (ex - 1) <= Rabs x)%R
Hey:(Rabs y < bpow ey)%R

(2 * bpow (ex - 2) <= x)%R
r:radix
Hbeta:(2 <= r)%Z
x, y:R
Px:(0 < x)%R
Py:(0 < y)%R
ex, ey:Z
Hln:(ey <= ex - 2)%Z
Oxy:(y < x)%R
Hex:(bpow (ex - 1) <= Rabs x)%R
Hey:(Rabs y < bpow ey)%R

(2 * bpow (ex - 2) <= bpow (ex - 1))%R
r:radix
Hbeta:(2 <= r)%Z
x, y:R
Px:(0 < x)%R
Py:(0 < y)%R
ex, ey:Z
Hln:(ey <= ex - 2)%Z
Oxy:(y < x)%R
Hex:(bpow (ex - 1) <= Rabs x)%R
Hey:(Rabs y < bpow ey)%R
(bpow (ex - 1) <= x)%R
r:radix
Hbeta:(2 <= r)%Z
x, y:R
Px:(0 < x)%R
Py:(0 < y)%R
ex, ey:Z
Hln:(ey <= ex - 2)%Z
Oxy:(y < x)%R
Hex:(bpow (ex - 1) <= Rabs x)%R
Hey:(Rabs y < bpow ey)%R

(2 * bpow (ex - 2) <= bpow (ex - 1))%R
r:radix
Hbeta:(2 <= r)%Z
x, y:R
Px:(0 < x)%R
Py:(0 < y)%R
ex, ey:Z
Hln:(ey <= ex - 2)%Z
Oxy:(y < x)%R
Hex:(bpow (ex - 1) <= Rabs x)%R
Hey:(Rabs y < bpow ey)%R

(2 * bpow (ex - 2) <= bpow (ex - 2 + 1))%R
r:radix
Hbeta:(2 <= r)%Z
x, y:R
Px:(0 < x)%R
Py:(0 < y)%R
ex, ey:Z
Hln:(ey <= ex - 2)%Z
Oxy:(y < x)%R
Hex:(bpow (ex - 1) <= Rabs x)%R
Hey:(Rabs y < bpow ey)%R

(2 * bpow (ex - 2) <= IZR r * bpow (ex - 2))%R
r:radix
Hbeta:(2 <= r)%Z
x, y:R
Px:(0 < x)%R
Py:(0 < y)%R
ex, ey:Z
Hln:(ey <= ex - 2)%Z
Oxy:(y < x)%R
Hex:(bpow (ex - 1) <= Rabs x)%R
Hey:(Rabs y < bpow ey)%R

(2 <= IZR r)%R
now apply IZR_le.
r:radix
Hbeta:(2 <= r)%Z
x, y:R
Px:(0 < x)%R
Py:(0 < y)%R
ex, ey:Z
Hln:(ey <= ex - 2)%Z
Oxy:(y < x)%R
Hex:(bpow (ex - 1) <= Rabs x)%R
Hey:(Rabs y < bpow ey)%R

(bpow (ex - 1) <= x)%R
now rewrite Rabs_right in Hex; [|apply Rle_ge; apply Rlt_le].
r:radix
Hbeta:(2 <= r)%Z
x, y:R
Px:(0 < x)%R
Py:(0 < y)%R
ex, ey:Z
Hln:(ey <= ex - 2)%Z
Oxy:(y < x)%R
Hex:(bpow (ex - 1) <= Rabs x)%R
Hey:(Rabs y < bpow ey)%R
Hbx:(bpow (ex - 2) + bpow (ex - 2) <= x)%R

(ex - 1 <= mag (x - y))%Z
r:radix
Hbeta:(2 <= r)%Z
x, y:R
Px:(0 < x)%R
Py:(0 < y)%R
ex, ey:Z
Hln:(ey <= ex - 2)%Z
Oxy:(y < x)%R
Hex:(bpow (ex - 1) <= Rabs x)%R
Hey:(Rabs y < bpow ey)%R
Hbx:(bpow (ex - 2) + bpow (ex - 2) <= x)%R

(y < bpow (ex - 2))%R
r:radix
Hbeta:(2 <= r)%Z
x, y:R
Px:(0 < x)%R
Py:(0 < y)%R
ex, ey:Z
Hln:(ey <= ex - 2)%Z
Oxy:(y < x)%R
Hex:(bpow (ex - 1) <= Rabs x)%R
Hey:(Rabs y < bpow ey)%R
Hbx:(bpow (ex - 2) + bpow (ex - 2) <= x)%R
Hby:(y < bpow (ex - 2))%R
(ex - 1 <= mag (x - y))%Z
r:radix
Hbeta:(2 <= r)%Z
x, y:R
Px:(0 < x)%R
Py:(0 < y)%R
ex, ey:Z
Hln:(ey <= ex - 2)%Z
Oxy:(y < x)%R
Hex:(bpow (ex - 1) <= Rabs x)%R
Hey:(Rabs y < bpow ey)%R
Hbx:(bpow (ex - 2) + bpow (ex - 2) <= x)%R

(y < bpow (ex - 2))%R
r:radix
Hbeta:(2 <= r)%Z
x, y:R
Px:(0 < x)%R
Py:(0 < y)%R
ex, ey:Z
Hln:(ey <= ex - 2)%Z
Oxy:(y < x)%R
Hex:(bpow (ex - 1) <= Rabs x)%R
Hey:(Rabs y < bpow ey)%R
Hbx:(bpow (ex - 2) + bpow (ex - 2) <= x)%R

(y < bpow ey)%R
r:radix
Hbeta:(2 <= r)%Z
x, y:R
Px:(0 < x)%R
Py:(0 < y)%R
ex, ey:Z
Hln:(ey <= ex - 2)%Z
Oxy:(y < x)%R
Hex:(bpow (ex - 1) <= Rabs x)%R
Hey:(Rabs y < bpow ey)%R
Hbx:(bpow (ex - 2) + bpow (ex - 2) <= x)%R
(bpow ey <= bpow (ex - 2))%R
r:radix
Hbeta:(2 <= r)%Z
x, y:R
Px:(0 < x)%R
Py:(0 < y)%R
ex, ey:Z
Hln:(ey <= ex - 2)%Z
Oxy:(y < x)%R
Hex:(bpow (ex - 1) <= Rabs x)%R
Hey:(Rabs y < bpow ey)%R
Hbx:(bpow (ex - 2) + bpow (ex - 2) <= x)%R

(bpow ey <= bpow (ex - 2))%R
now apply bpow_le.
r:radix
Hbeta:(2 <= r)%Z
x, y:R
Px:(0 < x)%R
Py:(0 < y)%R
ex, ey:Z
Hln:(ey <= ex - 2)%Z
Oxy:(y < x)%R
Hex:(bpow (ex - 1) <= Rabs x)%R
Hey:(Rabs y < bpow ey)%R
Hbx:(bpow (ex - 2) + bpow (ex - 2) <= x)%R
Hby:(y < bpow (ex - 2))%R

(ex - 1 <= mag (x - y))%Z
r:radix
Hbeta:(2 <= r)%Z
x, y:R
Px:(0 < x)%R
Py:(0 < y)%R
ex, ey:Z
Hln:(ey <= ex - 2)%Z
Oxy:(y < x)%R
Hex:(bpow (ex - 1) <= Rabs x)%R
Hey:(Rabs y < bpow ey)%R
Hbx:(bpow (ex - 2) + bpow (ex - 2) <= x)%R
Hby:(y < bpow (ex - 2))%R

(bpow (ex - 2) <= x - y)%R
r:radix
Hbeta:(2 <= r)%Z
x, y:R
Px:(0 < x)%R
Py:(0 < y)%R
ex, ey:Z
Hln:(ey <= ex - 2)%Z
Oxy:(y < x)%R
Hex:(bpow (ex - 1) <= Rabs x)%R
Hey:(Rabs y < bpow ey)%R
Hbx:(bpow (ex - 2) + bpow (ex - 2) <= x)%R
Hby:(y < bpow (ex - 2))%R
Hbxy:(bpow (ex - 2) <= x - y)%R
(ex - 1 <= mag (x - y))%Z
r:radix
Hbeta:(2 <= r)%Z
x, y:R
Px:(0 < x)%R
Py:(0 < y)%R
ex, ey:Z
Hln:(ey <= ex - 2)%Z
Oxy:(y < x)%R
Hex:(bpow (ex - 1) <= Rabs x)%R
Hey:(Rabs y < bpow ey)%R
Hbx:(bpow (ex - 2) + bpow (ex - 2) <= x)%R
Hby:(y < bpow (ex - 2))%R

(bpow (ex - 2) <= x - y)%R
r:radix
Hbeta:(2 <= r)%Z
x, y:R
Px:(0 < x)%R
Py:(0 < y)%R
ex, ey:Z
Hln:(ey <= ex - 2)%Z
Oxy:(y < x)%R
Hex:(bpow (ex - 1) <= Rabs x)%R
Hey:(Rabs y < bpow ey)%R
Hbx:(bpow (ex - 2) + bpow (ex - 2) <= x)%R
Hby:(- bpow (ex - 2) < - y)%R

(bpow (ex - 2) <= x - y)%R
r:radix
Hbeta:(2 <= r)%Z
x, y:R
Px:(0 < x)%R
Py:(0 < y)%R
ex, ey:Z
Hln:(ey <= ex - 2)%Z
Oxy:(y < x)%R
Hex:(bpow (ex - 1) <= Rabs x)%R
Hey:(Rabs y < bpow ey)%R
Hbx:(bpow (ex - 2) + bpow (ex - 2) <= x)%R
Hby:(- bpow (ex - 2) <= - y)%R

(bpow (ex - 2) <= x - y)%R
r:radix
Hbeta:(2 <= r)%Z
x, y:R
Px:(0 < x)%R
Py:(0 < y)%R
ex, ey:Z
Hln:(ey <= ex - 2)%Z
Oxy:(y < x)%R
Hex:(bpow (ex - 1) <= Rabs x)%R
Hey:(Rabs y < bpow ey)%R
Hbx:(bpow (ex - 2) + bpow (ex - 2) <= x)%R
Hby:(- bpow (ex - 2) <= - y)%R

(bpow (ex - 2) + bpow (ex - 2) - bpow (ex - 2) <= x - y)%R
now apply Rplus_le_compat.
r:radix
Hbeta:(2 <= r)%Z
x, y:R
Px:(0 < x)%R
Py:(0 < y)%R
ex, ey:Z
Hln:(ey <= ex - 2)%Z
Oxy:(y < x)%R
Hex:(bpow (ex - 1) <= Rabs x)%R
Hey:(Rabs y < bpow ey)%R
Hbx:(bpow (ex - 2) + bpow (ex - 2) <= x)%R
Hby:(y < bpow (ex - 2))%R
Hbxy:(bpow (ex - 2) <= x - y)%R

(ex - 1 <= mag (x - y))%Z
r:radix
Hbeta:(2 <= r)%Z
x, y:R
Px:(0 < x)%R
Py:(0 < y)%R
ex, ey:Z
Hln:(ey <= ex - 2)%Z
Oxy:(y < x)%R
Hex:(bpow (ex - 1) <= Rabs x)%R
Hey:(Rabs y < bpow ey)%R
Hbx:(bpow (ex - 2) + bpow (ex - 2) <= x)%R
Hby:(y < bpow (ex - 2))%R
Hbxy:(bpow (ex - 2) <= x - y)%R

(bpow (ex - 1 - 1) <= Rabs (x - y))%R
r:radix
Hbeta:(2 <= r)%Z
x, y:R
Px:(0 < x)%R
Py:(0 < y)%R
ex, ey:Z
Hln:(ey <= ex - 2)%Z
Oxy:(y < x)%R
Hex:(bpow (ex - 1) <= Rabs x)%R
Hey:(Rabs y < bpow ey)%R
Hbx:(bpow (ex - 2) + bpow (ex - 2) <= x)%R
Hby:(y < bpow (ex - 2))%R
Hbxy:(bpow (ex - 2) <= x - y)%R

(bpow (ex - 2) <= Rabs (x - y))%R
now apply Rabs_ge; right. Qed.
r:radix

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

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

(mag x - mag y <= mag (x / y) <= mag x - mag y + 1)%Z
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R

(Build_mag_prop x ex Hex - mag y <= mag (x / y) <= Build_mag_prop x ex Hex - mag y + 1)%Z
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R

(Build_mag_prop x ex Hex - Build_mag_prop y ey Hey <= mag (x / y) <= Build_mag_prop x ex Hex - Build_mag_prop y ey Hey + 1)%Z
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R

(ex - ey <= mag (x / y) <= ex - ey + 1)%Z
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R

(ex - ey <= mag (x * / y) <= ex - ey + 1)%Z
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R

(bpow (- ey) < Rabs (/ y) <= bpow (- ey + 1))%R
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Heiy:(bpow (- ey) < Rabs (/ y) <= bpow (- ey + 1))%R
(ex - ey <= mag (x * / y) <= ex - ey + 1)%Z
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R

(bpow (- ey) < Rabs (/ y) <= bpow (- ey + 1))%R
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R

(bpow (- ey) < / Rabs y <= bpow (- ey + 1))%R
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R

(bpow (- ey) < / Rabs y)%R
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
(/ Rabs y <= bpow (- ey + 1))%R
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R

(bpow (- ey) < / Rabs y)%R
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R

(/ bpow ey < / Rabs y)%R
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R

(0 < Rabs y * bpow ey)%R
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
(Rabs y < bpow ey)%R
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R

(0 < Rabs y * bpow ey)%R
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R

(0 < Rabs y)%R
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
(0 < bpow ey)%R
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R

(0 < bpow ey)%R
now apply bpow_gt_0.
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R

(Rabs y < bpow ey)%R
now apply Hey.
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R

(/ Rabs y <= bpow (- ey + 1))%R
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R

(/ Rabs y <= bpow (- (ey - 1)))%R
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R

(/ Rabs y <= / bpow (ey - 1))%R
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R

(bpow (ey - 1) <= Rabs y)%R
now apply Hey.
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Heiy:(bpow (- ey) < Rabs (/ y) <= bpow (- ey + 1))%R

(ex - ey <= mag (x * / y) <= ex - ey + 1)%Z
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Heiy:(bpow (- ey) < Rabs (/ y) <= bpow (- ey + 1))%R

(ex - ey <= mag (x * / y))%Z
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Heiy:(bpow (- ey) < Rabs (/ y) <= bpow (- ey + 1))%R
(mag (x * / y) <= ex - ey + 1)%Z
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Heiy:(bpow (- ey) < Rabs (/ y) <= bpow (- ey + 1))%R

(ex - ey <= mag (x * / y))%Z
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Heiy:(bpow (- ey) < Rabs (/ y) <= bpow (- ey + 1))%R

(bpow (ex - ey - 1) <= Rabs (x * / y))%R
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Heiy:(bpow (- ey) < Rabs (/ y) <= bpow (- ey + 1))%R

(bpow (ex - 1 - ey) <= Rabs (x * / y))%R
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Heiy:(bpow (- ey) < Rabs (/ y) <= bpow (- ey + 1))%R

(bpow (ex - 1) * bpow (- ey) <= Rabs (x * / y))%R
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Heiy:(bpow (- ey) < Rabs (/ y) <= bpow (- ey + 1))%R

(bpow (ex - 1) * bpow (- ey) <= Rabs x * Rabs (/ y))%R
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Heiy:(bpow (- ey) < Rabs (/ y) <= bpow (- ey + 1))%R

(0 <= bpow (ex - 1))%R
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Heiy:(bpow (- ey) < Rabs (/ y) <= bpow (- ey + 1))%R
(0 <= bpow (- ey))%R
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Heiy:(bpow (- ey) < Rabs (/ y) <= bpow (- ey + 1))%R
(bpow (ex - 1) <= Rabs x)%R
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Heiy:(bpow (- ey) < Rabs (/ y) <= bpow (- ey + 1))%R
(bpow (- ey) <= Rabs (/ y))%R
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Heiy:(bpow (- ey) < Rabs (/ y) <= bpow (- ey + 1))%R

(0 <= bpow (ex - 1))%R
now apply bpow_ge_0.
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Heiy:(bpow (- ey) < Rabs (/ y) <= bpow (- ey + 1))%R

(0 <= bpow (- ey))%R
now apply bpow_ge_0.
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Heiy:(bpow (- ey) < Rabs (/ y) <= bpow (- ey + 1))%R

(bpow (ex - 1) <= Rabs x)%R
now apply Hex.
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Heiy:(bpow (- ey) < Rabs (/ y) <= bpow (- ey + 1))%R

(bpow (- ey) <= Rabs (/ y))%R
now apply Rlt_le; apply Heiy.
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Heiy:(bpow (- ey) < Rabs (/ y) <= bpow (- ey + 1))%R

(mag (x * / y) <= ex - ey + 1)%Z
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Heiy:(bpow (- ey) < Rabs (/ y) <= bpow (- ey + 1))%R

(x * / y)%R <> 0%R
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Heiy:(bpow (- ey) < Rabs (/ y) <= bpow (- ey + 1))%R
(Rabs (x * / y) < bpow (ex - ey + 1))%R
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Heiy:(bpow (- ey) < Rabs (/ y) <= bpow (- ey + 1))%R

(x * / y)%R <> 0%R
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Heiy:(bpow (- ey) < Rabs (/ y) <= bpow (- ey + 1))%R

x <> 0%R
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Heiy:(bpow (- ey) < Rabs (/ y) <= bpow (- ey + 1))%R
(/ y)%R <> 0%R
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Heiy:(bpow (- ey) < Rabs (/ y) <= bpow (- ey + 1))%R

(/ y)%R <> 0%R
now apply Rinv_neq_0_compat.
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Heiy:(bpow (- ey) < Rabs (/ y) <= bpow (- ey + 1))%R

(Rabs (x * / y) < bpow (ex - ey + 1))%R
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Heiy:(bpow (- ey) < Rabs (/ y) <= bpow (- ey + 1))%R

(Rabs (x * / y) < bpow (ex + (- ey + 1)))%R
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Heiy:(bpow (- ey) < Rabs (/ y) <= bpow (- ey + 1))%R

(Rabs (x * / y) < bpow ex * bpow (- ey + 1))%R
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Heiy:(bpow (- ey) < Rabs (/ y) <= bpow (- ey + 1))%R

(Rabs (x * / y) < bpow ex * Rabs (/ y))%R
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Heiy:(bpow (- ey) < Rabs (/ y) <= bpow (- ey + 1))%R
(bpow ex * Rabs (/ y) <= bpow ex * bpow (- ey + 1))%R
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Heiy:(bpow (- ey) < Rabs (/ y) <= bpow (- ey + 1))%R

(Rabs (x * / y) < bpow ex * Rabs (/ y))%R
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Heiy:(bpow (- ey) < Rabs (/ y) <= bpow (- ey + 1))%R

(Rabs x * Rabs (/ y) < bpow ex * Rabs (/ y))%R
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Heiy:(bpow (- ey) < Rabs (/ y) <= bpow (- ey + 1))%R

(0 < Rabs (/ y))%R
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Heiy:(bpow (- ey) < Rabs (/ y) <= bpow (- ey + 1))%R
(Rabs x < bpow ex)%R
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Heiy:(bpow (- ey) < Rabs (/ y) <= bpow (- ey + 1))%R

(/ y)%R <> 0%R
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Heiy:(bpow (- ey) < Rabs (/ y) <= bpow (- ey + 1))%R
(Rabs x < bpow ex)%R
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Heiy:(bpow (- ey) < Rabs (/ y) <= bpow (- ey + 1))%R

(Rabs x < bpow ex)%R
now apply Hex.
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Heiy:(bpow (- ey) < Rabs (/ y) <= bpow (- ey + 1))%R

(bpow ex * Rabs (/ y) <= bpow ex * bpow (- ey + 1))%R
r:radix
x, y:R
Px:x <> 0%R
Py:y <> 0%R
ex:Z
Hex:x <> 0%R -> (bpow (ex - 1) <= Rabs x < bpow ex)%R
ey:Z
Hey:y <> 0%R -> (bpow (ey - 1) <= Rabs y < bpow ey)%R
Heiy:(bpow (- ey) < Rabs (/ y) <= bpow (- ey + 1))%R

(Rabs (/ y) <= bpow (- ey + 1))%R
apply Heiy. Qed.
r:radix

forall x : R, (0 < x)%R -> mag (sqrt x) = Z.div2 (mag x + 1)
r:radix

forall x : R, (0 < x)%R -> mag (sqrt x) = Z.div2 (mag x + 1)
r:radix
x:R
Px:(0 < x)%R

mag (sqrt x) = Z.div2 (mag x + 1)
r:radix
x:R
Px:(0 < x)%R

(bpow (Z.div2 (mag x + 1) - 1) <= Rabs (sqrt x) < bpow (Z.div2 (mag x + 1)))%R
r:radix
x:R
Px:(0 < x)%R
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R

(bpow (Z.div2 (Build_mag_prop x e He + 1) - 1) <= Rabs (sqrt x) < bpow (Z.div2 (Build_mag_prop x e He + 1)))%R
r:radix
x:R
Px:(0 < x)%R
e:Z
He:x <> 0%R -> (bpow (e - 1) <= Rabs x < bpow e)%R

(bpow (Z.div2 (e + 1) - 1) <= Rabs (sqrt x) < bpow (Z.div2 (e + 1)))%R
r:radix
x:R
Px:(0 < x)%R
e:Z
He:(bpow (e - 1) <= Rabs x < bpow e)%R

(bpow (Z.div2 (e + 1) - 1) <= Rabs (sqrt x) < bpow (Z.div2 (e + 1)))%R
r:radix
x:R
Px:(0 < x)%R
e:Z
He:(bpow (e - 1) <= x < bpow e)%R

(bpow (Z.div2 (e + 1) - 1) <= Rabs (sqrt x) < bpow (Z.div2 (e + 1)))%R
r:radix
x:R
Px:(0 < x)%R
e:Z
He:(bpow (e - 1) <= x < bpow e)%R

(bpow (Z.div2 (e + 1) - 1) <= Rabs (sqrt x))%R
r:radix
x:R
Px:(0 < x)%R
e:Z
He:(bpow (e - 1) <= x < bpow e)%R
(Rabs (sqrt x) < bpow (Z.div2 (e + 1)))%R
r:radix
x:R
Px:(0 < x)%R
e:Z
He:(bpow (e - 1) <= x < bpow e)%R

(bpow (Z.div2 (e + 1) - 1) <= Rabs (sqrt x))%R
r:radix
x:R
Px:(0 < x)%R
e:Z
He:(bpow (e - 1) <= x < bpow e)%R

(Rabs (bpow (Z.div2 (e + 1) - 1)) <= Rabs (sqrt x))%R
r:radix
x:R
Px:(0 < x)%R
e:Z
He:(bpow (e - 1) <= x < bpow e)%R

((bpow (Z.div2 (e + 1) - 1))² <= (sqrt x)²)%R
r:radix
x:R
Px:(0 < x)%R
e:Z
He:(bpow (e - 1) <= x < bpow e)%R

((bpow (Z.div2 (e + 1) - 1))² <= x)%R
r:radix
x:R
Px:(0 < x)%R
e:Z
He:(bpow (e - 1) <= x < bpow e)%R

((bpow (Z.div2 (e + 1) - 1))² <= bpow (e - 1))%R
r:radix
x:R
Px:(0 < x)%R
e:Z
He:(bpow (e - 1) <= x < bpow e)%R

(bpow (Z.div2 (e + 1) - 1 + (Z.div2 (e + 1) - 1)) <= bpow (e - 1))%R
r:radix
x:R
Px:(0 < x)%R
e:Z
He:(bpow (e - 1) <= x < bpow e)%R

(Z.div2 (e + 1) - 1 + (Z.div2 (e + 1) - 1) <= e - 1)%Z
r:radix
x:R
Px:(0 < x)%R
e:Z
He:(bpow (e - 1) <= x < bpow e)%R

(e + 1)%Z = (2 * Z.div2 (e + 1) + (if Z.odd (e + 1) then 1 else 0))%Z -> (Z.div2 (e + 1) - 1 + (Z.div2 (e + 1) - 1) <= e - 1)%Z
destruct Z.odd ; intros ; omega.
r:radix
x:R
Px:(0 < x)%R
e:Z
He:(bpow (e - 1) <= x < bpow e)%R

(Rabs (sqrt x) < bpow (Z.div2 (e + 1)))%R
r:radix
x:R
Px:(0 < x)%R
e:Z
He:(bpow (e - 1) <= x < bpow e)%R

(Rabs (sqrt x) < Rabs (bpow (Z.div2 (e + 1))))%R
r:radix
x:R
Px:(0 < x)%R
e:Z
He:(bpow (e - 1) <= x < bpow e)%R

((sqrt x)² < (bpow (Z.div2 (e + 1)))²)%R
r:radix
x:R
Px:(0 < x)%R
e:Z
He:(bpow (e - 1) <= x < bpow e)%R

(x < (bpow (Z.div2 (e + 1)))²)%R
r:radix
x:R
Px:(0 < x)%R
e:Z
He:(bpow (e - 1) <= x < bpow e)%R

(bpow e <= (bpow (Z.div2 (e + 1)))²)%R
r:radix
x:R
Px:(0 < x)%R
e:Z
He:(bpow (e - 1) <= x < bpow e)%R

(bpow e <= bpow (Z.div2 (e + 1) + Z.div2 (e + 1)))%R
r:radix
x:R
Px:(0 < x)%R
e:Z
He:(bpow (e - 1) <= x < bpow e)%R

(e <= Z.div2 (e + 1) + Z.div2 (e + 1))%Z
r:radix
x:R
Px:(0 < x)%R
e:Z
He:(bpow (e - 1) <= x < bpow e)%R

(e + 1)%Z = (2 * Z.div2 (e + 1) + (if Z.odd (e + 1) then 1 else 0))%Z -> (e <= Z.div2 (e + 1) + Z.div2 (e + 1))%Z
destruct Z.odd ; intros ; omega. Qed.
r:radix

mag 1 = 1%Z
r:radix

mag 1 = 1%Z
r:radix

(1 < r)%Z
r:radix

(2 <= r)%Z -> (1 < r)%Z
now apply Z.lt_le_trans. Qed. End pow. Section Bool.

forall x y : bool, Bool.eqb x y = Bool.eqb y x

forall x y : bool, Bool.eqb x y = Bool.eqb y x
now intros [|] [|]. Qed.

forall x y : bool, x = negb y -> Bool.eqb x y = false

forall x y : bool, x = negb y -> Bool.eqb x y = false
now intros [|] [|]. Qed.

forall x y : bool, x = y -> Bool.eqb x y = true

forall x y : bool, x = y -> Bool.eqb x y = true
now intros [|] [|]. Qed. End Bool. Section cond_Ropp. Definition cond_Ropp (b : bool) m := if b then Ropp m else m.

forall (b : bool) (m : Z), IZR (SpecFloatCopy.cond_Zopp b m) = cond_Ropp b (IZR m)

forall (b : bool) (m : Z), IZR (SpecFloatCopy.cond_Zopp b m) = cond_Ropp b (IZR m)
m:Z

IZR (SpecFloatCopy.cond_Zopp true m) = cond_Ropp true (IZR m)
m:Z
IZR (SpecFloatCopy.cond_Zopp false m) = cond_Ropp false (IZR m)
m:Z

IZR (SpecFloatCopy.cond_Zopp false m) = cond_Ropp false (IZR m)
apply refl_equal. Qed.

forall (b : bool) (m : R), Rabs (cond_Ropp b m) = Rabs m

forall (b : bool) (m : R), Rabs (cond_Ropp b m) = Rabs m
m:R

Rabs (cond_Ropp true m) = Rabs m
m:R
Rabs (cond_Ropp false m) = Rabs m
m:R

Rabs (cond_Ropp false m) = Rabs m
apply refl_equal. Qed.

forall m : R, cond_Ropp (Rlt_bool m 0) m = Rabs m

forall m : R, cond_Ropp (Rlt_bool m 0) m = Rabs m
m:R

cond_Ropp (Rlt_bool m 0) m = Rabs m
m:R

Rabs m = cond_Ropp (Rlt_bool m 0) m
m:R
Hm:(m < 0)%R

Rabs m = cond_Ropp true m
m:R
Hm:(0 <= m)%R
Rabs m = cond_Ropp false m
m:R
Hm:(0 <= m)%R

Rabs m = cond_Ropp false m
now apply Rabs_pos_eq. Qed.

forall (b : bool) (x : R), cond_Ropp b (cond_Ropp b x) = x

forall (b : bool) (x : R), cond_Ropp b (cond_Ropp b x) = x
x:R

cond_Ropp true (cond_Ropp true x) = x
x:R
cond_Ropp false (cond_Ropp false x) = x
x:R

cond_Ropp false (cond_Ropp false x) = x
apply refl_equal. Qed.

forall (b : bool) (x y : R), cond_Ropp b x = cond_Ropp b y -> x = y

forall (b : bool) (x y : R), cond_Ropp b x = cond_Ropp b y -> x = y
b:bool
x, y:R
H:cond_Ropp b x = cond_Ropp b y

x = y
b:bool
x, y:R
H:cond_Ropp b x = cond_Ropp b y

cond_Ropp b (cond_Ropp b y) = y
apply cond_Ropp_involutive. Qed.

forall (b : bool) (x y : R), cond_Ropp b (x * y) = (cond_Ropp b x * y)%R

forall (b : bool) (x y : R), cond_Ropp b (x * y) = (cond_Ropp b x * y)%R
x, y:R

cond_Ropp true (x * y) = (cond_Ropp true x * y)%R
x, y:R
cond_Ropp false (x * y) = (cond_Ropp false x * y)%R
x, y:R

(cond_Ropp true x * y)%R = cond_Ropp true (x * y)
x, y:R
cond_Ropp false (x * y) = (cond_Ropp false x * y)%R
x, y:R

cond_Ropp false (x * y) = (cond_Ropp false x * y)%R
apply refl_equal. Qed.

forall (b : bool) (x y : R), cond_Ropp b (x * y) = (x * cond_Ropp b y)%R

forall (b : bool) (x y : R), cond_Ropp b (x * y) = (x * cond_Ropp b y)%R
x, y:R

cond_Ropp true (x * y) = (x * cond_Ropp true y)%R
x, y:R
cond_Ropp false (x * y) = (x * cond_Ropp false y)%R
x, y:R

(x * cond_Ropp true y)%R = cond_Ropp true (x * y)
x, y:R
cond_Ropp false (x * y) = (x * cond_Ropp false y)%R
x, y:R

cond_Ropp false (x * y) = (x * cond_Ropp false y)%R
apply refl_equal. Qed.

forall (b : bool) (x y : R), cond_Ropp b (x + y) = (cond_Ropp b x + cond_Ropp b y)%R

forall (b : bool) (x y : R), cond_Ropp b (x + y) = (cond_Ropp b x + cond_Ropp b y)%R
x, y:R

cond_Ropp true (x + y) = (cond_Ropp true x + cond_Ropp true y)%R
x, y:R
cond_Ropp false (x + y) = (cond_Ropp false x + cond_Ropp false y)%R
x, y:R

cond_Ropp false (x + y) = (cond_Ropp false x + cond_Ropp false y)%R
apply refl_equal. Qed. End cond_Ropp.
LPO taken from Coquelicot

forall P : nat -> Prop, (forall n : nat, P n \/ ~ P n) -> {n : nat | P n /\ (forall i : nat, i < n -> ~ P i)} + {forall n : nat, ~ P n}

forall P : nat -> Prop, (forall n : nat, P n \/ ~ P n) -> {n : nat | P n /\ (forall i : nat, i < n -> ~ P i)} + {forall n : nat, ~ P n}

forall n : nat, (0 < INR n + 1)%R
Hi:forall n : nat, (0 < INR n + 1)%R
forall P : nat -> Prop, (forall n : nat, P n \/ ~ P n) -> {n : nat | P n /\ (forall i : nat, i < n -> ~ P i)} + {forall n : nat, ~ P n}
N:nat

(0 < INR N + 1)%R
Hi:forall n : nat, (0 < INR n + 1)%R
forall P : nat -> Prop, (forall n : nat, P n \/ ~ P n) -> {n : nat | P n /\ (forall i : nat, i < n -> ~ P i)} + {forall n : nat, ~ P n}
N:nat

(0 < INR (S N))%R
Hi:forall n : nat, (0 < INR n + 1)%R
forall P : nat -> Prop, (forall n : nat, P n \/ ~ P n) -> {n : nat | P n /\ (forall i : nat, i < n -> ~ P i)} + {forall n : nat, ~ P n}
N:nat

0 < S N
Hi:forall n : nat, (0 < INR n + 1)%R
forall P : nat -> Prop, (forall n : nat, P n \/ ~ P n) -> {n : nat | P n /\ (forall i : nat, i < n -> ~ P i)} + {forall n : nat, ~ P n}
Hi:forall n : nat, (0 < INR n + 1)%R

forall P : nat -> Prop, (forall n : nat, P n \/ ~ P n) -> {n : nat | P n /\ (forall i : nat, i < n -> ~ P i)} + {forall n : nat, ~ P n}
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n

{n : nat | P n /\ (forall i : nat, i < n -> ~ P i)} + {forall n : nat, ~ P n}
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop

{n : nat | P n /\ (forall i : nat, i < n -> ~ P i)} + {forall n : nat, ~ P n}
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop

forall n : nat, P n -> E (/ (INR n + 1))%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
{n : nat | P n /\ (forall i : nat, i < n -> ~ P i)} + {forall n : nat, ~ P n}
Hi:forall n0 : nat, (0 < INR n0 + 1)%R
P:nat -> Prop
HP:forall n0 : nat, P n0 \/ ~ P n0
E:=fun y : R => exists n0 : nat, P n0 /\ y = (/ (INR n0 + 1))%R \/ ~ P n0 /\ y = 0%R:R -> Prop
n:nat
Pn:P n

E (/ (INR n + 1))%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
{n : nat | P n /\ (forall i : nat, i < n -> ~ P i)} + {forall n : nat, ~ P n}
Hi:forall n0 : nat, (0 < INR n0 + 1)%R
P:nat -> Prop
HP:forall n0 : nat, P n0 \/ ~ P n0
E:=fun y : R => exists n0 : nat, P n0 /\ y = (/ (INR n0 + 1))%R \/ ~ P n0 /\ y = 0%R:R -> Prop
n:nat
Pn:P n

P n /\ (/ (INR n + 1))%R = (/ (INR n + 1))%R \/ ~ P n /\ (/ (INR n + 1))%R = 0%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
{n : nat | P n /\ (forall i : nat, i < n -> ~ P i)} + {forall n : nat, ~ P n}
Hi:forall n0 : nat, (0 < INR n0 + 1)%R
P:nat -> Prop
HP:forall n0 : nat, P n0 \/ ~ P n0
E:=fun y : R => exists n0 : nat, P n0 /\ y = (/ (INR n0 + 1))%R \/ ~ P n0 /\ y = 0%R:R -> Prop
n:nat
Pn:P n

P n /\ (/ (INR n + 1))%R = (/ (INR n + 1))%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
{n : nat | P n /\ (forall i : nat, i < n -> ~ P i)} + {forall n : nat, ~ P n}
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R

{n : nat | P n /\ (forall i : nat, i < n -> ~ P i)} + {forall n : nat, ~ P n}
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R

is_upper_bound E 1
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
{n : nat | P n /\ (forall i : nat, i < n -> ~ P i)} + {forall n : nat, ~ P n}
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y0 : R => exists n : nat, P n /\ y0 = (/ (INR n + 1))%R \/ ~ P n /\ y0 = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
y:nat

(/ (INR y + 1) <= 1)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y0 : R => exists n : nat, P n /\ y0 = (/ (INR n + 1))%R \/ ~ P n /\ y0 = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
y:nat
(0 <= 1)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
{n : nat | P n /\ (forall i : nat, i < n -> ~ P i)} + {forall n : nat, ~ P n}
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y0 : R => exists n : nat, P n /\ y0 = (/ (INR n + 1))%R \/ ~ P n /\ y0 = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
y:nat

(/ (INR y + 1) <= / 1)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y0 : R => exists n : nat, P n /\ y0 = (/ (INR n + 1))%R \/ ~ P n /\ y0 = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
y:nat
(0 <= 1)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
{n : nat | P n /\ (forall i : nat, i < n -> ~ P i)} + {forall n : nat, ~ P n}
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y0 : R => exists n : nat, P n /\ y0 = (/ (INR n + 1))%R \/ ~ P n /\ y0 = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
y:nat

(0 < 1)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y0 : R => exists n : nat, P n /\ y0 = (/ (INR n + 1))%R \/ ~ P n /\ y0 = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
y:nat
(1 <= INR y + 1)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y0 : R => exists n : nat, P n /\ y0 = (/ (INR n + 1))%R \/ ~ P n /\ y0 = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
y:nat
(0 <= 1)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
{n : nat | P n /\ (forall i : nat, i < n -> ~ P i)} + {forall n : nat, ~ P n}
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y0 : R => exists n : nat, P n /\ y0 = (/ (INR n + 1))%R \/ ~ P n /\ y0 = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
y:nat

(1 <= INR y + 1)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y0 : R => exists n : nat, P n /\ y0 = (/ (INR n + 1))%R \/ ~ P n /\ y0 = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
y:nat
(0 <= 1)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
{n : nat | P n /\ (forall i : nat, i < n -> ~ P i)} + {forall n : nat, ~ P n}
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y0 : R => exists n : nat, P n /\ y0 = (/ (INR n + 1))%R \/ ~ P n /\ y0 = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
y:nat

(1 <= INR (S y))%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y0 : R => exists n : nat, P n /\ y0 = (/ (INR n + 1))%R \/ ~ P n /\ y0 = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
y:nat
(0 <= 1)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
{n : nat | P n /\ (forall i : nat, i < n -> ~ P i)} + {forall n : nat, ~ P n}
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y0 : R => exists n : nat, P n /\ y0 = (/ (INR n + 1))%R \/ ~ P n /\ y0 = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
y:nat

(0 <= 1)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
{n : nat | P n /\ (forall i : nat, i < n -> ~ P i)} + {forall n : nat, ~ P n}
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1

{n : nat | P n /\ (forall i : nat, i < n -> ~ P i)} + {forall n : nat, ~ P n}
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1

bound E
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
exists x : R, E x
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
{n : nat | P n /\ (forall i : nat, i < n -> ~ P i)} + {forall n : nat, ~ P n}
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1

exists x : R, E x
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
{n : nat | P n /\ (forall i : nat, i < n -> ~ P i)} + {forall n : nat, ~ P n}
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
H0:P 0

exists x : R, E x
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
H0:~ P 0
exists x : R, E x
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
{n : nat | P n /\ (forall i : nat, i < n -> ~ P i)} + {forall n : nat, ~ P n}
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
H0:P 0

E 1%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
H0:~ P 0
exists x : R, E x
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
{n : nat | P n /\ (forall i : nat, i < n -> ~ P i)} + {forall n : nat, ~ P n}
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
H0:P 0

P 0 /\ 1%R = (/ (INR 0 + 1))%R \/ ~ P 0 /\ 1%R = 0%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
H0:~ P 0
exists x : R, E x
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
{n : nat | P n /\ (forall i : nat, i < n -> ~ P i)} + {forall n : nat, ~ P n}
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
H0:P 0

P 0 /\ 1%R = (/ (INR 0 + 1))%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
H0:~ P 0
exists x : R, E x
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
{n : nat | P n /\ (forall i : nat, i < n -> ~ P i)} + {forall n : nat, ~ P n}
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
H0:P 0

1%R = (/ (INR 0 + 1))%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
H0:~ P 0
exists x : R, E x
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
{n : nat | P n /\ (forall i : nat, i < n -> ~ P i)} + {forall n : nat, ~ P n}
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
H0:P 0

1%R = (/ 1)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
H0:~ P 0
exists x : R, E x
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
{n : nat | P n /\ (forall i : nat, i < n -> ~ P i)} + {forall n : nat, ~ P n}
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
H0:~ P 0

exists x : R, E x
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
{n : nat | P n /\ (forall i : nat, i < n -> ~ P i)} + {forall n : nat, ~ P n}
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
H0:~ P 0

E 0%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
{n : nat | P n /\ (forall i : nat, i < n -> ~ P i)} + {forall n : nat, ~ P n}
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
H0:~ P 0

P 0 /\ 0%R = (/ (INR 0 + 1))%R \/ ~ P 0 /\ 0%R = 0%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
{n : nat | P n /\ (forall i : nat, i < n -> ~ P i)} + {forall n : nat, ~ P n}
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
H0:~ P 0

~ P 0 /\ 0%R = 0%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
{n : nat | P n /\ (forall i : nat, i < n -> ~ P i)} + {forall n : nat, ~ P n}
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R

{n : nat | P n /\ (forall i : nat, i < n -> ~ P i)} + {forall n : nat, ~ P n}
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(l <= 0)%R

{n : nat | P n /\ (forall i : nat, i < n -> ~ P i)} + {forall n : nat, ~ P n}
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
{n : nat | P n /\ (forall i : nat, i < n -> ~ P i)} + {forall n : nat, ~ P n}
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(l <= 0)%R

forall n : nat, ~ P n
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
{n : nat | P n /\ (forall i : nat, i < n -> ~ P i)} + {forall n : nat, ~ P n}
Hi:forall n0 : nat, (0 < INR n0 + 1)%R
P:nat -> Prop
HP:forall n0 : nat, P n0 \/ ~ P n0
E:=fun y : R => exists n0 : nat, P n0 /\ y = (/ (INR n0 + 1))%R \/ ~ P n0 /\ y = 0%R:R -> Prop
HE:forall n0 : nat, P n0 -> E (/ (INR n0 + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(l <= 0)%R
n:nat
Pn:P n

False
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
{n : nat | P n /\ (forall i : nat, i < n -> ~ P i)} + {forall n : nat, ~ P n}
Hi:forall n0 : nat, (0 < INR n0 + 1)%R
P:nat -> Prop
HP:forall n0 : nat, P n0 \/ ~ P n0
E:=fun y : R => exists n0 : nat, P n0 /\ y = (/ (INR n0 + 1))%R \/ ~ P n0 /\ y = 0%R:R -> Prop
HE:forall n0 : nat, P n0 -> E (/ (INR n0 + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(l <= 0)%R
n:nat
Pn:P n

(0 < l)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
{n : nat | P n /\ (forall i : nat, i < n -> ~ P i)} + {forall n : nat, ~ P n}
Hi:forall n0 : nat, (0 < INR n0 + 1)%R
P:nat -> Prop
HP:forall n0 : nat, P n0 \/ ~ P n0
E:=fun y : R => exists n0 : nat, P n0 /\ y = (/ (INR n0 + 1))%R \/ ~ P n0 /\ y = 0%R:R -> Prop
HE:forall n0 : nat, P n0 -> E (/ (INR n0 + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(l <= 0)%R
n:nat
Pn:P n

(0 < / (INR n + 1))%R
Hi:forall n0 : nat, (0 < INR n0 + 1)%R
P:nat -> Prop
HP:forall n0 : nat, P n0 \/ ~ P n0
E:=fun y : R => exists n0 : nat, P n0 /\ y = (/ (INR n0 + 1))%R \/ ~ P n0 /\ y = 0%R:R -> Prop
HE:forall n0 : nat, P n0 -> E (/ (INR n0 + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(l <= 0)%R
n:nat
Pn:P n
(/ (INR n + 1) <= l)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
{n : nat | P n /\ (forall i : nat, i < n -> ~ P i)} + {forall n : nat, ~ P n}
Hi:forall n0 : nat, (0 < INR n0 + 1)%R
P:nat -> Prop
HP:forall n0 : nat, P n0 \/ ~ P n0
E:=fun y : R => exists n0 : nat, P n0 /\ y = (/ (INR n0 + 1))%R \/ ~ P n0 /\ y = 0%R:R -> Prop
HE:forall n0 : nat, P n0 -> E (/ (INR n0 + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(l <= 0)%R
n:nat
Pn:P n

(/ (INR n + 1) <= l)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
{n : nat | P n /\ (forall i : nat, i < n -> ~ P i)} + {forall n : nat, ~ P n}
Hi:forall n0 : nat, (0 < INR n0 + 1)%R
P:nat -> Prop
HP:forall n0 : nat, P n0 \/ ~ P n0
E:=fun y : R => exists n0 : nat, P n0 /\ y = (/ (INR n0 + 1))%R \/ ~ P n0 /\ y = 0%R:R -> Prop
HE:forall n0 : nat, P n0 -> E (/ (INR n0 + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(l <= 0)%R
n:nat
Pn:P n

E (/ (INR n + 1))%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
{n : nat | P n /\ (forall i : nat, i < n -> ~ P i)} + {forall n : nat, ~ P n}
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R

{n : nat | P n /\ (forall i : nat, i < n -> ~ P i)} + {forall n : nat, ~ P n}
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R

{n : nat | P n /\ (forall i : nat, i < n -> ~ P i)}
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat

{n : nat | P n /\ (forall i : nat, i < n -> ~ P i)}
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat

P N /\ (forall i : nat, i < N -> ~ P i)
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat

(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
P N /\ (forall i : nat, i < N -> ~ P i)
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat

(INR (Z.abs_nat (up (/ l) - 2)) + 1)%R = (IZR (up (/ l)) - 1)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
P N /\ (forall i : nat, i < N -> ~ P i)
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat

(IZR (Z.of_nat (Z.abs_nat (up (/ l) - 2))) + 1)%R = (IZR (up (/ l)) - 1)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
P N /\ (forall i : nat, i < N -> ~ P i)
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat

(IZR (Z.abs (up (/ l) - 2)) + 1)%R = (IZR (up (/ l)) - 1)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
P N /\ (forall i : nat, i < N -> ~ P i)
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat

(IZR (Z.abs (up (/ l) - 2)) + 1)%R = (IZR (up (/ l) - 2) + 1)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
(IZR (up (/ l) - 2) + 1)%R = (IZR (up (/ l)) - 1)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
P N /\ (forall i : nat, i < N -> ~ P i)
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat

Z.abs (up (/ l) - 2) = (up (/ l) - 2)%Z
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
(IZR (up (/ l) - 2) + 1)%R = (IZR (up (/ l)) - 1)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
P N /\ (forall i : nat, i < N -> ~ P i)
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat

(0 <= up (/ l) - 2)%Z
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
(IZR (up (/ l) - 2) + 1)%R = (IZR (up (/ l)) - 1)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
P N /\ (forall i : nat, i < N -> ~ P i)
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat

(2 <= up (/ l))%Z
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
(IZR (up (/ l) - 2) + 1)%R = (IZR (up (/ l)) - 1)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
P N /\ (forall i : nat, i < N -> ~ P i)
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat

(1 < up (/ l))%Z
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
(IZR (up (/ l) - 2) + 1)%R = (IZR (up (/ l)) - 1)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
P N /\ (forall i : nat, i < N -> ~ P i)
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat

(1 < IZR (up (/ l)))%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
(IZR (up (/ l) - 2) + 1)%R = (IZR (up (/ l)) - 1)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
P N /\ (forall i : nat, i < N -> ~ P i)
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat

(1 <= / l)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
(/ l < IZR (up (/ l)))%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
(IZR (up (/ l) - 2) + 1)%R = (IZR (up (/ l)) - 1)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
P N /\ (forall i : nat, i < N -> ~ P i)
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat

(1 * l <= / l * l)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
(/ l < IZR (up (/ l)))%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
(IZR (up (/ l) - 2) + 1)%R = (IZR (up (/ l)) - 1)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
P N /\ (forall i : nat, i < N -> ~ P i)
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat

(l <= 1)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
(/ l < IZR (up (/ l)))%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
(IZR (up (/ l) - 2) + 1)%R = (IZR (up (/ l)) - 1)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
P N /\ (forall i : nat, i < N -> ~ P i)
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat

is_upper_bound E 1
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
(/ l < IZR (up (/ l)))%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
(IZR (up (/ l) - 2) + 1)%R = (IZR (up (/ l)) - 1)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
P N /\ (forall i : nat, i < N -> ~ P i)
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat

(/ l < IZR (up (/ l)))%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
(IZR (up (/ l) - 2) + 1)%R = (IZR (up (/ l)) - 1)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
P N /\ (forall i : nat, i < N -> ~ P i)
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat

(IZR (up (/ l) - 2) + 1)%R = (IZR (up (/ l)) - 1)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
P N /\ (forall i : nat, i < N -> ~ P i)
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat

(IZR (up (/ l)) - 2 + 1)%R = (IZR (up (/ l)) - 1)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
P N /\ (forall i : nat, i < N -> ~ P i)
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat

(IZR (up (/ l)) - 2 + 1)%R = (IZR (up (/ l)) - 1)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
P N /\ (forall i : nat, i < N -> ~ P i)
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R

P N /\ (forall i : nat, i < N -> ~ P i)
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R

forall i : nat, i < N -> ~ P i
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
P N /\ (forall i : nat, i < N -> ~ P i)
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
i:nat
HiN:i < N
Pi:P i

False
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
P N /\ (forall i : nat, i < N -> ~ P i)
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:forall x : R, E x -> (x <= l)%R
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
i:nat
HiN:i < N
Pi:P i

False
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
P N /\ (forall i : nat, i < N -> ~ P i)
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:forall x : R, E x -> (x <= l)%R
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
i:nat
HiN:i < N
Pi:P i

E (/ (INR i + 1))%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:forall x : R, E x -> (x <= l)%R
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
i:nat
HiN:i < N
Pi:P i
(l < / (INR i + 1))%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
P N /\ (forall i : nat, i < N -> ~ P i)
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:forall x : R, E x -> (x <= l)%R
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
i:nat
HiN:i < N
Pi:P i

(l < / (INR i + 1))%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
P N /\ (forall i : nat, i < N -> ~ P i)
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:forall x : R, E x -> (x <= l)%R
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
i:nat
HiN:i < N
Pi:P i

(/ / l < / (INR i + 1))%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
P N /\ (forall i : nat, i < N -> ~ P i)
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:forall x : R, E x -> (x <= l)%R
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
i:nat
HiN:i < N
Pi:P i

(1 <= INR i + 1)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:forall x : R, E x -> (x <= l)%R
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
i:nat
HiN:i < N
Pi:P i
(INR i + 1 < / l)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
P N /\ (forall i : nat, i < N -> ~ P i)
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:forall x : R, E x -> (x <= l)%R
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
i:nat
HiN:i < N
Pi:P i

(1 <= INR (S i))%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:forall x : R, E x -> (x <= l)%R
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
i:nat
HiN:i < N
Pi:P i
(INR i + 1 < / l)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
P N /\ (forall i : nat, i < N -> ~ P i)
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:forall x : R, E x -> (x <= l)%R
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
i:nat
HiN:i < N
Pi:P i

1 <= S i
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:forall x : R, E x -> (x <= l)%R
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
i:nat
HiN:i < N
Pi:P i
(INR i + 1 < / l)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
P N /\ (forall i : nat, i < N -> ~ P i)
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:forall x : R, E x -> (x <= l)%R
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
i:nat
HiN:i < N
Pi:P i

0 <= i
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:forall x : R, E x -> (x <= l)%R
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
i:nat
HiN:i < N
Pi:P i
(INR i + 1 < / l)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
P N /\ (forall i : nat, i < N -> ~ P i)
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:forall x : R, E x -> (x <= l)%R
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
i:nat
HiN:i < N
Pi:P i

(INR i + 1 < / l)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
P N /\ (forall i : nat, i < N -> ~ P i)
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:forall x : R, E x -> (x <= l)%R
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
i:nat
HiN:i < N
Pi:P i

(INR i + 1 < INR N + 1)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:forall x : R, E x -> (x <= l)%R
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
i:nat
HiN:i < N
Pi:P i
(INR N + 1 <= / l)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
P N /\ (forall i : nat, i < N -> ~ P i)
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:forall x : R, E x -> (x <= l)%R
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
i:nat
HiN:i < N
Pi:P i

(INR i < INR N)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:forall x : R, E x -> (x <= l)%R
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
i:nat
HiN:i < N
Pi:P i
(INR N + 1 <= / l)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
P N /\ (forall i : nat, i < N -> ~ P i)
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:forall x : R, E x -> (x <= l)%R
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
i:nat
HiN:i < N
Pi:P i

(INR N + 1 <= / l)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
P N /\ (forall i : nat, i < N -> ~ P i)
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:forall x : R, E x -> (x <= l)%R
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
i:nat
HiN:i < N
Pi:P i

(IZR (up (/ l)) - 1 <= / l)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
P N /\ (forall i : nat, i < N -> ~ P i)
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:forall x : R, E x -> (x <= l)%R
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
i:nat
HiN:i < N
Pi:P i

(IZR (up (/ l)) - 1 + (- / l + 1) <= / l + (- / l + 1))%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
P N /\ (forall i : nat, i < N -> ~ P i)
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:forall x : R, E x -> (x <= l)%R
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
i:nat
HiN:i < N
Pi:P i

(IZR (up (/ l)) - / l <= 1)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
P N /\ (forall i : nat, i < N -> ~ P i)
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i

P N /\ (forall i : nat, i < N -> ~ P i)
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:P N

P N /\ (forall i : nat, i < N -> ~ P i)
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N
P N /\ (forall i : nat, i < N -> ~ P i)
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N

P N /\ (forall i : nat, i < N -> ~ P i)
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N

False
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N

is_upper_bound E (/ (INR (S N) + 1))
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N
(/ (INR (S N) + 1) < l)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y0 : R => exists n : nat, P n /\ y0 = (/ (INR n + 1))%R \/ ~ P n /\ y0 = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N
y:nat
Py:P y

(/ (INR y + 1) <= / (INR (S N) + 1))%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y0 : R => exists n : nat, P n /\ y0 = (/ (INR n + 1))%R \/ ~ P n /\ y0 = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N
y:nat
(0 <= / (INR (S N) + 1))%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N
(/ (INR (S N) + 1) < l)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y0 : R => exists n : nat, P n /\ y0 = (/ (INR n + 1))%R \/ ~ P n /\ y0 = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N
y:nat
Py:P y
HyN:y = N

(/ (INR y + 1) <= / (INR (S N) + 1))%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y0 : R => exists n : nat, P n /\ y0 = (/ (INR n + 1))%R \/ ~ P n /\ y0 = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N
y:nat
Py:P y
HyN:y <> N
(/ (INR y + 1) <= / (INR (S N) + 1))%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y0 : R => exists n : nat, P n /\ y0 = (/ (INR n + 1))%R \/ ~ P n /\ y0 = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N
y:nat
(0 <= / (INR (S N) + 1))%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N
(/ (INR (S N) + 1) < l)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y0 : R => exists n : nat, P n /\ y0 = (/ (INR n + 1))%R \/ ~ P n /\ y0 = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N
y:nat
Py:P y
HyN:y = N

P N
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y0 : R => exists n : nat, P n /\ y0 = (/ (INR n + 1))%R \/ ~ P n /\ y0 = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N
y:nat
Py:P y
HyN:y <> N
(/ (INR y + 1) <= / (INR (S N) + 1))%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y0 : R => exists n : nat, P n /\ y0 = (/ (INR n + 1))%R \/ ~ P n /\ y0 = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N
y:nat
(0 <= / (INR (S N) + 1))%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N
(/ (INR (S N) + 1) < l)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y0 : R => exists n : nat, P n /\ y0 = (/ (INR n + 1))%R \/ ~ P n /\ y0 = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N
y:nat
Py:P y
HyN:y <> N

(/ (INR y + 1) <= / (INR (S N) + 1))%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y0 : R => exists n : nat, P n /\ y0 = (/ (INR n + 1))%R \/ ~ P n /\ y0 = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N
y:nat
(0 <= / (INR (S N) + 1))%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N
(/ (INR (S N) + 1) < l)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y0 : R => exists n : nat, P n /\ y0 = (/ (INR n + 1))%R \/ ~ P n /\ y0 = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N
y:nat
Py:P y
HyN:y <> N

(0 < INR (S N) + 1)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y0 : R => exists n : nat, P n /\ y0 = (/ (INR n + 1))%R \/ ~ P n /\ y0 = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N
y:nat
Py:P y
HyN:y <> N
(INR (S N) + 1 <= INR y + 1)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y0 : R => exists n : nat, P n /\ y0 = (/ (INR n + 1))%R \/ ~ P n /\ y0 = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N
y:nat
(0 <= / (INR (S N) + 1))%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N
(/ (INR (S N) + 1) < l)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y0 : R => exists n : nat, P n /\ y0 = (/ (INR n + 1))%R \/ ~ P n /\ y0 = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N
y:nat
Py:P y
HyN:y <> N

(INR (S N) + 1 <= INR y + 1)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y0 : R => exists n : nat, P n /\ y0 = (/ (INR n + 1))%R \/ ~ P n /\ y0 = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N
y:nat
(0 <= / (INR (S N) + 1))%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N
(/ (INR (S N) + 1) < l)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y0 : R => exists n : nat, P n /\ y0 = (/ (INR n + 1))%R \/ ~ P n /\ y0 = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N
y:nat
Py:P y
HyN:y <> N

(INR (S N) <= INR y)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y0 : R => exists n : nat, P n /\ y0 = (/ (INR n + 1))%R \/ ~ P n /\ y0 = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N
y:nat
(0 <= / (INR (S N) + 1))%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N
(/ (INR (S N) + 1) < l)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y0 : R => exists n : nat, P n /\ y0 = (/ (INR n + 1))%R \/ ~ P n /\ y0 = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N
y:nat
Py:P y
HyN:y <> N

~ (INR y < INR (S N))%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y0 : R => exists n : nat, P n /\ y0 = (/ (INR n + 1))%R \/ ~ P n /\ y0 = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N
y:nat
(0 <= / (INR (S N) + 1))%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N
(/ (INR (S N) + 1) < l)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y0 : R => exists n : nat, P n /\ y0 = (/ (INR n + 1))%R \/ ~ P n /\ y0 = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N
y:nat
Py:P y
HyN:y <> N
Hy:(INR y < INR (S N))%R

False
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y0 : R => exists n : nat, P n /\ y0 = (/ (INR n + 1))%R \/ ~ P n /\ y0 = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N
y:nat
(0 <= / (INR (S N) + 1))%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N
(/ (INR (S N) + 1) < l)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y0 : R => exists n : nat, P n /\ y0 = (/ (INR n + 1))%R \/ ~ P n /\ y0 = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N
y:nat
Py:P y
HyN:y <> N
Hy:(INR y < INR (S N))%R

y < N
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y0 : R => exists n : nat, P n /\ y0 = (/ (INR n + 1))%R \/ ~ P n /\ y0 = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N
y:nat
(0 <= / (INR (S N) + 1))%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N
(/ (INR (S N) + 1) < l)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y0 : R => exists n : nat, P n /\ y0 = (/ (INR n + 1))%R \/ ~ P n /\ y0 = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N
y:nat
Py:P y
HyN:y <> N
Hy:y < S N

y < N
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y0 : R => exists n : nat, P n /\ y0 = (/ (INR n + 1))%R \/ ~ P n /\ y0 = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N
y:nat
(0 <= / (INR (S N) + 1))%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N
(/ (INR (S N) + 1) < l)%R
l:R
N:=Z.abs_nat (up (/ l) - 2):nat
y:nat
HyN:y <> N
Hy:y < S N

y < N
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y0 : R => exists n : nat, P n /\ y0 = (/ (INR n + 1))%R \/ ~ P n /\ y0 = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N
y:nat
(0 <= / (INR (S N) + 1))%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N
(/ (INR (S N) + 1) < l)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y0 : R => exists n : nat, P n /\ y0 = (/ (INR n + 1))%R \/ ~ P n /\ y0 = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N
y:nat

(0 <= / (INR (S N) + 1))%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N
(/ (INR (S N) + 1) < l)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N

(/ (INR (S N) + 1) < l)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N

(/ (IZR (up (/ l)) - 1 + 1) < l)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N

(/ IZR (up (/ l)) < l)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N

(/ IZR (up (/ l)) < / / l)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N

(1 <= / l)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N
(/ l < IZR (up (/ l)))%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N

(/ 1 <= / l)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N
(/ l < IZR (up (/ l)))%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N

(0 < l)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N
(l <= 1)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N
(/ l < IZR (up (/ l)))%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N

(l <= 1)%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N
(/ l < IZR (up (/ l)))%R
Hi:forall n : nat, (0 < INR n + 1)%R
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
E:=fun y : R => exists n : nat, P n /\ y = (/ (INR n + 1))%R \/ ~ P n /\ y = 0%R:R -> Prop
HE:forall n : nat, P n -> E (/ (INR n + 1))%R
BE:is_upper_bound E 1
l:R
ub:is_upper_bound E l
lub:forall b : R, is_upper_bound E b -> (l <= b)%R
Hl:(0 < l)%R
N:=Z.abs_nat (up (/ l) - 2):nat
HN:(INR N + 1)%R = (IZR (up (/ l)) - 1)%R
H:forall i : nat, i < N -> ~ P i
PN:~ P N

(/ l < IZR (up (/ l)))%R
apply archimed. Qed.

forall P : nat -> Prop, (forall n : nat, P n \/ ~ P n) -> {n : nat | P n} + {forall n : nat, ~ P n}

forall P : nat -> Prop, (forall n : nat, P n \/ ~ P n) -> {n : nat | P n} + {forall n : nat, ~ P n}
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n

{n : nat | P n} + {forall n : nat, ~ P n}
P:nat -> Prop
HP:forall n0 : nat, P n0 \/ ~ P n0
n:nat
Pn:P n

{n0 : nat | P n0} + {forall n0 : nat, ~ P n0}
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
Pn:forall n : nat, ~ P n
{n : nat | P n} + {forall n : nat, ~ P n}
P:nat -> Prop
HP:forall n0 : nat, P n0 \/ ~ P n0
n:nat
Pn:P n

{n0 : nat | P n0}
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
Pn:forall n : nat, ~ P n
{n : nat | P n} + {forall n : nat, ~ P n}
P:nat -> Prop
HP:forall n : nat, P n \/ ~ P n
Pn:forall n : nat, ~ P n

{n : nat | P n} + {forall n : nat, ~ P n}
now right. Qed.

forall P : Z -> Prop, (forall n : Z, P n \/ ~ P n) -> {n : Z | P n} + {forall n : Z, ~ P n}

forall P : Z -> Prop, (forall n : Z, P n \/ ~ P n) -> {n : Z | P n} + {forall n : Z, ~ P n}
P:Z -> Prop
H:forall n : Z, P n \/ ~ P n

{n : Z | P n} + {forall n : Z, ~ P n}
P:Z -> Prop
H:forall n : Z, P n \/ ~ P n

forall n : nat, P (Z.of_nat n) \/ ~ P (Z.of_nat n)
P:Z -> Prop
H:forall n : Z, P n \/ ~ P n
J:{n : nat | P (Z.of_nat n)}
{n : Z | P n} + {forall n : Z, ~ P n}
P:Z -> Prop
H:forall n : Z, P n \/ ~ P n
J:forall n : nat, ~ P (Z.of_nat n)
{n : Z | P n} + {forall n : Z, ~ P n}
P:Z -> Prop
H:forall n : Z, P n \/ ~ P n
J:{n : nat | P (Z.of_nat n)}

{n : Z | P n} + {forall n : Z, ~ P n}
P:Z -> Prop
H:forall n : Z, P n \/ ~ P n
J:forall n : nat, ~ P (Z.of_nat n)
{n : Z | P n} + {forall n : Z, ~ P n}
P:Z -> Prop
H:forall n0 : Z, P n0 \/ ~ P n0
n:nat
Hn:P (Z.of_nat n)

{n0 : Z | P n0} + {forall n0 : Z, ~ P n0}
P:Z -> Prop
H:forall n : Z, P n \/ ~ P n
J:forall n : nat, ~ P (Z.of_nat n)
{n : Z | P n} + {forall n : Z, ~ P n}
P:Z -> Prop
H:forall n : Z, P n \/ ~ P n
J:forall n : nat, ~ P (Z.of_nat n)

{n : Z | P n} + {forall n : Z, ~ P n}
P:Z -> Prop
H:forall n : Z, P n \/ ~ P n
J:forall n : nat, ~ P (Z.of_nat n)

forall n : nat, P (- Z.of_nat n)%Z \/ ~ P (- Z.of_nat n)%Z
P:Z -> Prop
H:forall n : Z, P n \/ ~ P n
J:forall n : nat, ~ P (Z.of_nat n)
K:{n : nat | P (- Z.of_nat n)%Z}
{n : Z | P n} + {forall n : Z, ~ P n}
P:Z -> Prop
H:forall n : Z, P n \/ ~ P n
J:forall n : nat, ~ P (Z.of_nat n)
K:forall n : nat, ~ P (- Z.of_nat n)%Z
{n : Z | P n} + {forall n : Z, ~ P n}
P:Z -> Prop
H:forall n : Z, P n \/ ~ P n
J:forall n : nat, ~ P (Z.of_nat n)
K:{n : nat | P (- Z.of_nat n)%Z}

{n : Z | P n} + {forall n : Z, ~ P n}
P:Z -> Prop
H:forall n : Z, P n \/ ~ P n
J:forall n : nat, ~ P (Z.of_nat n)
K:forall n : nat, ~ P (- Z.of_nat n)%Z
{n : Z | P n} + {forall n : Z, ~ P n}
P:Z -> Prop
H:forall n0 : Z, P n0 \/ ~ P n0
J:forall n0 : nat, ~ P (Z.of_nat n0)
n:nat
Hn:P (- Z.of_nat n)%Z

{n0 : Z | P n0} + {forall n0 : Z, ~ P n0}
P:Z -> Prop
H:forall n : Z, P n \/ ~ P n
J:forall n : nat, ~ P (Z.of_nat n)
K:forall n : nat, ~ P (- Z.of_nat n)%Z
{n : Z | P n} + {forall n : Z, ~ P n}
P:Z -> Prop
H:forall n : Z, P n \/ ~ P n
J:forall n : nat, ~ P (Z.of_nat n)
K:forall n : nat, ~ P (- Z.of_nat n)%Z

{n : Z | P n} + {forall n : Z, ~ P n}
P:Z -> Prop
H:forall n0 : Z, P n0 \/ ~ P n0
J:forall n0 : nat, ~ P (Z.of_nat n0)
K:forall n0 : nat, ~ P (- Z.of_nat n0)%Z
n:Z
M:(0 <= n)%Z

~ P n
P:Z -> Prop
H:forall n0 : Z, P n0 \/ ~ P n0
J:forall n0 : nat, ~ P (Z.of_nat n0)
K:forall n0 : nat, ~ P (- Z.of_nat n0)%Z
n:Z
M:(n < 0)%Z
~ P n
P:Z -> Prop
H:forall n0 : Z, P n0 \/ ~ P n0
J:forall n0 : nat, ~ P (Z.of_nat n0)
K:forall n0 : nat, ~ P (- Z.of_nat n0)%Z
n:Z
M:(0 <= n)%Z

~ P (Z.abs n)
P:Z -> Prop
H:forall n0 : Z, P n0 \/ ~ P n0
J:forall n0 : nat, ~ P (Z.of_nat n0)
K:forall n0 : nat, ~ P (- Z.of_nat n0)%Z
n:Z
M:(n < 0)%Z
~ P n
P:Z -> Prop
H:forall n0 : Z, P n0 \/ ~ P n0
J:forall n0 : nat, ~ P (Z.of_nat n0)
K:forall n0 : nat, ~ P (- Z.of_nat n0)%Z
n:Z
M:(0 <= n)%Z

~ P (Z.of_nat (Z.abs_nat n))
P:Z -> Prop
H:forall n0 : Z, P n0 \/ ~ P n0
J:forall n0 : nat, ~ P (Z.of_nat n0)
K:forall n0 : nat, ~ P (- Z.of_nat n0)%Z
n:Z
M:(n < 0)%Z
~ P n
P:Z -> Prop
H:forall n0 : Z, P n0 \/ ~ P n0
J:forall n0 : nat, ~ P (Z.of_nat n0)
K:forall n0 : nat, ~ P (- Z.of_nat n0)%Z
n:Z
M:(n < 0)%Z

~ P n
P:Z -> Prop
H:forall n0 : Z, P n0 \/ ~ P n0
J:forall n0 : nat, ~ P (Z.of_nat n0)
K:forall n0 : nat, ~ P (- Z.of_nat n0)%Z
n:Z
M:(n < 0)%Z

~ P (- - n)%Z
P:Z -> Prop
H:forall n0 : Z, P n0 \/ ~ P n0
J:forall n0 : nat, ~ P (Z.of_nat n0)
K:forall n0 : nat, ~ P (- Z.of_nat n0)%Z
n:Z
M:(n < 0)%Z

~ P (- Z.abs n)%Z
P:Z -> Prop
H:forall n0 : Z, P n0 \/ ~ P n0
J:forall n0 : nat, ~ P (Z.of_nat n0)
K:forall n0 : nat, ~ P (- Z.of_nat n0)%Z
n:Z
M:(n < 0)%Z
(n <= 0)%Z
P:Z -> Prop
H:forall n0 : Z, P n0 \/ ~ P n0
J:forall n0 : nat, ~ P (Z.of_nat n0)
K:forall n0 : nat, ~ P (- Z.of_nat n0)%Z
n:Z
M:(n < 0)%Z

~ P (- Z.of_nat (Z.abs_nat n))%Z
P:Z -> Prop
H:forall n0 : Z, P n0 \/ ~ P n0
J:forall n0 : nat, ~ P (Z.of_nat n0)
K:forall n0 : nat, ~ P (- Z.of_nat n0)%Z
n:Z
M:(n < 0)%Z
(n <= 0)%Z
P:Z -> Prop
H:forall n0 : Z, P n0 \/ ~ P n0
J:forall n0 : nat, ~ P (Z.of_nat n0)
K:forall n0 : nat, ~ P (- Z.of_nat n0)%Z
n:Z
M:(n < 0)%Z

(n <= 0)%Z
omega. Qed.
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 [(bpow _ _ * bpow _ _)] =>
        rewrite <- bpow_plus
      | |- context [(?X1 * bpow _ _ * bpow _ _)] =>
        rewrite (Rmult_assoc X1); rewrite <- bpow_plus
      | |- context [(?X1 * (?X2 * bpow _ _) * 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 [(bpow _ ?X)] =>
        progress ring_simplify X
    end;
  (* bpow 0 ~~> 1 *)
  change (bpow _ 0) with 1;
  repeat
    match goal with
      | |- context [(_ * 1)] =>
        rewrite Rmult_1_r
    end.