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.
(************************************************************************)
(*         *   The Coq Proof Assistant / The Coq Development Team       *)
(*  v      *   INRIA, CNRS and contributors - Copyright 1999-2018       *)
(* <O___,, *       (see CREDITS file for the list of authors)           *)
(*   \VV/  **************************************************************)
(*    //   *    This file is distributed under the terms of the         *)
(*         *     GNU Lesser General Public License Version 2.1          *)
(*         *     (see LICENSE file for the text of the license)         *)
(************************************************************************)

Require Import Rbase.
Require Import Rfunctions.
Require Import Compare.
Local Open Scope R_scope.

Implicit Type r : R.

(* classical is needed for [Un_cv_crit] *)
(*********************************************************)

Definition of sequence and properties

(*                                                       *)
(*********************************************************)

Section sequence.

(*********)
  Variable Un : nat -> R.

(*********)
  Fixpoint Rmax_N (N:nat) : R :=
    match N with
      | O => Un 0
      | S n => Rmax (Un (S n)) (Rmax_N n)
    end.

(*********)
  Definition EUn r : Prop :=  exists i : nat, r = Un i.

(*********)
  Definition Un_cv (l:R) : Prop :=
    forall eps:R,
      eps > 0 ->
      exists N : nat, (forall n:nat, (n >= N)%nat -> R_dist (Un n) l < eps).

(*********)
  Definition Cauchy_crit : Prop :=
    forall eps:R,
      eps > 0 ->
      exists N : nat,
        (forall n m:nat,
          (n >= N)%nat -> (m >= N)%nat -> R_dist (Un n) (Un m) < eps).

(*********)
  Definition Un_growing : Prop := forall n:nat, Un n <= Un (S n).

(*********)
  
Un:nat -> R

exists r : R, EUn r
Un:nat -> R

exists r : R, EUn r
unfold EUn; split with (Un 0); split with 0%nat; trivial. Qed. (*********)
Un:nat -> R

forall n : nat, EUn (Un n)
Un:nat -> R

forall n : nat, EUn (Un n)
intro; unfold EUn; split with n; trivial. Qed. (*********)
Un:nat -> R

forall x : R, (forall n : nat, Un n <= x) -> is_upper_bound EUn x
Un:nat -> R

forall x : R, (forall n : nat, Un n <= x) -> is_upper_bound EUn x
intros; unfold is_upper_bound; intros; unfold EUn in H0; elim H0; clear H0; intros; generalize (H x1); intro; rewrite <- H0 in H1; trivial. Qed. (*********)
Un:nat -> R

forall n m : nat, Un_growing -> (n >= m)%nat -> Un n >= Un m
Un:nat -> R

forall n m : nat, Un_growing -> (n >= m)%nat -> Un n >= Un m
Un:nat -> R
n, m:nat
H:Un_growing
H0:(0 >= 0)%nat

Un 0 >= Un 0
Un:nat -> R
n, m, n0:nat
H:Un_growing -> (0 >= n0)%nat -> Un 0 >= Un n0
H0:Un_growing
H1:(0 >= S n0)%nat
Un 0 >= Un (S n0)
Un:nat -> R
n, m, n0:nat
H:forall m0 : nat, Un_growing -> (n0 >= m0)%nat -> Un n0 >= Un m0
H0:Un_growing
H1:(S n0 >= 0)%nat
Un (S n0) >= Un 0
Un:nat -> R
n, m, n0:nat
H:forall n2 : nat, (forall m0 : nat, Un_growing -> (n2 >= m0)%nat -> Un n2 >= Un m0) -> Un_growing -> (S n2 >= n0)%nat -> Un (S n2) >= Un n0
n1:nat
H0:forall m0 : nat, Un_growing -> (n1 >= m0)%nat -> Un n1 >= Un m0
H1:Un_growing
H2:(S n1 >= S n0)%nat
Un (S n1) >= Un (S n0)
Un:nat -> R
n, m, n0:nat
H:Un_growing -> (0 >= n0)%nat -> Un 0 >= Un n0
H0:Un_growing
H1:(0 >= S n0)%nat

Un 0 >= Un (S n0)
Un:nat -> R
n, m, n0:nat
H:forall m0 : nat, Un_growing -> (n0 >= m0)%nat -> Un n0 >= Un m0
H0:Un_growing
H1:(S n0 >= 0)%nat
Un (S n0) >= Un 0
Un:nat -> R
n, m, n0:nat
H:forall n2 : nat, (forall m0 : nat, Un_growing -> (n2 >= m0)%nat -> Un n2 >= Un m0) -> Un_growing -> (S n2 >= n0)%nat -> Un (S n2) >= Un n0
n1:nat
H0:forall m0 : nat, Un_growing -> (n1 >= m0)%nat -> Un n1 >= Un m0
H1:Un_growing
H2:(S n1 >= S n0)%nat
Un (S n1) >= Un (S n0)
Un:nat -> R
n, m, n0:nat
H:forall m0 : nat, Un_growing -> (n0 >= m0)%nat -> Un n0 >= Un m0
H0:Un_growing
H1:(S n0 >= 0)%nat

Un (S n0) >= Un 0
Un:nat -> R
n, m, n0:nat
H:forall n2 : nat, (forall m0 : nat, Un_growing -> (n2 >= m0)%nat -> Un n2 >= Un m0) -> Un_growing -> (S n2 >= n0)%nat -> Un (S n2) >= Un n0
n1:nat
H0:forall m0 : nat, Un_growing -> (n1 >= m0)%nat -> Un n1 >= Un m0
H1:Un_growing
H2:(S n1 >= S n0)%nat
Un (S n1) >= Un (S n0)
Un:nat -> R
n, m, n0:nat
H:forall m0 : nat, Un_growing -> (n0 >= m0)%nat -> Un n0 >= Un m0
H0:Un_growing
H1:(S n0 >= 0)%nat

(n0 >= 0)%nat -> Un (S n0) >= Un 0
Un:nat -> R
n, m, n0:nat
H:forall m0 : nat, Un_growing -> (n0 >= m0)%nat -> Un n0 >= Un m0
H0:Un_growing
H1:(S n0 >= 0)%nat
(n0 >= 0)%nat
Un:nat -> R
n, m, n0:nat
H:forall n2 : nat, (forall m0 : nat, Un_growing -> (n2 >= m0)%nat -> Un n2 >= Un m0) -> Un_growing -> (S n2 >= n0)%nat -> Un (S n2) >= Un n0
n1:nat
H0:forall m0 : nat, Un_growing -> (n1 >= m0)%nat -> Un n1 >= Un m0
H1:Un_growing
H2:(S n1 >= S n0)%nat
Un (S n1) >= Un (S n0)
Un:nat -> R
n, m, n0:nat
H:forall m0 : nat, Un_growing -> (n0 >= m0)%nat -> Un n0 >= Un m0
H0:Un_growing
H1:(S n0 >= 0)%nat

(n0 >= 0)%nat
Un:nat -> R
n, m, n0:nat
H:forall n2 : nat, (forall m0 : nat, Un_growing -> (n2 >= m0)%nat -> Un n2 >= Un m0) -> Un_growing -> (S n2 >= n0)%nat -> Un (S n2) >= Un n0
n1:nat
H0:forall m0 : nat, Un_growing -> (n1 >= m0)%nat -> Un n1 >= Un m0
H1:Un_growing
H2:(S n1 >= S n0)%nat
Un (S n1) >= Un (S n0)
Un:nat -> R
n, m, n0:nat
H:forall n2 : nat, (forall m0 : nat, Un_growing -> (n2 >= m0)%nat -> Un n2 >= Un m0) -> Un_growing -> (S n2 >= n0)%nat -> Un (S n2) >= Un n0
n1:nat
H0:forall m0 : nat, Un_growing -> (n1 >= m0)%nat -> Un n1 >= Un m0
H1:Un_growing
H2:(S n1 >= S n0)%nat

Un (S n1) >= Un (S n0)
Un:nat -> R
n, m, n0:nat
H:forall n2 : nat, (forall m0 : nat, Un_growing -> (n2 >= m0)%nat -> Un n2 >= Un m0) -> Un_growing -> (S n2 >= n0)%nat -> Un (S n2) >= Un n0
n1:nat
H0:forall m0 : nat, Un_growing -> (n1 >= m0)%nat -> Un n1 >= Un m0
H1:Un_growing
H2:(S n1 >= S n0)%nat
y:{(n1 < n0)%nat} + {n1 = n0}

Un (S n1) >= Un (S n0)
Un:nat -> R
n, m, n0:nat
H:forall n2 : nat, (forall m0 : nat, Un_growing -> (n2 >= m0)%nat -> Un n2 >= Un m0) -> Un_growing -> (S n2 >= n0)%nat -> Un (S n2) >= Un n0
n1:nat
H0:forall m0 : nat, Un_growing -> (n1 >= m0)%nat -> Un n1 >= Un m0
H1:Un_growing
H2:(S n1 >= S n0)%nat
y:(n0 < n1)%nat
Un (S n1) >= Un (S n0)
Un:nat -> R
n, m, n0:nat
H:forall n2 : nat, (forall m0 : nat, Un_growing -> (n2 >= m0)%nat -> Un n2 >= Un m0) -> Un_growing -> (S n2 >= n0)%nat -> Un (S n2) >= Un n0
n1:nat
H0:forall m0 : nat, Un_growing -> (n1 >= m0)%nat -> Un n1 >= Un m0
H1:Un_growing
H2:(S n1 >= S n0)%nat
y:(n1 < n0)%nat

Un (S n1) >= Un (S n0)
Un:nat -> R
n, m, n0:nat
H:forall n2 : nat, (forall m0 : nat, Un_growing -> (n2 >= m0)%nat -> Un n2 >= Un m0) -> Un_growing -> (S n2 >= n0)%nat -> Un (S n2) >= Un n0
n1:nat
H0:forall m0 : nat, Un_growing -> (n1 >= m0)%nat -> Un n1 >= Un m0
H1:Un_growing
H2:(S n1 >= S n0)%nat
y:n1 = n0
Un (S n1) >= Un (S n0)
Un:nat -> R
n, m, n0:nat
H:forall n2 : nat, (forall m0 : nat, Un_growing -> (n2 >= m0)%nat -> Un n2 >= Un m0) -> Un_growing -> (S n2 >= n0)%nat -> Un (S n2) >= Un n0
n1:nat
H0:forall m0 : nat, Un_growing -> (n1 >= m0)%nat -> Un n1 >= Un m0
H1:Un_growing
H2:(S n1 >= S n0)%nat
y:(n0 < n1)%nat
Un (S n1) >= Un (S n0)
Un:nat -> R
n, m, n0:nat
H:forall n2 : nat, (forall m0 : nat, Un_growing -> (n2 >= m0)%nat -> Un n2 >= Un m0) -> Un_growing -> (S n2 >= n0)%nat -> Un (S n2) >= Un n0
n1:nat
H0:forall m0 : nat, Un_growing -> (n1 >= m0)%nat -> Un n1 >= Un m0
H1:Un_growing
H2:(S n1 >= S n0)%nat
y:n1 = n0

Un (S n1) >= Un (S n0)
Un:nat -> R
n, m, n0:nat
H:forall n2 : nat, (forall m0 : nat, Un_growing -> (n2 >= m0)%nat -> Un n2 >= Un m0) -> Un_growing -> (S n2 >= n0)%nat -> Un (S n2) >= Un n0
n1:nat
H0:forall m0 : nat, Un_growing -> (n1 >= m0)%nat -> Un n1 >= Un m0
H1:Un_growing
H2:(S n1 >= S n0)%nat
y:(n0 < n1)%nat
Un (S n1) >= Un (S n0)
Un:nat -> R
n, m, n0:nat
H:forall n2 : nat, (forall m0 : nat, Un_growing -> (n2 >= m0)%nat -> Un n2 >= Un m0) -> Un_growing -> (S n2 >= n0)%nat -> Un (S n2) >= Un n0
n1:nat
H0:forall m0 : nat, Un_growing -> (n1 >= m0)%nat -> Un n1 >= Un m0
H1:Un_growing
H2:(S n1 >= S n0)%nat
y:(n0 < n1)%nat

Un (S n1) >= Un (S n0)
unfold ge in H0; generalize (H0 (S n0) H1 (lt_le_S n0 n1 y)); intro; unfold Un_growing in H1; apply (Rge_trans (Un (S n1)) (Un n1) (Un (S n0)) (Rle_ge (Un n1) (Un (S n1)) (H1 n1)) H3). Qed. (*********)
Un:nat -> R

Un_growing -> forall l : R, is_lub EUn l -> Un_cv l
Un:nat -> R

Un_growing -> forall l : R, is_lub EUn l -> Un_cv l
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0

exists N : nat, forall n : nat, (n >= N)%nat -> R_dist (Un n) l < eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0

(exists N : nat, Un N > l - eps) -> exists N : nat, forall n : nat, (n >= N)%nat -> R_dist (Un n) l < eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
N:nat
H3:Un N > l - eps

exists N0 : nat, forall n : nat, (n >= N0)%nat -> R_dist (Un n) l < eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
N:nat
H3:Un N > l - eps

forall n : nat, (n >= N)%nat -> R_dist (Un n) l < eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
N:nat
H3:Un N > l - eps
n:nat
H4:(n >= N)%nat

R_dist (Un n) l < eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
N:nat
H3:Un N > l - eps
n:nat
H4:(n >= N)%nat

Rabs (Un n - l) < eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
N:nat
H3:Un N > l - eps
n:nat
H4:(n >= N)%nat

l - Un n < eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
N:nat
H3:Un N > l - eps
n:nat
H4:(n >= N)%nat
Un n - l <= 0
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
N:nat
H3:Un N > l - eps
n:nat
H4:(n >= N)%nat

Un n - eps + (l - Un n) < Un n - eps + eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
N:nat
H3:Un N > l - eps
n:nat
H4:(n >= N)%nat
Un n - l <= 0
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
N:nat
H3:Un N > l - eps
n:nat
H4:(n >= N)%nat

Un n - eps + (l - Un n) < Un N
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
N:nat
H3:Un N > l - eps
n:nat
H4:(n >= N)%nat
Un N <= Un n - eps + eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
N:nat
H3:Un N > l - eps
n:nat
H4:(n >= N)%nat
Un n - l <= 0
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
N:nat
H3:Un N > l - eps
n:nat
H4:(n >= N)%nat

Un N <= Un n - eps + eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
N:nat
H3:Un N > l - eps
n:nat
H4:(n >= N)%nat
Un n - l <= 0
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
N:nat
H3:Un N > l - eps
n:nat
H4:(n >= N)%nat

Un N <= Un n
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
N:nat
H3:Un N > l - eps
n:nat
H4:(n >= N)%nat
Un n - l <= 0
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
N:nat
H3:Un N > l - eps
n:nat
H4:(n >= N)%nat

Un n >= Un N
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
N:nat
H3:Un N > l - eps
n:nat
H4:(n >= N)%nat
Un n - l <= 0
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
N:nat
H3:Un N > l - eps
n:nat
H4:(n >= N)%nat

Un n - l <= 0
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
N:nat
H3:Un N > l - eps
n:nat
H4:(n >= N)%nat

Un n <= l
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
N:nat
H3:Un N > l - eps
n:nat
H4:(n >= N)%nat

EUn (Un n)
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0

exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0

forall n : nat, 0 < (/ 2) ^ n
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
exists N : nat, Un N > l - eps

forall n : nat, 0 < (/ 2) ^ n
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
exists N : nat, Un N > l - eps
n:nat

0 < (/ 2) ^ n
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
exists N : nat, Un N > l - eps
n:nat

0 < / 2
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
exists N : nat, Un N > l - eps
n:nat

0 < 2
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n

exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool

exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R

exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R

forall m n : nat, sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m n : nat, sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R

forall m n : nat, sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m n : nat, sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
exists N : nat, Un N > l - eps
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R

forall m n : nat, sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m n : nat, sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
exists N : nat, Un N > l - eps
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
m:nat

forall n : nat, sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m n : nat, sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
exists N : nat, Un N > l - eps
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
m:nat

sum m <= sum (m + 0)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + 0)
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
m, n:nat
IHn:sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
sum m <= sum (m + S n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + S n)
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m n : nat, sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
exists N : nat, Un N > l - eps
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
m:nat

sum m <= sum m <= sum m + (/ 2) ^ m - (/ 2) ^ m
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
m, n:nat
IHn:sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
sum m <= sum (m + S n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + S n)
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m n : nat, sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
exists N : nat, Un N > l - eps
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
m:nat

sum m <= sum m <= sum m
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
m, n:nat
IHn:sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
sum m <= sum (m + S n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + S n)
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m n : nat, sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
exists N : nat, Un N > l - eps
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
m, n:nat
IHn:sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)

sum m <= sum (m + S n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + S n)
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m n : nat, sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
exists N : nat, Un N > l - eps
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
m, n:nat
IHn:sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)

sum m <= sum (S (m + n)) <= sum m + (/ 2) ^ m - (/ 2) ^ S (m + n)
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m n : nat, sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
exists N : nat, Un N > l - eps
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
m, n:nat
IHn:sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)

sum m <= sum (m + n)%nat + (if test (m + n)%nat then / 2 * (/ 2) ^ (m + n) else 0) <= sum m + (/ 2) ^ m - / 2 * (/ 2) ^ (m + n)
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m n : nat, sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
exists N : nat, Un N > l - eps
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
m, n:nat
IHn:sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)

sum m <= sum (m + n)%nat + (if test (m + n)%nat then / 2 * (/ 2) ^ (m + n) else 0)
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
m, n:nat
IHn:sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
sum (m + n)%nat + (if test (m + n)%nat then / 2 * (/ 2) ^ (m + n) else 0) <= sum m + (/ 2) ^ m - / 2 * (/ 2) ^ (m + n)
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m n : nat, sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
exists N : nat, Un N > l - eps
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
m, n:nat
IHn:sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)

sum m <= sum (m + n)%nat + 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
m, n:nat
IHn:sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
sum (m + n)%nat + 0 <= sum (m + n)%nat + (if test (m + n)%nat then / 2 * (/ 2) ^ (m + n) else 0)
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
m, n:nat
IHn:sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
sum (m + n)%nat + (if test (m + n)%nat then / 2 * (/ 2) ^ (m + n) else 0) <= sum m + (/ 2) ^ m - / 2 * (/ 2) ^ (m + n)
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m n : nat, sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
exists N : nat, Un N > l - eps
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
m, n:nat
IHn:sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)

sum m <= sum (m + n)%nat
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
m, n:nat
IHn:sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
sum (m + n)%nat + 0 <= sum (m + n)%nat + (if test (m + n)%nat then / 2 * (/ 2) ^ (m + n) else 0)
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
m, n:nat
IHn:sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
sum (m + n)%nat + (if test (m + n)%nat then / 2 * (/ 2) ^ (m + n) else 0) <= sum m + (/ 2) ^ m - / 2 * (/ 2) ^ (m + n)
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m n : nat, sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
exists N : nat, Un N > l - eps
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
m, n:nat
IHn:sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)

sum (m + n)%nat + 0 <= sum (m + n)%nat + (if test (m + n)%nat then / 2 * (/ 2) ^ (m + n) else 0)
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
m, n:nat
IHn:sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
sum (m + n)%nat + (if test (m + n)%nat then / 2 * (/ 2) ^ (m + n) else 0) <= sum m + (/ 2) ^ m - / 2 * (/ 2) ^ (m + n)
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m n : nat, sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
exists N : nat, Un N > l - eps
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
m, n:nat
IHn:sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)

0 <= (if test (m + n)%nat then / 2 * (/ 2) ^ (m + n) else 0)
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
m, n:nat
IHn:sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
sum (m + n)%nat + (if test (m + n)%nat then / 2 * (/ 2) ^ (m + n) else 0) <= sum m + (/ 2) ^ m - / 2 * (/ 2) ^ (m + n)
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m n : nat, sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
exists N : nat, Un N > l - eps
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
m, n:nat
IHn:sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)

0 <= / 2 * (/ 2) ^ (m + n)
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
m, n:nat
IHn:sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
0 <= 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
m, n:nat
IHn:sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
sum (m + n)%nat + (if test (m + n)%nat then / 2 * (/ 2) ^ (m + n) else 0) <= sum m + (/ 2) ^ m - / 2 * (/ 2) ^ (m + n)
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m n : nat, sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
exists N : nat, Un N > l - eps
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
m, n:nat
IHn:sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)

0 < / 2 * (/ 2) ^ (m + n)
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
m, n:nat
IHn:sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
0 <= 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
m, n:nat
IHn:sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
sum (m + n)%nat + (if test (m + n)%nat then / 2 * (/ 2) ^ (m + n) else 0) <= sum m + (/ 2) ^ m - / 2 * (/ 2) ^ (m + n)
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m n : nat, sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
exists N : nat, Un N > l - eps
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
m, n:nat
IHn:sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)

0 <= 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
m, n:nat
IHn:sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
sum (m + n)%nat + (if test (m + n)%nat then / 2 * (/ 2) ^ (m + n) else 0) <= sum m + (/ 2) ^ m - / 2 * (/ 2) ^ (m + n)
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m n : nat, sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
exists N : nat, Un N > l - eps
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
m, n:nat
IHn:sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)

sum (m + n)%nat + (if test (m + n)%nat then / 2 * (/ 2) ^ (m + n) else 0) <= sum m + (/ 2) ^ m - / 2 * (/ 2) ^ (m + n)
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m n : nat, sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
exists N : nat, Un N > l - eps
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
m, n:nat
IHn:sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)

sum (m + n)%nat + (if test (m + n)%nat then / 2 * (/ 2) ^ (m + n) else 0) <= sum (m + n)%nat + / 2 * (/ 2) ^ (m + n)
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
m, n:nat
IHn:sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
sum (m + n)%nat + / 2 * (/ 2) ^ (m + n) <= sum m + (/ 2) ^ m - / 2 * (/ 2) ^ (m + n)
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m n : nat, sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
exists N : nat, Un N > l - eps
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
m, n:nat
IHn:sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)

(if test (m + n)%nat then / 2 * (/ 2) ^ (m + n) else 0) <= / 2 * (/ 2) ^ (m + n)
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
m, n:nat
IHn:sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
sum (m + n)%nat + / 2 * (/ 2) ^ (m + n) <= sum m + (/ 2) ^ m - / 2 * (/ 2) ^ (m + n)
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m n : nat, sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
exists N : nat, Un N > l - eps
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
m, n:nat
IHn:sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)

/ 2 * (/ 2) ^ (m + n) <= / 2 * (/ 2) ^ (m + n)
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
m, n:nat
IHn:sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
0 <= / 2 * (/ 2) ^ (m + n)
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
m, n:nat
IHn:sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
sum (m + n)%nat + / 2 * (/ 2) ^ (m + n) <= sum m + (/ 2) ^ m - / 2 * (/ 2) ^ (m + n)
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m n : nat, sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
exists N : nat, Un N > l - eps
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
m, n:nat
IHn:sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)

0 <= / 2 * (/ 2) ^ (m + n)
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
m, n:nat
IHn:sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
sum (m + n)%nat + / 2 * (/ 2) ^ (m + n) <= sum m + (/ 2) ^ m - / 2 * (/ 2) ^ (m + n)
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m n : nat, sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
exists N : nat, Un N > l - eps
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
m, n:nat
IHn:sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)

0 < / 2 * (/ 2) ^ (m + n)
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
m, n:nat
IHn:sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
sum (m + n)%nat + / 2 * (/ 2) ^ (m + n) <= sum m + (/ 2) ^ m - / 2 * (/ 2) ^ (m + n)
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m n : nat, sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
exists N : nat, Un N > l - eps
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
m, n:nat
IHn:sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)

sum (m + n)%nat + / 2 * (/ 2) ^ (m + n) <= sum m + (/ 2) ^ m - / 2 * (/ 2) ^ (m + n)
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m n : nat, sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
exists N : nat, Un N > l - eps
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
m, n:nat
IHn:sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)

sum (m + n)%nat + / 2 * (/ 2) ^ (m + n) + - (/ 2 * (/ 2) ^ (m + n)) <= sum m + (/ 2) ^ m - / 2 * (/ 2) ^ (m + n) + - (/ 2 * (/ 2) ^ (m + n))
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m n : nat, sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
exists N : nat, Un N > l - eps
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
m, n:nat
IHn:sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)

sum (m + n)%nat <= sum m + (/ 2) ^ m - / 2 * (/ 2) ^ (m + n) + - (/ 2 * (/ 2) ^ (m + n))
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m n : nat, sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
exists N : nat, Un N > l - eps
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
m, n:nat
IHn:sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)

sum m + (/ 2) ^ m - (/ 2) ^ (m + n) <= sum m + (/ 2) ^ m - / 2 * (/ 2) ^ (m + n) + - (/ 2 * (/ 2) ^ (m + n))
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m n : nat, sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
exists N : nat, Un N > l - eps
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
m, n:nat
IHn:sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)

sum m + (/ 2) ^ m - (/ 2) ^ (m + n) = sum m + (/ 2) ^ m - / 2 * (/ 2) ^ (m + n) + - (/ 2 * (/ 2) ^ (m + n))
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m n : nat, sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m n : nat, sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)

exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m n : nat, sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)

forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m n : nat, sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m n : nat, sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
N:nat

0 <= sum N <= 1 - (/ 2) ^ N
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m n : nat, sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m n : nat, sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
N:nat

sum 0%nat <= sum (0 + N)%nat <= sum 0%nat + (/ 2) ^ 0 - (/ 2) ^ (0 + N) -> 0 <= sum N <= 1 - (/ 2) ^ N
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m n : nat, sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m n : nat, sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
N:nat

0 <= sum N <= 0 + 1 - (/ 2) ^ N -> 0 <= sum N <= 1 - (/ 2) ^ N
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m n : nat, sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m n : nat, sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n

exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m n : nat, sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n

bound (fun x : R => exists n : nat, x = sum n)
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m n : nat, sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
exists (x : R) (n : nat), x = sum n
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m n : nat, sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n

is_upper_bound (fun x : R => exists n : nat, x = sum n) 1
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m n : nat, sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
exists (x : R) (n : nat), x = sum n
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m n0 : nat, sum m <= sum (m + n0)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
x:R
n:nat
H1:x = sum n

x <= 1
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m n : nat, sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
exists (x : R) (n : nat), x = sum n
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m n0 : nat, sum m <= sum (m + n0)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
x:R
n:nat
H1:x = sum n

sum n <= 1
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m n : nat, sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
exists (x : R) (n : nat), x = sum n
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m n0 : nat, sum m <= sum (m + n0)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
x:R
n:nat
H1:x = sum n

1 - (/ 2) ^ n <= 1
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m n : nat, sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
exists (x : R) (n : nat), x = sum n
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m n0 : nat, sum m <= sum (m + n0)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
x:R
n:nat
H1:x = sum n

1 - (/ 2) ^ n < 1
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m n : nat, sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
exists (x : R) (n : nat), x = sum n
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m n0 : nat, sum m <= sum (m + n0)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
x:R
n:nat
H1:x = sum n

(/ 2) ^ n - 1 + (1 - (/ 2) ^ n) < (/ 2) ^ n - 1 + 1
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m n : nat, sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
exists (x : R) (n : nat), x = sum n
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m n : nat, sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n

exists (x : R) (n : nat), x = sum n
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m n : nat, sum m <= sum (m + n)%nat <= sum m + (/ 2) ^ m - (/ 2) ^ (m + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n

exists n : nat, 0 = sum n
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b

exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:m < 0

exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:m = 0
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:m < 0

0 <= m
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:m = 0
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:m < 0

exists n : nat, 0 = sum n
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:m = 0
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:m = 0

exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:m = 0

forall n : nat, sum n = 0
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:m = 0
Hs0:forall n : nat, sum n = 0
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:m = 0
n:nat

sum n = 0
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:m = 0
Hs0:forall n : nat, sum n = 0
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
n:nat
Hm1:sum n <= m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:m = 0

sum n = 0
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:m = 0
Hs0:forall n : nat, sum n = 0
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
n:nat
Hm1:sum n <= m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:m = 0

sum n <= 0
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:m = 0
Hs0:forall n : nat, sum n = 0
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:m = 0
Hs0:forall n : nat, sum n = 0

exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:m = 0
Hs0:forall n : nat, sum n = 0

forall n : nat, Un n <= l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:m = 0
Hs0:forall n : nat, sum n = 0
Hub:forall n : nat, Un n <= l - eps
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:m = 0
Hs0:forall n0 : nat, sum n0 = 0
n:nat

Un n <= l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:m = 0
Hs0:forall n : nat, sum n = 0
Hub:forall n : nat, Un n <= l - eps
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:m = 0
Hs0:forall n0 : nat, sum n0 = 0
n:nat

sum (S n) = sum (S n) -> Un n <= l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:m = 0
Hs0:forall n : nat, sum n = 0
Hub:forall n : nat, Un n <= l - eps
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:m = 0
Hs0:forall n0 : nat, sum n0 = 0
n:nat

sum n + (if test n then / 2 * (/ 2) ^ n else 0) = sum (S n) -> Un n <= l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:m = 0
Hs0:forall n : nat, sum n = 0
Hub:forall n : nat, Un n <= l - eps
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:m = 0
Hs0:forall n0 : nat, sum n0 = 0
n:nat

(if test n then / 2 * (/ 2) ^ n else 0) = 0 -> Un n <= l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:m = 0
Hs0:forall n : nat, sum n = 0
Hub:forall n : nat, Un n <= l - eps
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:m = 0
Hs0:forall n0 : nat, sum n0 = 0
n:nat

(if if Rle_lt_dec (Un n) (l - eps) then false else true then / 2 * (/ 2) ^ n else 0) = 0 -> Un n <= l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:m = 0
Hs0:forall n : nat, sum n = 0
Hub:forall n : nat, Un n <= l - eps
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:m = 0
Hs0:forall n0 : nat, sum n0 = 0
n:nat
r:Un n <= l - eps

0 = 0 -> Un n <= l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:m = 0
Hs0:forall n0 : nat, sum n0 = 0
n:nat
r:l - eps < Un n
/ 2 * (/ 2) ^ n = 0 -> Un n <= l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:m = 0
Hs0:forall n : nat, sum n = 0
Hub:forall n : nat, Un n <= l - eps
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:m = 0
Hs0:forall n0 : nat, sum n0 = 0
n:nat
r:l - eps < Un n

/ 2 * (/ 2) ^ n = 0 -> Un n <= l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:m = 0
Hs0:forall n : nat, sum n = 0
Hub:forall n : nat, Un n <= l - eps
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:m = 0
Hs0:forall n0 : nat, sum n0 = 0
n:nat
r:l - eps < Un n
H':/ 2 * (/ 2) ^ n = 0

Un n <= l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:m = 0
Hs0:forall n : nat, sum n = 0
Hub:forall n : nat, Un n <= l - eps
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:m = 0
Hs0:forall n0 : nat, sum n0 = 0
n:nat
r:l - eps < Un n
H':/ 2 * (/ 2) ^ n = 0

/ 2 * (/ 2) ^ n > 0
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:m = 0
Hs0:forall n : nat, sum n = 0
Hub:forall n : nat, Un n <= l - eps
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:m = 0
Hs0:forall n : nat, sum n = 0
Hub:forall n : nat, Un n <= l - eps

exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m
exists N : nat, Un N > l - eps
Un:nat -> R
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hub:forall n : nat, Un n <= l - eps

exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m
exists N : nat, Un N > l - eps
Un:nat -> R
l:R
H:forall b : R, is_upper_bound EUn b -> l <= b
eps:R
Heps:eps > 0
Hub:forall n : nat, Un n <= l - eps

exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m
exists N : nat, Un N > l - eps
Un:nat -> R
l:R
H:forall b : R, is_upper_bound EUn b -> l <= b
eps:R
Heps:eps > 0
Hub:forall n : nat, Un n <= l - eps

is_upper_bound EUn (l - eps)
Un:nat -> R
l:R
H:forall b : R, is_upper_bound EUn b -> l <= b
eps:R
Heps:eps > 0
Hub:forall n : nat, Un n <= l - eps
l - eps < l
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m
exists N : nat, Un N > l - eps
Un:nat -> R
l:R
H:forall b : R, is_upper_bound EUn b -> l <= b
eps:R
Heps:eps > 0
Hub:forall n0 : nat, Un n0 <= l - eps
x:R
n:nat
H1:x = Un n

x <= l - eps
Un:nat -> R
l:R
H:forall b : R, is_upper_bound EUn b -> l <= b
eps:R
Heps:eps > 0
Hub:forall n : nat, Un n <= l - eps
l - eps < l
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m
exists N : nat, Un N > l - eps
Un:nat -> R
l:R
H:forall b : R, is_upper_bound EUn b -> l <= b
eps:R
Heps:eps > 0
Hub:forall n : nat, Un n <= l - eps

l - eps < l
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m
exists N : nat, Un N > l - eps
Un:nat -> R
l:R
H:forall b : R, is_upper_bound EUn b -> l <= b
eps:R
Heps:eps > 0
Hub:forall n : nat, Un n <= l - eps

eps - l + (l - eps) < eps - l + l
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m

exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m

Rabs (/ 2) < 1
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m

/ 2 < 1
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m
0 <= / 2
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m

/ 2 < / 1
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m
0 <= / 2
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m

0 < 1 * 2
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m
1 < 2
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m
0 <= / 2
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m

0 < 2
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m
1 < 2
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m
0 <= / 2
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m

1 < 2
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m
0 <= / 2
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m

0 <= / 2
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m

0 < / 2
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m

0 < 2
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1

exists N : nat, Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H4:forall n : nat, (n >= N)%nat -> Rabs ((/ 2) ^ n) < m

exists N0 : nat, Un N0 > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H4:forall n : nat, (n >= N)%nat -> Rabs ((/ 2) ^ n) < m

Un N > l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H4:forall n : nat, (n >= N)%nat -> Rabs ((/ 2) ^ n) < m

~ Un N <= l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H4:forall n : nat, (n >= N)%nat -> Rabs ((/ 2) ^ n) < m
H5:Un N <= l - eps

False
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H4:forall n : nat, (n >= N)%nat -> Rabs ((/ 2) ^ n) < m
H5:Un N <= l - eps

m <= Rabs ((/ 2) ^ N)
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H4:forall n : nat, (n >= N)%nat -> Rabs ((/ 2) ^ n) < m
H5:Un N <= l - eps

m <= (/ 2) ^ N
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H4:forall n : nat, (n >= N)%nat -> Rabs ((/ 2) ^ n) < m
H5:Un N <= l - eps
0 <= (/ 2) ^ N
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H4:forall n : nat, (n >= N)%nat -> Rabs ((/ 2) ^ n) < m
H5:Un N <= l - eps

m <= (/ 2) ^ N
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n : nat, 0 < (/ 2) ^ n
test:=fun n : nat => if Rle_lt_dec (Un n) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n : nat) : R := match n with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n else 0) end in aux:nat -> R
Hsum':forall m0 n : nat, sum m0 <= sum (m0 + n)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n)
Hsum:forall n : nat, 0 <= sum n <= 1 - (/ 2) ^ n
m:R
Hm1:is_upper_bound (fun x : R => exists n : nat, x = sum n) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n : nat, x = sum n) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H4:forall n : nat, (n >= N)%nat -> Rabs ((/ 2) ^ n) < m
H5:Un N <= l - eps

is_upper_bound (fun x : R => exists n : nat, x = sum n) ((/ 2) ^ N)
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x0 : R => exists n0 : nat, x0 = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x0 : R => exists n0 : nat, x0 = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H4:forall n0 : nat, (n0 >= N)%nat -> Rabs ((/ 2) ^ n0) < m
H5:Un N <= l - eps
x:R
n:nat
H6:x = sum n

x <= (/ 2) ^ N
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x0 : R => exists n0 : nat, x0 = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x0 : R => exists n0 : nat, x0 = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H4:forall n0 : nat, (n0 >= N)%nat -> Rabs ((/ 2) ^ n0) < m
H5:Un N <= l - eps
x:R
n:nat
H6:x = sum n

sum n <= (/ 2) ^ N
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H4:forall n0 : nat, (n0 >= N)%nat -> Rabs ((/ 2) ^ n0) < m
H5:Un N <= l - eps
n:nat

sum n <= (/ 2) ^ N
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H4:forall n0 : nat, (n0 >= N)%nat -> Rabs ((/ 2) ^ n0) < m
H5:Un N <= l - eps
n:nat

sum N = 0
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H4:forall n0 : nat, (n0 >= N)%nat -> Rabs ((/ 2) ^ n0) < m
H5:Un N <= l - eps
n:nat
Hs:sum N = 0
sum n <= (/ 2) ^ N
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H5:Un N <= l - eps
n:nat

sum N = 0
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H4:forall n0 : nat, (n0 >= N)%nat -> Rabs ((/ 2) ^ n0) < m
H5:Un N <= l - eps
n:nat
Hs:sum N = 0
sum n <= (/ 2) ^ N
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
H5:Un 0 <= l - eps
n:nat

sum 0%nat = 0
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H5:Un (S N) <= l - eps
n:nat
IHN:Un N <= l - eps -> sum N = 0
sum (S N) = 0
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H4:forall n0 : nat, (n0 >= N)%nat -> Rabs ((/ 2) ^ n0) < m
H5:Un N <= l - eps
n:nat
Hs:sum N = 0
sum n <= (/ 2) ^ N
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H5:Un (S N) <= l - eps
n:nat
IHN:Un N <= l - eps -> sum N = 0

sum (S N) = 0
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H4:forall n0 : nat, (n0 >= N)%nat -> Rabs ((/ 2) ^ n0) < m
H5:Un N <= l - eps
n:nat
Hs:sum N = 0
sum n <= (/ 2) ^ N
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H5:Un (S N) <= l - eps
n:nat
IHN:Un N <= l - eps -> sum N = 0

sum N + (if test N then / 2 * (/ 2) ^ N else 0) = 0
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H4:forall n0 : nat, (n0 >= N)%nat -> Rabs ((/ 2) ^ n0) < m
H5:Un N <= l - eps
n:nat
Hs:sum N = 0
sum n <= (/ 2) ^ N
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H5:Un (S N) <= l - eps
n:nat
IHN:Un N <= l - eps -> sum N = 0

Un N <= l - eps
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H5:Un (S N) <= l - eps
n:nat
IHN:Un N <= l - eps -> sum N = 0
H6:Un N <= l - eps
sum N + (if test N then / 2 * (/ 2) ^ N else 0) = 0
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H4:forall n0 : nat, (n0 >= N)%nat -> Rabs ((/ 2) ^ n0) < m
H5:Un N <= l - eps
n:nat
Hs:sum N = 0
sum n <= (/ 2) ^ N
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H5:Un (S N) <= l - eps
n:nat
IHN:Un N <= l - eps -> sum N = 0

Un N <= Un (S N)
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H5:Un (S N) <= l - eps
n:nat
IHN:Un N <= l - eps -> sum N = 0
H6:Un N <= l - eps
sum N + (if test N then / 2 * (/ 2) ^ N else 0) = 0
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H4:forall n0 : nat, (n0 >= N)%nat -> Rabs ((/ 2) ^ n0) < m
H5:Un N <= l - eps
n:nat
Hs:sum N = 0
sum n <= (/ 2) ^ N
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H5:Un (S N) <= l - eps
n:nat
IHN:Un N <= l - eps -> sum N = 0

Un (S N) >= Un N
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H5:Un (S N) <= l - eps
n:nat
IHN:Un N <= l - eps -> sum N = 0
H6:Un N <= l - eps
sum N + (if test N then / 2 * (/ 2) ^ N else 0) = 0
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H4:forall n0 : nat, (n0 >= N)%nat -> Rabs ((/ 2) ^ n0) < m
H5:Un N <= l - eps
n:nat
Hs:sum N = 0
sum n <= (/ 2) ^ N
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H5:Un (S N) <= l - eps
n:nat
IHN:Un N <= l - eps -> sum N = 0

(S N >= N)%nat
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H5:Un (S N) <= l - eps
n:nat
IHN:Un N <= l - eps -> sum N = 0
H6:Un N <= l - eps
sum N + (if test N then / 2 * (/ 2) ^ N else 0) = 0
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H4:forall n0 : nat, (n0 >= N)%nat -> Rabs ((/ 2) ^ n0) < m
H5:Un N <= l - eps
n:nat
Hs:sum N = 0
sum n <= (/ 2) ^ N
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H5:Un (S N) <= l - eps
n:nat
IHN:Un N <= l - eps -> sum N = 0
H6:Un N <= l - eps

sum N + (if test N then / 2 * (/ 2) ^ N else 0) = 0
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H4:forall n0 : nat, (n0 >= N)%nat -> Rabs ((/ 2) ^ n0) < m
H5:Un N <= l - eps
n:nat
Hs:sum N = 0
sum n <= (/ 2) ^ N
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H5:Un (S N) <= l - eps
n:nat
IHN:Un N <= l - eps -> sum N = 0
H6:Un N <= l - eps

(if test N then / 2 * (/ 2) ^ N else 0) = 0
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H4:forall n0 : nat, (n0 >= N)%nat -> Rabs ((/ 2) ^ n0) < m
H5:Un N <= l - eps
n:nat
Hs:sum N = 0
sum n <= (/ 2) ^ N
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H5:Un (S N) <= l - eps
n:nat
IHN:Un N <= l - eps -> sum N = 0
H6:Un N <= l - eps

(if if Rle_lt_dec (Un N) (l - eps) then false else true then / 2 * (/ 2) ^ N else 0) = 0
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H4:forall n0 : nat, (n0 >= N)%nat -> Rabs ((/ 2) ^ n0) < m
H5:Un N <= l - eps
n:nat
Hs:sum N = 0
sum n <= (/ 2) ^ N
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H5:Un (S N) <= l - eps
n:nat
IHN:Un N <= l - eps -> sum N = 0
H6, Hle:Un N <= l - eps

0 = 0
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H5:Un (S N) <= l - eps
n:nat
IHN:Un N <= l - eps -> sum N = 0
H6:Un N <= l - eps
Hlt:l - eps < Un N
/ 2 * (/ 2) ^ N = 0
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H4:forall n0 : nat, (n0 >= N)%nat -> Rabs ((/ 2) ^ n0) < m
H5:Un N <= l - eps
n:nat
Hs:sum N = 0
sum n <= (/ 2) ^ N
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H5:Un (S N) <= l - eps
n:nat
IHN:Un N <= l - eps -> sum N = 0
H6:Un N <= l - eps
Hlt:l - eps < Un N

/ 2 * (/ 2) ^ N = 0
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H4:forall n0 : nat, (n0 >= N)%nat -> Rabs ((/ 2) ^ n0) < m
H5:Un N <= l - eps
n:nat
Hs:sum N = 0
sum n <= (/ 2) ^ N
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H4:forall n0 : nat, (n0 >= N)%nat -> Rabs ((/ 2) ^ n0) < m
H5:Un N <= l - eps
n:nat
Hs:sum N = 0

sum n <= (/ 2) ^ N
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H4:forall n0 : nat, (n0 >= N)%nat -> Rabs ((/ 2) ^ n0) < m
H5:Un N <= l - eps
n:nat
Hs:sum N = 0
Hn:(N <= n)%nat

sum n <= (/ 2) ^ N
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H4:forall n0 : nat, (n0 >= N)%nat -> Rabs ((/ 2) ^ n0) < m
H5:Un N <= l - eps
n:nat
Hs:sum N = 0
Hn:(n < N)%nat
sum n <= (/ 2) ^ N
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H4:forall n0 : nat, (n0 >= N)%nat -> Rabs ((/ 2) ^ n0) < m
H5:Un N <= l - eps
n:nat
Hs:sum N = 0
Hn:(N <= n)%nat

sum (N + (n - N))%nat <= (/ 2) ^ N
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H4:forall n0 : nat, (n0 >= N)%nat -> Rabs ((/ 2) ^ n0) < m
H5:Un N <= l - eps
n:nat
Hs:sum N = 0
Hn:(n < N)%nat
sum n <= (/ 2) ^ N
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H4:forall n0 : nat, (n0 >= N)%nat -> Rabs ((/ 2) ^ n0) < m
H5:Un N <= l - eps
n:nat
Hs:sum N = 0
Hn:(N <= n)%nat

sum N + (/ 2) ^ N - (/ 2) ^ (N + (n - N)) <= (/ 2) ^ N
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H4:forall n0 : nat, (n0 >= N)%nat -> Rabs ((/ 2) ^ n0) < m
H5:Un N <= l - eps
n:nat
Hs:sum N = 0
Hn:(n < N)%nat
sum n <= (/ 2) ^ N
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H4:forall n0 : nat, (n0 >= N)%nat -> Rabs ((/ 2) ^ n0) < m
H5:Un N <= l - eps
n:nat
Hs:sum N = 0
Hn:(N <= n)%nat

(/ 2) ^ N - (/ 2) ^ (N + (n - N)) <= (/ 2) ^ N
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H4:forall n0 : nat, (n0 >= N)%nat -> Rabs ((/ 2) ^ n0) < m
H5:Un N <= l - eps
n:nat
Hs:sum N = 0
Hn:(n < N)%nat
sum n <= (/ 2) ^ N
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H4:forall n0 : nat, (n0 >= N)%nat -> Rabs ((/ 2) ^ n0) < m
H5:Un N <= l - eps
n:nat
Hs:sum N = 0
Hn:(N <= n)%nat
k:=(N + (n - N))%nat:nat

(/ 2) ^ N - (/ 2) ^ k <= (/ 2) ^ N
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H4:forall n0 : nat, (n0 >= N)%nat -> Rabs ((/ 2) ^ n0) < m
H5:Un N <= l - eps
n:nat
Hs:sum N = 0
Hn:(n < N)%nat
sum n <= (/ 2) ^ N
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H4:forall n0 : nat, (n0 >= N)%nat -> Rabs ((/ 2) ^ n0) < m
H5:Un N <= l - eps
n:nat
Hs:sum N = 0
Hn:(N <= n)%nat
k:=(N + (n - N))%nat:nat

(/ 2) ^ N - (/ 2) ^ k < (/ 2) ^ N
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H4:forall n0 : nat, (n0 >= N)%nat -> Rabs ((/ 2) ^ n0) < m
H5:Un N <= l - eps
n:nat
Hs:sum N = 0
Hn:(n < N)%nat
sum n <= (/ 2) ^ N
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H4:forall n0 : nat, (n0 >= N)%nat -> Rabs ((/ 2) ^ n0) < m
H5:Un N <= l - eps
n:nat
Hs:sum N = 0
Hn:(N <= n)%nat
k:=(N + (n - N))%nat:nat

(/ 2) ^ k - (/ 2) ^ N + ((/ 2) ^ N - (/ 2) ^ k) < (/ 2) ^ k - (/ 2) ^ N + (/ 2) ^ N
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H4:forall n0 : nat, (n0 >= N)%nat -> Rabs ((/ 2) ^ n0) < m
H5:Un N <= l - eps
n:nat
Hs:sum N = 0
Hn:(n < N)%nat
sum n <= (/ 2) ^ N
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H4:forall n0 : nat, (n0 >= N)%nat -> Rabs ((/ 2) ^ n0) < m
H5:Un N <= l - eps
n:nat
Hs:sum N = 0
Hn:(n < N)%nat

sum n <= (/ 2) ^ N
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H4:forall n0 : nat, (n0 >= N)%nat -> Rabs ((/ 2) ^ n0) < m
H5:Un N <= l - eps
n:nat
Hs:sum N = 0
Hn:(n < N)%nat

sum n <= sum N
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H4:forall n0 : nat, (n0 >= N)%nat -> Rabs ((/ 2) ^ n0) < m
H5:Un N <= l - eps
n:nat
Hs:sum N = 0
Hn:(n < N)%nat
sum N <= (/ 2) ^ N
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H4:forall n0 : nat, (n0 >= N)%nat -> Rabs ((/ 2) ^ n0) < m
H5:Un N <= l - eps
n:nat
Hs:sum N = 0
Hn:(n < N)%nat

sum n <= sum (S n + (N - S n))%nat
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H4:forall n0 : nat, (n0 >= N)%nat -> Rabs ((/ 2) ^ n0) < m
H5:Un N <= l - eps
n:nat
Hs:sum N = 0
Hn:(n < N)%nat
sum N <= (/ 2) ^ N
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H4:forall n0 : nat, (n0 >= N)%nat -> Rabs ((/ 2) ^ n0) < m
H5:Un N <= l - eps
n:nat
Hs:sum N = 0
Hn:(n < N)%nat

sum n <= sum (n + S (N - S n))%nat
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H4:forall n0 : nat, (n0 >= N)%nat -> Rabs ((/ 2) ^ n0) < m
H5:Un N <= l - eps
n:nat
Hs:sum N = 0
Hn:(n < N)%nat
sum N <= (/ 2) ^ N
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H4:forall n0 : nat, (n0 >= N)%nat -> Rabs ((/ 2) ^ n0) < m
H5:Un N <= l - eps
n:nat
Hs:sum N = 0
Hn:(n < N)%nat

sum N <= (/ 2) ^ N
Un:nat -> R
Hug:Un_growing
l:R
H:is_lub EUn l
eps:R
Heps:eps > 0
Hi2pn:forall n0 : nat, 0 < (/ 2) ^ n0
test:=fun n0 : nat => if Rle_lt_dec (Un n0) (l - eps) then false else true:nat -> bool
sum:=let fix aux (n0 : nat) : R := match n0 with | 0%nat => 0 | S n' => aux n' + (if test n' then (/ 2) ^ n0 else 0) end in aux:nat -> R
Hsum':forall m0 n0 : nat, sum m0 <= sum (m0 + n0)%nat <= sum m0 + (/ 2) ^ m0 - (/ 2) ^ (m0 + n0)
Hsum:forall n0 : nat, 0 <= sum n0 <= 1 - (/ 2) ^ n0
m:R
Hm1:is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) m
Hm2:forall b : R, is_upper_bound (fun x : R => exists n0 : nat, x = sum n0) b -> m <= b
Hm:0 < m
H0:Rabs (/ 2) < 1
N:nat
H4:forall n0 : nat, (n0 >= N)%nat -> Rabs ((/ 2) ^ n0) < m
H5:Un N <= l - eps
n:nat
Hs:sum N = 0
Hn:(n < N)%nat

0 <= (/ 2) ^ N
now apply Rlt_le. Qed. (*********)
Un:nat -> R

Un_growing -> bound EUn -> exists l : R, Un_cv l
Un:nat -> R

Un_growing -> bound EUn -> exists l : R, Un_cv l
Un:nat -> R
Hug:Un_growing
Heub:bound EUn

exists l : R, Un_cv l
Un:nat -> R
Hug:Un_growing
Heub:bound EUn

Un_cv (proj1_sig (completeness EUn Heub EUn_noempty))
Un:nat -> R
Hug:Un_growing
Heub:bound EUn
l:R
H:is_lub EUn l

Un_cv (proj1_sig (exist (fun m : R => is_lub EUn m) l H))
now apply Un_cv_crit_lub. Qed. (*********)
Un:nat -> R

forall N : nat, exists M : R, forall n : nat, (n <= N)%nat -> Un n <= M
Un:nat -> R

forall N : nat, exists M : R, forall n : nat, (n <= N)%nat -> Un n <= M
Un:nat -> R

exists M : R, forall n : nat, (n <= 0)%nat -> Un n <= M
Un:nat -> R
N:nat
HrecN:exists M : R, forall n : nat, (n <= N)%nat -> Un n <= M
exists M : R, forall n : nat, (n <= S N)%nat -> Un n <= M
Un:nat -> R
N:nat
HrecN:exists M : R, forall n : nat, (n <= N)%nat -> Un n <= M

exists M : R, forall n : nat, (n <= S N)%nat -> Un n <= M
Un:nat -> R
N:nat
x:R
H:forall n0 : nat, (n0 <= N)%nat -> Un n0 <= x
n:nat
H0:(n <= S N)%nat
H2:Un n <= Un (S N) \/ Un n <= x -> Un n <= Rmax (Un (S N)) x
H1:n = S N

Un (S N) <= Rmax (Un (S N)) x
Un:nat -> R
N:nat
x:R
H:forall n0 : nat, (n0 <= N)%nat -> Un n0 <= x
n:nat
H0:(n <= S N)%nat
H2:Un n <= Un (S N) \/ Un n <= x -> Un n <= Rmax (Un (S N)) x
m:nat
H3:(n <= N)%nat
H1:m = N
Un n <= Rmax (Un (S N)) x
Un:nat -> R
N:nat
x:R
H:forall n0 : nat, (n0 <= N)%nat -> Un n0 <= x
n:nat
H0:(n <= S N)%nat
H2:Un n <= Un (S N) \/ Un n <= x -> Un n <= Rmax (Un (S N)) x
m:nat
H3:(n <= N)%nat
H1:m = N

Un n <= Rmax (Un (S N)) x
apply (H2 (or_intror (Un n <= Un (S N)) (H n H3))). Qed. (*********)
Un:nat -> R

Cauchy_crit -> bound EUn
Un:nat -> R

Cauchy_crit -> bound EUn
Un:nat -> R
x:nat
H0:forall m : nat, (x >= x)%nat -> (m >= x)%nat -> R_dist (Un x) (Un m) < 1
x0:R
H2:forall n : nat, (n <= x)%nat -> Un n <= x0
x1:R
x2:nat
H:x1 = Un x2
y:(x <= x2)%nat

x1 <= Rmax x0 (Un x + 1)
Un:nat -> R
x:nat
H0:forall m : nat, (x >= x)%nat -> (m >= x)%nat -> R_dist (Un x) (Un m) < 1
x0:R
H2:forall n : nat, (n <= x)%nat -> Un n <= x0
x1:R
x2:nat
H:x1 = Un x2
y:(x2 <= x)%nat
x1 <= Rmax x0 (Un x + 1)
Un:nat -> R
x:nat
H0:forall m : nat, (x >= x)%nat -> (m >= x)%nat -> R_dist (Un x) (Un m) < 1
x0:R
H2:forall n : nat, (n <= x)%nat -> Un n <= x0
x1:R
x2:nat
H:x1 = Un x2
y:(x2 <= x)%nat

x1 <= Rmax x0 (Un x + 1)
generalize (H2 x2 y); clear H2 H0; intro; rewrite <- H in H0; elim (Rmax_Rle x0 (Un x + 1) x1); intros; clear H1; apply H2; left; assumption. Qed. End sequence. (*****************************************************************)

Definition of Power Series and properties

(*                                                               *)
(*****************************************************************)

Section Isequence.

(*********)
  Variable An : nat -> R.

(*********)
  Definition Pser (x l:R) : Prop := infinite_sum (fun n:nat => An n * x ^ n) l.

End Isequence.


forall x : R, Rabs x < 1 -> Pser (fun _ : nat => 1) x (/ (1 - x))

forall x : R, Rabs x < 1 -> Pser (fun _ : nat => 1) x (/ (1 - x))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0

x = 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> R_dist (sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n) (/ (1 - x)) < eps
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
x <> 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> R_dist (sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n) (/ (1 - x)) < eps
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x = 0
n:nat
H2:(n >= 0)%nat

sum_f_R0 (fun n0 : nat => 1 * 0 ^ n0) n = 1 -> R_dist (sum_f_R0 (fun n0 : nat => 1 * 0 ^ n0) n) 1 < eps
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x = 0
n:nat
H2:(n >= 0)%nat
sum_f_R0 (fun n0 : nat => 1 * 0 ^ n0) n = 1
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
x <> 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> R_dist (sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n) (/ (1 - x)) < eps
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x = 0
n:nat
H2:(n >= 0)%nat

sum_f_R0 (fun n0 : nat => 1 * 0 ^ n0) n = 1
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
x <> 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> R_dist (sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n) (/ (1 - x)) < eps
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x = 0
n:nat
H2:(n >= 0)%nat

1 * 1 = 1
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x = 0
n:nat
H2:(n >= 0)%nat
forall n0 : nat, sum_f_R0 (fun n1 : nat => 1 * 0 ^ n1) n0 = 1 -> sum_f_R0 (fun n1 : nat => 1 * 0 ^ n1) n0 + 1 * (0 * 0 ^ n0) = 1
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
x <> 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> R_dist (sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n) (/ (1 - x)) < eps
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x = 0
n:nat
H2:(n >= 0)%nat

forall n0 : nat, sum_f_R0 (fun n1 : nat => 1 * 0 ^ n1) n0 = 1 -> sum_f_R0 (fun n1 : nat => 1 * 0 ^ n1) n0 + 1 * (0 * 0 ^ n0) = 1
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
x <> 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> R_dist (sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n) (/ (1 - x)) < eps
x:R
H:Rabs x < 1
eps:R
H0:eps > 0

x <> 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> R_dist (sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n) (/ (1 - x)) < eps
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0

0 < eps * (Rabs (1 - x) * Rabs (/ x)) -> exists N : nat, forall n : nat, (n >= N)%nat -> R_dist (sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n) (/ (1 - x)) < eps
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
0 < eps * (Rabs (1 - x) * Rabs (/ x))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat

sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n -> R_dist (sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n) (/ (1 - x)) < eps
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
0 < eps * (Rabs (1 - x) * Rabs (/ x))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n

0 < Rabs (1 - x)
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
Rabs (1 - x) * R_dist (sum_f_R0 (fun n0 : nat => x ^ n0) n) (/ (1 - x)) < Rabs (1 - x) * eps
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
0 < eps * (Rabs (1 - x) * Rabs (/ x))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n

1 - x <> 0
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
Rabs (1 - x) * R_dist (sum_f_R0 (fun n0 : nat => x ^ n0) n) (/ (1 - x)) < Rabs (1 - x) * eps
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
0 < eps * (Rabs (1 - x) * Rabs (/ x))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n

1 <> x
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
Rabs (1 - x) * R_dist (sum_f_R0 (fun n0 : nat => x ^ n0) n) (/ (1 - x)) < Rabs (1 - x) * eps
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
0 < eps * (Rabs (1 - x) * Rabs (/ x))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n

1 < x \/ 1 > x
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
Rabs (1 - x) * R_dist (sum_f_R0 (fun n0 : nat => x ^ n0) n) (/ (1 - x)) < Rabs (1 - x) * eps
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
0 < eps * (Rabs (1 - x) * Rabs (/ x))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n

x < 1
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
Rabs (1 - x) * R_dist (sum_f_R0 (fun n0 : nat => x ^ n0) n) (/ (1 - x)) < Rabs (1 - x) * eps
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
0 < eps * (Rabs (1 - x) * Rabs (/ x))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n

x <= Rabs x
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
Rabs x < 1
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
Rabs (1 - x) * R_dist (sum_f_R0 (fun n0 : nat => x ^ n0) n) (/ (1 - x)) < Rabs (1 - x) * eps
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
0 < eps * (Rabs (1 - x) * Rabs (/ x))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n

Rabs x < 1
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
Rabs (1 - x) * R_dist (sum_f_R0 (fun n0 : nat => x ^ n0) n) (/ (1 - x)) < Rabs (1 - x) * eps
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
0 < eps * (Rabs (1 - x) * Rabs (/ x))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n

Rabs (1 - x) * R_dist (sum_f_R0 (fun n0 : nat => x ^ n0) n) (/ (1 - x)) < Rabs (1 - x) * eps
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
0 < eps * (Rabs (1 - x) * Rabs (/ x))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n

Rabs ((1 - x) * (sum_f_R0 (fun n0 : nat => x ^ n0) n - / (1 - x))) < Rabs (1 - x) * eps
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
0 < eps * (Rabs (1 - x) * Rabs (/ x))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n

Rabs ((1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n - (1 - x) * / (1 - x)) < Rabs (1 - x) * eps
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
0 < eps * (Rabs (1 - x) * Rabs (/ x))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n

(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1)) -> Rabs ((1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n - (1 - x) * / (1 - x)) < Rabs (1 - x) * eps
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
0 < eps * (Rabs (1 - x) * Rabs (/ x))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))

Rabs (- (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1)) - (1 - x) * / (1 - x)) < Rabs (1 - x) * eps
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
0 < eps * (Rabs (1 - x) * Rabs (/ x))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))

Rabs (- (x ^ (n + 1) - 1) - (1 - x) * / (1 - x)) < Rabs (1 - x) * eps
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
0 < eps * (Rabs (1 - x) * Rabs (/ x))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))

Rabs (- (x ^ (n + 1) - 1) - 1) < Rabs (1 - x) * eps
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
1 - x <> 0
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
0 < eps * (Rabs (1 - x) * Rabs (/ x))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))

- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1) -> Rabs (- (x ^ (n + 1) - 1) - 1) < Rabs (1 - x) * eps
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
1 - x <> 0
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
0 < eps * (Rabs (1 - x) * Rabs (/ x))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
H7:- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)

Rabs (- x ^ (n + 1)) < Rabs (1 - x) * eps
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
1 - x <> 0
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
0 < eps * (Rabs (1 - x) * Rabs (/ x))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
H7:- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)

(n + 1)%nat = S n -> Rabs (x ^ (n + 1)) < Rabs (1 - x) * eps
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
H7:- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)
(n + 1)%nat = S n
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
1 - x <> 0
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
0 < eps * (Rabs (1 - x) * Rabs (/ x))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
H7:- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)
H8:(n + 1)%nat = S n

Rabs x * Rabs (x ^ n) < Rabs x * (eps * (Rabs (1 - x) * Rabs (/ x)))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
H7:- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)
H8:(n + 1)%nat = S n
Rabs x * (eps * (Rabs (1 - x) * Rabs (/ x))) <= Rabs (1 - x) * eps
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
H7:- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)
(n + 1)%nat = S n
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
1 - x <> 0
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
0 < eps * (Rabs (1 - x) * Rabs (/ x))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
H7:- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)
H8:(n + 1)%nat = S n

0 < Rabs x
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
H7:- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)
H8:(n + 1)%nat = S n
Rabs (x ^ n) < eps * (Rabs (1 - x) * Rabs (/ x))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
H7:- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)
H8:(n + 1)%nat = S n
Rabs x * (eps * (Rabs (1 - x) * Rabs (/ x))) <= Rabs (1 - x) * eps
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
H7:- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)
(n + 1)%nat = S n
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
1 - x <> 0
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
0 < eps * (Rabs (1 - x) * Rabs (/ x))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
H7:- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)
H8:(n + 1)%nat = S n

x <> 0
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
H7:- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)
H8:(n + 1)%nat = S n
Rabs (x ^ n) < eps * (Rabs (1 - x) * Rabs (/ x))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
H7:- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)
H8:(n + 1)%nat = S n
Rabs x * (eps * (Rabs (1 - x) * Rabs (/ x))) <= Rabs (1 - x) * eps
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
H7:- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)
(n + 1)%nat = S n
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
1 - x <> 0
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
0 < eps * (Rabs (1 - x) * Rabs (/ x))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
H7:- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)
H8:(n + 1)%nat = S n

Rabs (x ^ n) < eps * (Rabs (1 - x) * Rabs (/ x))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
H7:- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)
H8:(n + 1)%nat = S n
Rabs x * (eps * (Rabs (1 - x) * Rabs (/ x))) <= Rabs (1 - x) * eps
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
H7:- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)
(n + 1)%nat = S n
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
1 - x <> 0
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
0 < eps * (Rabs (1 - x) * Rabs (/ x))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
H7:- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)
H8:(n + 1)%nat = S n

Rabs x * (eps * (Rabs (1 - x) * Rabs (/ x))) <= Rabs (1 - x) * eps
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
H7:- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)
(n + 1)%nat = S n
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
1 - x <> 0
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
0 < eps * (Rabs (1 - x) * Rabs (/ x))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
H7:- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)
H8:(n + 1)%nat = S n

Rabs x * (eps * (Rabs (1 - x) * Rabs (/ x))) = Rabs x * Rabs (/ x) * (eps * Rabs (1 - x)) -> Rabs x * (eps * (Rabs (1 - x) * Rabs (/ x))) <= Rabs (1 - x) * eps
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
H7:- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)
H8:(n + 1)%nat = S n
Rabs x * (eps * (Rabs (1 - x) * Rabs (/ x))) = Rabs x * Rabs (/ x) * (eps * Rabs (1 - x))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
H7:- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)
(n + 1)%nat = S n
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
1 - x <> 0
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
0 < eps * (Rabs (1 - x) * Rabs (/ x))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
H7:- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)
H8:Rabs x * (eps * (Rabs (1 - x) * Rabs (/ x))) = Rabs x * Rabs (/ x) * (eps * Rabs (1 - x))

Rabs 1 * (eps * Rabs (1 - x)) <= Rabs (1 - x) * eps
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
H7:- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)
H8:Rabs x * (eps * (Rabs (1 - x) * Rabs (/ x))) = Rabs x * Rabs (/ x) * (eps * Rabs (1 - x))
x <> 0
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
H7:- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)
H8:(n + 1)%nat = S n
Rabs x * (eps * (Rabs (1 - x) * Rabs (/ x))) = Rabs x * Rabs (/ x) * (eps * Rabs (1 - x))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
H7:- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)
(n + 1)%nat = S n
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
1 - x <> 0
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
0 < eps * (Rabs (1 - x) * Rabs (/ x))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
H7:- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)
H8:Rabs x * (eps * (Rabs (1 - x) * Rabs (/ x))) = Rabs x * Rabs (/ x) * (eps * Rabs (1 - x))

1 * (eps * Rabs (1 - x)) = Rabs (1 - x) * eps -> 1 * (eps * Rabs (1 - x)) <= Rabs (1 - x) * eps
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
H7:- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)
H8:Rabs x * (eps * (Rabs (1 - x) * Rabs (/ x))) = Rabs x * Rabs (/ x) * (eps * Rabs (1 - x))
1 * (eps * Rabs (1 - x)) = Rabs (1 - x) * eps
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
H7:- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)
H8:Rabs x * (eps * (Rabs (1 - x) * Rabs (/ x))) = Rabs x * Rabs (/ x) * (eps * Rabs (1 - x))
x <> 0
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
H7:- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)
H8:(n + 1)%nat = S n
Rabs x * (eps * (Rabs (1 - x) * Rabs (/ x))) = Rabs x * Rabs (/ x) * (eps * Rabs (1 - x))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
H7:- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)
(n + 1)%nat = S n
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
1 - x <> 0
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
0 < eps * (Rabs (1 - x) * Rabs (/ x))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
H7:- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)
H8:Rabs x * (eps * (Rabs (1 - x) * Rabs (/ x))) = Rabs x * Rabs (/ x) * (eps * Rabs (1 - x))

1 * (eps * Rabs (1 - x)) = Rabs (1 - x) * eps
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
H7:- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)
H8:Rabs x * (eps * (Rabs (1 - x) * Rabs (/ x))) = Rabs x * Rabs (/ x) * (eps * Rabs (1 - x))
x <> 0
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
H7:- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)
H8:(n + 1)%nat = S n
Rabs x * (eps * (Rabs (1 - x) * Rabs (/ x))) = Rabs x * Rabs (/ x) * (eps * Rabs (1 - x))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
H7:- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)
(n + 1)%nat = S n
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
1 - x <> 0
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
0 < eps * (Rabs (1 - x) * Rabs (/ x))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
H7:- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)
H8:Rabs x * (eps * (Rabs (1 - x) * Rabs (/ x))) = Rabs x * Rabs (/ x) * (eps * Rabs (1 - x))

x <> 0
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
H7:- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)
H8:(n + 1)%nat = S n
Rabs x * (eps * (Rabs (1 - x) * Rabs (/ x))) = Rabs x * Rabs (/ x) * (eps * Rabs (1 - x))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
H7:- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)
(n + 1)%nat = S n
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
1 - x <> 0
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
0 < eps * (Rabs (1 - x) * Rabs (/ x))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
H7:- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)
H8:(n + 1)%nat = S n

Rabs x * (eps * (Rabs (1 - x) * Rabs (/ x))) = Rabs x * Rabs (/ x) * (eps * Rabs (1 - x))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
H7:- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)
(n + 1)%nat = S n
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
1 - x <> 0
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
0 < eps * (Rabs (1 - x) * Rabs (/ x))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
H7:- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)

(n + 1)%nat = S n
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
1 - x <> 0
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
0 < eps * (Rabs (1 - x) * Rabs (/ x))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))

- (x ^ (n + 1) - 1) - 1 = - x ^ (n + 1)
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
1 - x <> 0
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
0 < eps * (Rabs (1 - x) * Rabs (/ x))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))

1 - x <> 0
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
0 < eps * (Rabs (1 - x) * Rabs (/ x))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))

1 <> x
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
0 < eps * (Rabs (1 - x) * Rabs (/ x))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))

1 < x \/ 1 > x
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
0 < eps * (Rabs (1 - x) * Rabs (/ x))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))

x < 1
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
0 < eps * (Rabs (1 - x) * Rabs (/ x))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))

x <= Rabs x
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
Rabs x < 1
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
0 < eps * (Rabs (1 - x) * Rabs (/ x))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
H6:(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))

Rabs x < 1
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
0 < eps * (Rabs (1 - x) * Rabs (/ x))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
H5:sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n

(1 - x) * sum_f_R0 (fun n0 : nat => x ^ n0) n = - (sum_f_R0 (fun n0 : nat => x ^ n0) n * (x - 1))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
0 < eps * (Rabs (1 - x) * Rabs (/ x))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat

sum_f_R0 (fun n0 : nat => 1 * x ^ n0) n = sum_f_R0 (fun n0 : nat => x ^ n0) n
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
0 < eps * (Rabs (1 - x) * Rabs (/ x))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat

1 * 1 = 1
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
forall n0 : nat, sum_f_R0 (fun n1 : nat => 1 * x ^ n1) n0 = sum_f_R0 (fun n1 : nat => x ^ n1) n0 -> sum_f_R0 (fun n1 : nat => 1 * x ^ n1) n0 + 1 * (x * x ^ n0) = sum_f_R0 (fun n1 : nat => x ^ n1) n0 + x * x ^ n0
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
0 < eps * (Rabs (1 - x) * Rabs (/ x))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n0 : nat, (n0 >= N)%nat -> Rabs (x ^ n0) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat

forall n0 : nat, sum_f_R0 (fun n1 : nat => 1 * x ^ n1) n0 = sum_f_R0 (fun n1 : nat => x ^ n1) n0 -> sum_f_R0 (fun n1 : nat => 1 * x ^ n1) n0 + 1 * (x * x ^ n0) = sum_f_R0 (fun n1 : nat => x ^ n1) n0 + x * x ^ n0
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
0 < eps * (Rabs (1 - x) * Rabs (/ x))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
H2:0 < eps * (Rabs (1 - x) * Rabs (/ x))
N:nat
H3:forall n1 : nat, (n1 >= N)%nat -> Rabs (x ^ n1) < eps * (Rabs (1 - x) * Rabs (/ x))
n:nat
H4:(n >= N)%nat
n0:nat
H5:sum_f_R0 (fun n1 : nat => 1 * x ^ n1) n0 = sum_f_R0 (fun n1 : nat => x ^ n1) n0

sum_f_R0 (fun n1 : nat => x ^ n1) n0 + 1 * (x * x ^ n0) = sum_f_R0 (fun n1 : nat => x ^ n1) n0 + x * x ^ n0
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
0 < eps * (Rabs (1 - x) * Rabs (/ x))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0

0 < eps * (Rabs (1 - x) * Rabs (/ x))
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0

0 < eps
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
0 < Rabs (1 - x) * Rabs (/ x)
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0

0 < Rabs (1 - x) * Rabs (/ x)
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0

0 < Rabs (1 - x)
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
0 < Rabs (/ x)
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0

1 - x <> 0
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
0 < Rabs (/ x)
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0

1 <> x
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
0 < Rabs (/ x)
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0

1 < x \/ 1 > x
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
0 < Rabs (/ x)
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0

x < 1
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
0 < Rabs (/ x)
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0

x <= Rabs x
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
Rabs x < 1
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
0 < Rabs (/ x)
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0

Rabs x < 1
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0
0 < Rabs (/ x)
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0

0 < Rabs (/ x)
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0

/ x <> 0
x:R
H:Rabs x < 1
eps:R
H0:eps > 0
H1:x <> 0

x <> 0
assumption. Qed. (* Convergence is preserved after shifting the indices. *)

forall (f : nat -> R) (k : nat) (l : R), Un_cv (fun n : nat => f (n + k)%nat) l -> Un_cv f l
f':nat -> R
k:nat
l:R
cvfk:Un_cv (fun n : nat => f' (n + k)%nat) l
eps:R
ep:eps > 0
N:nat
Pn:forall n : nat, (n >= N)%nat -> R_dist (f' (n + k)%nat) l < eps

exists N0 : nat, forall n : nat, (n >= N0)%nat -> R_dist (f' n) l < eps
f':nat -> R
k:nat
l:R
cvfk:Un_cv (fun n0 : nat => f' (n0 + k)%nat) l
eps:R
ep:eps > 0
N:nat
Pn:forall n0 : nat, (n0 >= N)%nat -> R_dist (f' (n0 + k)%nat) l < eps
n:nat
nN:(n >= N + k)%nat

n = (n - k + k)%nat
f':nat -> R
k:nat
l:R
cvfk:Un_cv (fun n0 : nat => f' (n0 + k)%nat) l
eps:R
ep:eps > 0
N:nat
Pn:forall n0 : nat, (n0 >= N)%nat -> R_dist (f' (n0 + k)%nat) l < eps
n:nat
nN:(n >= N + k)%nat
tmp:n = (n - k + k)%nat
R_dist (f' n) l < eps
f':nat -> R
k:nat
l:R
cvfk:Un_cv (fun n0 : nat => f' (n0 + k)%nat) l
eps:R
ep:eps > 0
N:nat
Pn:forall n0 : nat, (n0 >= N)%nat -> R_dist (f' (n0 + k)%nat) l < eps
n:nat
nN:(n >= N + k)%nat
tmp:n = (n - k + k)%nat

R_dist (f' n) l < eps
rewrite tmp; apply Pn; apply Nat.le_add_le_sub_r; assumption. Qed.

forall (f : nat -> R) (k : nat) (l : R), Un_cv f l -> Un_cv (fun n : nat => f (n + k)%nat) l
f':nat -> R
k:nat
l:R
cvf:Un_cv f' l
eps:R
ep:eps > 0
N:nat
Pn:forall n : nat, (n >= N)%nat -> R_dist (f' n) l < eps

exists N0 : nat, forall n : nat, (n >= N0)%nat -> R_dist (f' (n + k)%nat) l < eps
exists N; intros n nN; apply Pn; auto with arith. Qed. (* Growing property is preserved after shifting the indices (one way only) *)

forall (k : nat) (un : nat -> R), Un_growing un -> Un_growing (fun n : nat => un (n + k)%nat)

forall (k : nat) (un : nat -> R), Un_growing un -> Un_growing (fun n : nat => un (n + k)%nat)
intros k un P n; apply P. Qed.