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 Rseries.
Require Import SeqProp.
Require Import PartSum.
Require Import Max.
Local Open Scope R_scope.

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

Formalization of alternated series

Definition tg_alt (Un:nat -> R) (i:nat) : R := (-1) ^ i * Un i.
Definition positivity_seq (Un:nat -> R) : Prop := forall n:nat, 0 <= Un n.


forall Un : nat -> R, Un_decreasing Un -> Un_growing (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N)))

forall Un : nat -> R, Un_decreasing Un -> Un_growing (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N)))
Un:nat -> R
H:Un_decreasing Un
n:nat

sum_f_R0 (tg_alt Un) (S (2 * n)) <= sum_f_R0 (tg_alt Un) (S (2 * S n))
Un:nat -> R
H:Un_decreasing Un
n:nat

(2 * S n)%nat = S (S (2 * n)) -> sum_f_R0 (tg_alt Un) (S (2 * n)) <= sum_f_R0 (tg_alt Un) (S (2 * S n))
Un:nat -> R
H:Un_decreasing Un
n:nat
(2 * S n)%nat = S (S (2 * n))
Un:nat -> R
H:Un_decreasing Un
n:nat
H0:(2 * S n)%nat = S (S (2 * n))

sum_f_R0 (tg_alt Un) (S (2 * n)) <= sum_f_R0 (tg_alt Un) (S (S (S (2 * n))))
Un:nat -> R
H:Un_decreasing Un
n:nat
(2 * S n)%nat = S (S (2 * n))
Un:nat -> R
H:Un_decreasing Un
n:nat
H0:(2 * S n)%nat = S (S (2 * n))

tg_alt Un (S (2 * n)) <= tg_alt Un (S (2 * n)) + (tg_alt Un (S (S (2 * n))) + tg_alt Un (S (S (S (2 * n)))))
Un:nat -> R
H:Un_decreasing Un
n:nat
(2 * S n)%nat = S (S (2 * n))
Un:nat -> R
H:Un_decreasing Un
n:nat
H0:(2 * S n)%nat = S (S (2 * n))

tg_alt Un (S (2 * n)) + 0 <= tg_alt Un (S (2 * n)) + (tg_alt Un (S (S (2 * n))) + tg_alt Un (S (S (S (2 * n)))))
Un:nat -> R
H:Un_decreasing Un
n:nat
(2 * S n)%nat = S (S (2 * n))
Un:nat -> R
H:Un_decreasing Un
n:nat
H0:(2 * S n)%nat = S (S (2 * n))

0 <= tg_alt Un (S (S (2 * n))) + tg_alt Un (S (S (S (2 * n))))
Un:nat -> R
H:Un_decreasing Un
n:nat
(2 * S n)%nat = S (S (2 * n))
Un:nat -> R
H:Un_decreasing Un
n:nat
H0:(2 * S n)%nat = S (S (2 * n))

0 <= Un (2 * S n)%nat + -1 * Un (S (2 * S n))
Un:nat -> R
H:Un_decreasing Un
n:nat
(2 * S n)%nat = S (S (2 * n))
Un:nat -> R
H:Un_decreasing Un
n:nat
H0:(2 * S n)%nat = S (S (2 * n))

Un (S (2 * S n)) + 0 <= Un (S (2 * S n)) + (Un (2 * S n)%nat + -1 * Un (S (2 * S n)))
Un:nat -> R
H:Un_decreasing Un
n:nat
(2 * S n)%nat = S (S (2 * n))
Un:nat -> R
H:Un_decreasing Un
n:nat
H0:(2 * S n)%nat = S (S (2 * n))

Un (S (2 * S n)) <= Un (2 * S n)%nat
Un:nat -> R
H:Un_decreasing Un
n:nat
(2 * S n)%nat = S (S (2 * n))
Un:nat -> R
H:Un_decreasing Un
n:nat

(2 * S n)%nat = S (S (2 * n))
Un:nat -> R
H:Un_decreasing Un
n:nat
H0:forall n0 : nat, S n0 = (n0 + 1)%nat

(2 * S n)%nat = S (S (2 * n))
rewrite (H0 n); rewrite (H0 (S (2 * n))); rewrite (H0 (2 * n)%nat); ring. Qed.

forall Un : nat -> R, Un_decreasing Un -> Un_decreasing (fun N : nat => sum_f_R0 (tg_alt Un) (2 * N))

forall Un : nat -> R, Un_decreasing Un -> Un_decreasing (fun N : nat => sum_f_R0 (tg_alt Un) (2 * N))
Un:nat -> R
H:Un_decreasing Un
n:nat

sum_f_R0 (tg_alt Un) (2 * S n) <= sum_f_R0 (tg_alt Un) (2 * n)
Un:nat -> R
H:Un_decreasing Un
n:nat

(2 * S n)%nat = S (S (2 * n)) -> sum_f_R0 (tg_alt Un) (2 * S n) <= sum_f_R0 (tg_alt Un) (2 * n)
Un:nat -> R
H:Un_decreasing Un
n:nat
(2 * S n)%nat = S (S (2 * n))
Un:nat -> R
H:Un_decreasing Un
n:nat
H0:(2 * S n)%nat = S (S (2 * n))

sum_f_R0 (tg_alt Un) (2 * n) + (tg_alt Un (S (2 * n)) + tg_alt Un (S (S (2 * n)))) <= sum_f_R0 (tg_alt Un) (2 * n)
Un:nat -> R
H:Un_decreasing Un
n:nat
(2 * S n)%nat = S (S (2 * n))
Un:nat -> R
H:Un_decreasing Un
n:nat
H0:(2 * S n)%nat = S (S (2 * n))

sum_f_R0 (tg_alt Un) (2 * n) + (tg_alt Un (S (2 * n)) + tg_alt Un (S (S (2 * n)))) <= sum_f_R0 (tg_alt Un) (2 * n) + 0
Un:nat -> R
H:Un_decreasing Un
n:nat
(2 * S n)%nat = S (S (2 * n))
Un:nat -> R
H:Un_decreasing Un
n:nat
H0:(2 * S n)%nat = S (S (2 * n))

tg_alt Un (S (2 * n)) + tg_alt Un (S (S (2 * n))) <= 0
Un:nat -> R
H:Un_decreasing Un
n:nat
(2 * S n)%nat = S (S (2 * n))
Un:nat -> R
H:Un_decreasing Un
n:nat
H0:(2 * S n)%nat = S (S (2 * n))

-1 * Un (S (2 * n)) + Un (2 * S n)%nat <= 0
Un:nat -> R
H:Un_decreasing Un
n:nat
(2 * S n)%nat = S (S (2 * n))
Un:nat -> R
H:Un_decreasing Un
n:nat
H0:(2 * S n)%nat = S (S (2 * n))

Un (S (2 * n)) + (-1 * Un (S (2 * n)) + Un (2 * S n)%nat) <= Un (S (2 * n)) + 0
Un:nat -> R
H:Un_decreasing Un
n:nat
(2 * S n)%nat = S (S (2 * n))
Un:nat -> R
H:Un_decreasing Un
n:nat
H0:(2 * S n)%nat = S (S (2 * n))

Un (2 * S n)%nat <= Un (S (2 * n))
Un:nat -> R
H:Un_decreasing Un
n:nat
(2 * S n)%nat = S (S (2 * n))
Un:nat -> R
H:Un_decreasing Un
n:nat

(2 * S n)%nat = S (S (2 * n))
Un:nat -> R
H:Un_decreasing Un
n:nat
H0:forall n0 : nat, S n0 = (n0 + 1)%nat

(2 * S n)%nat = S (S (2 * n))
rewrite (H0 n); rewrite (H0 (S (2 * n))); rewrite (H0 (2 * n)%nat); ring. Qed. (**********)

forall (Un : nat -> R) (N : nat), Un_decreasing Un -> positivity_seq Un -> sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0

forall (Un : nat -> R) (N : nat), Un_decreasing Un -> positivity_seq Un -> sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un

sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * 0)) <= 0
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * S N)) <= 0
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un

-1 * Un 1%nat + -1 * -1 * Un 2%nat <= 0
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * S N)) <= 0
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un

-1 * Un 1%nat + Un 2%nat <= 0
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * S N)) <= 0
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un

Un 1%nat + (-1 * Un 1%nat + Un 2%nat) <= Un 1%nat
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * S N)) <= 0
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0

sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * S N)) <= 0
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0

S (2 * S N) = S (S (S (2 * N))) -> sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * S N)) <= 0
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
S (2 * S N) = S (S (S (2 * N)))
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
H1:S (2 * S N) = S (S (S (2 * N)))

sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) + tg_alt Un (S (S (S (2 * N)))) + tg_alt Un (S (S (S (S (2 * N))))) <= 0
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
S (2 * S N) = S (S (S (2 * N)))
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
H1:S (2 * S N) = S (S (S (2 * N)))

sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) + tg_alt Un (S (S (S (2 * N)))) + tg_alt Un (S (S (S (S (2 * N))))) <= sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N))
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
H1:S (2 * S N) = S (S (S (2 * N)))
sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
S (2 * S N) = S (S (S (2 * N)))
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
H1:S (2 * S N) = S (S (S (2 * N)))

sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) + tg_alt Un (S (S (S (2 * N)))) + tg_alt Un (S (S (S (S (2 * N))))) <= sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) + 0
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
H1:S (2 * S N) = S (S (S (2 * N)))
sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
S (2 * S N) = S (S (S (2 * N)))
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
H1:S (2 * S N) = S (S (S (2 * N)))

tg_alt Un (S (S (S (2 * N)))) + tg_alt Un (S (S (S (S (2 * N))))) <= 0
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
H1:S (2 * S N) = S (S (S (2 * N)))
sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
S (2 * S N) = S (S (S (2 * N)))
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
H1:S (2 * S N) = S (S (S (2 * N)))

(-1) ^ S (2 * S N) * Un (S (2 * S N)) + (-1) ^ S (S (2 * S N)) * Un (S (S (2 * S N))) <= 0
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
H1:S (2 * S N) = S (S (S (2 * N)))
sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
S (2 * S N) = S (S (S (2 * N)))
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
H1:S (2 * S N) = S (S (S (2 * N)))

-1 * Un (S (2 * S N)) + (-1) ^ S (S (2 * S N)) * Un (S (S (2 * S N))) <= 0
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
H1:S (2 * S N) = S (S (S (2 * N)))
sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
S (2 * S N) = S (S (S (2 * N)))
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
H1:S (2 * S N) = S (S (S (2 * N)))

S (S (2 * S N)) = (2 * S (S N))%nat -> -1 * Un (S (2 * S N)) + (-1) ^ S (S (2 * S N)) * Un (S (S (2 * S N))) <= 0
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
H1:S (2 * S N) = S (S (S (2 * N)))
S (S (2 * S N)) = (2 * S (S N))%nat
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
H1:S (2 * S N) = S (S (S (2 * N)))
sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
S (2 * S N) = S (S (S (2 * N)))
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
H1:S (2 * S N) = S (S (S (2 * N)))
H2:S (S (2 * S N)) = (2 * S (S N))%nat

-1 * Un (S (2 * S N)) + Un (S (S (2 * S N))) <= 0
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
H1:S (2 * S N) = S (S (S (2 * N)))
S (S (2 * S N)) = (2 * S (S N))%nat
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
H1:S (2 * S N) = S (S (S (2 * N)))
sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
S (2 * S N) = S (S (S (2 * N)))
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
H1:S (2 * S N) = S (S (S (2 * N)))
H2:S (S (2 * S N)) = (2 * S (S N))%nat

Un (S (2 * S N)) + (-1 * Un (S (2 * S N)) + Un (S (S (2 * S N)))) <= Un (S (2 * S N)) + 0
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
H1:S (2 * S N) = S (S (S (2 * N)))
S (S (2 * S N)) = (2 * S (S N))%nat
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
H1:S (2 * S N) = S (S (S (2 * N)))
sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
S (2 * S N) = S (S (S (2 * N)))
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
H1:S (2 * S N) = S (S (S (2 * N)))
H2:S (S (2 * S N)) = (2 * S (S N))%nat

Un (S (S (2 * S N))) <= Un (S (2 * S N))
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
H1:S (2 * S N) = S (S (S (2 * N)))
S (S (2 * S N)) = (2 * S (S N))%nat
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
H1:S (2 * S N) = S (S (S (2 * N)))
sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
S (2 * S N) = S (S (S (2 * N)))
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
H1:S (2 * S N) = S (S (S (2 * N)))

S (S (2 * S N)) = (2 * S (S N))%nat
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
H1:S (2 * S N) = S (S (S (2 * N)))
sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
S (2 * S N) = S (S (S (2 * N)))
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
H1:S (2 * S N) = S (S (S (2 * N)))

sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0
S (2 * S N) = S (S (S (2 * N)))
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0

S (2 * S N) = S (S (S (2 * N)))
ring. Qed.
A more general inequality

forall (Un : nat -> R) (N : nat), Un_decreasing Un -> positivity_seq Un -> sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0

forall (Un : nat -> R) (N : nat), Un_decreasing Un -> positivity_seq Un -> sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un

sum_f_R0 (fun i : nat => tg_alt Un (S i)) 0 <= 0
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0
sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S N) <= 0
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un

-1 * Un 1%nat <= 0
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0
sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S N) <= 0
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un

Un 1%nat + -1 * Un 1%nat <= Un 1%nat + 0
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0
sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S N) <= 0
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0

sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S N) <= 0
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0
H1:exists p : nat, N = (2 * p)%nat \/ N = S (2 * p)

sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S N) <= 0
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0
H1:exists p : nat, N = (2 * p)%nat \/ N = S (2 * p)
x:nat
H2:N = (2 * x)%nat \/ N = S (2 * x)

sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S N) <= 0
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0
H1:exists p : nat, N = (2 * p)%nat \/ N = S (2 * p)
x:nat
H2:N = (2 * x)%nat \/ N = S (2 * x)
H3:N = (2 * x)%nat

sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S N) <= 0
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0
H1:exists p : nat, N = (2 * p)%nat \/ N = S (2 * p)
x:nat
H2:N = (2 * x)%nat \/ N = S (2 * x)
H3:N = S (2 * x)
sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S N) <= 0
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0
H1:exists p : nat, N = (2 * p)%nat \/ N = S (2 * p)
x:nat
H2:N = (2 * x)%nat \/ N = S (2 * x)
H3:N = S (2 * x)

sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S N) <= 0
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0
H1:exists p : nat, N = (2 * p)%nat \/ N = S (2 * p)
x:nat
H2:N = (2 * x)%nat \/ N = S (2 * x)
H3:N = S (2 * x)

sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * x)) + tg_alt Un (S (S (S (2 * x)))) <= 0
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0
H1:exists p : nat, N = (2 * p)%nat \/ N = S (2 * p)
x:nat
H2:N = (2 * x)%nat \/ N = S (2 * x)
H3:N = S (2 * x)

sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * x)) + tg_alt Un (S (S (S (2 * x)))) <= sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * x))
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0
H1:exists p : nat, N = (2 * p)%nat \/ N = S (2 * p)
x:nat
H2:N = (2 * x)%nat \/ N = S (2 * x)
H3:N = S (2 * x)
sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * x)) <= 0
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0
H1:exists p : nat, N = (2 * p)%nat \/ N = S (2 * p)
x:nat
H2:N = (2 * x)%nat \/ N = S (2 * x)
H3:N = S (2 * x)

sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * x)) + tg_alt Un (S (S (S (2 * x)))) <= sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * x)) + 0
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0
H1:exists p : nat, N = (2 * p)%nat \/ N = S (2 * p)
x:nat
H2:N = (2 * x)%nat \/ N = S (2 * x)
H3:N = S (2 * x)
sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * x)) <= 0
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0
H1:exists p : nat, N = (2 * p)%nat \/ N = S (2 * p)
x:nat
H2:N = (2 * x)%nat \/ N = S (2 * x)
H3:N = S (2 * x)

tg_alt Un (S (S (S (2 * x)))) <= 0
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0
H1:exists p : nat, N = (2 * p)%nat \/ N = S (2 * p)
x:nat
H2:N = (2 * x)%nat \/ N = S (2 * x)
H3:N = S (2 * x)
sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * x)) <= 0
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0
H1:exists p : nat, N = (2 * p)%nat \/ N = S (2 * p)
x:nat
H2:N = (2 * x)%nat \/ N = S (2 * x)
H3:N = S (2 * x)

-1 * (-1 * (-1 * (-1) ^ (x + (x + 0)))) * Un (S (S (S (x + (x + 0))))) <= 0
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0
H1:exists p : nat, N = (2 * p)%nat \/ N = S (2 * p)
x:nat
H2:N = (2 * x)%nat \/ N = S (2 * x)
H3:N = S (2 * x)
sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * x)) <= 0
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0
H1:exists p : nat, N = (2 * p)%nat \/ N = S (2 * p)
x:nat
H2:N = (2 * x)%nat \/ N = S (2 * x)
H3:N = S (2 * x)

-1 * (-1 * (-1 * (-1) ^ (2 * x))) * Un (S (S (S (2 * x)))) <= 0
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0
H1:exists p : nat, N = (2 * p)%nat \/ N = S (2 * p)
x:nat
H2:N = (2 * x)%nat \/ N = S (2 * x)
H3:N = S (2 * x)
sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * x)) <= 0
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0
H1:exists p : nat, N = (2 * p)%nat \/ N = S (2 * p)
x:nat
H2:N = (2 * x)%nat \/ N = S (2 * x)
H3:N = S (2 * x)

-1 * (-1 * (-1 * 1)) * Un (S (S (S (2 * x)))) <= 0
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0
H1:exists p : nat, N = (2 * p)%nat \/ N = S (2 * p)
x:nat
H2:N = (2 * x)%nat \/ N = S (2 * x)
H3:N = S (2 * x)
sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * x)) <= 0
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0
H1:exists p : nat, N = (2 * p)%nat \/ N = S (2 * p)
x:nat
H2:N = (2 * x)%nat \/ N = S (2 * x)
H3:N = S (2 * x)

- Un (S (S (S (2 * x)))) <= 0
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0
H1:exists p : nat, N = (2 * p)%nat \/ N = S (2 * p)
x:nat
H2:N = (2 * x)%nat \/ N = S (2 * x)
H3:N = S (2 * x)
sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * x)) <= 0
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0
H1:exists p : nat, N = (2 * p)%nat \/ N = S (2 * p)
x:nat
H2:N = (2 * x)%nat \/ N = S (2 * x)
H3:N = S (2 * x)

Un (S (S (S (2 * x)))) + - Un (S (S (S (2 * x)))) <= Un (S (S (S (2 * x)))) + 0
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0
H1:exists p : nat, N = (2 * p)%nat \/ N = S (2 * p)
x:nat
H2:N = (2 * x)%nat \/ N = S (2 * x)
H3:N = S (2 * x)
sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * x)) <= 0
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0
H1:exists p : nat, N = (2 * p)%nat \/ N = S (2 * p)
x:nat
H2:N = (2 * x)%nat \/ N = S (2 * x)
H3:N = S (2 * x)

0 <= Un (S (S (S (2 * x))))
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0
H1:exists p : nat, N = (2 * p)%nat \/ N = S (2 * p)
x:nat
H2:N = (2 * x)%nat \/ N = S (2 * x)
H3:N = S (2 * x)
sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * x)) <= 0
Un:nat -> R
N:nat
H:Un_decreasing Un
H0:positivity_seq Un
HrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0
H1:exists p : nat, N = (2 * p)%nat \/ N = S (2 * p)
x:nat
H2:N = (2 * x)%nat \/ N = S (2 * x)
H3:N = S (2 * x)

sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * x)) <= 0
apply CV_ALT_step2; assumption. Qed. (**********)

forall Un : nat -> R, Un_decreasing Un -> positivity_seq Un -> has_ub (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N)))

forall Un : nat -> R, Un_decreasing Un -> positivity_seq Un -> has_ub (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N)))
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un

exists m : R, is_upper_bound (EUn (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N)))) m
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un

is_upper_bound (EUn (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N)))) (Un 0%nat)
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
x:R
H1:EUn (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N))) x
x0:nat
H2:x = sum_f_R0 (tg_alt Un) (S (2 * x0))

x <= Un 0%nat
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
x:R
H1:EUn (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N))) x
x0:nat
H2:x = sum_f_R0 (tg_alt Un) (S (2 * x0))

tg_alt Un 0 + sum_f_R0 (fun i : nat => tg_alt Un (S i)) (Init.Nat.pred (S (2 * x0))) <= Un 0%nat
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
x:R
H1:EUn (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N))) x
x0:nat
H2:x = sum_f_R0 (tg_alt Un) (S (2 * x0))
(0 < S (2 * x0))%nat
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
x:R
H1:EUn (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N))) x
x0:nat
H2:x = sum_f_R0 (tg_alt Un) (S (2 * x0))

Un 0%nat + sum_f_R0 (fun i : nat => tg_alt Un (S i)) (Init.Nat.pred (S (2 * x0))) <= Un 0%nat
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
x:R
H1:EUn (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N))) x
x0:nat
H2:x = sum_f_R0 (tg_alt Un) (S (2 * x0))
Un 0%nat = tg_alt Un 0
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
x:R
H1:EUn (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N))) x
x0:nat
H2:x = sum_f_R0 (tg_alt Un) (S (2 * x0))
(0 < S (2 * x0))%nat
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
x:R
H1:EUn (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N))) x
x0:nat
H2:x = sum_f_R0 (tg_alt Un) (S (2 * x0))

Un 0%nat + sum_f_R0 (fun i : nat => tg_alt Un (S i)) (Init.Nat.pred (S (2 * x0))) <= Un 0%nat + 0
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
x:R
H1:EUn (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N))) x
x0:nat
H2:x = sum_f_R0 (tg_alt Un) (S (2 * x0))
Un 0%nat = tg_alt Un 0
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
x:R
H1:EUn (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N))) x
x0:nat
H2:x = sum_f_R0 (tg_alt Un) (S (2 * x0))
(0 < S (2 * x0))%nat
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
x:R
H1:EUn (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N))) x
x0:nat
H2:x = sum_f_R0 (tg_alt Un) (S (2 * x0))

sum_f_R0 (fun i : nat => tg_alt Un (S i)) (Init.Nat.pred (S (2 * x0))) <= 0
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
x:R
H1:EUn (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N))) x
x0:nat
H2:x = sum_f_R0 (tg_alt Un) (S (2 * x0))
Un 0%nat = tg_alt Un 0
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
x:R
H1:EUn (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N))) x
x0:nat
H2:x = sum_f_R0 (tg_alt Un) (S (2 * x0))
(0 < S (2 * x0))%nat
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
x:R
H1:EUn (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N))) x
x0:nat
H2:x = sum_f_R0 (tg_alt Un) (S (2 * x0))

Un 0%nat = tg_alt Un 0
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
x:R
H1:EUn (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N))) x
x0:nat
H2:x = sum_f_R0 (tg_alt Un) (S (2 * x0))
(0 < S (2 * x0))%nat
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
x:R
H1:EUn (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N))) x
x0:nat
H2:x = sum_f_R0 (tg_alt Un) (S (2 * x0))

(0 < S (2 * x0))%nat
apply lt_O_Sn. Qed.
This lemma gives an interesting result about alternated series

forall Un : nat -> R, Un_decreasing Un -> positivity_seq Un -> Un_cv Un 0 -> {l : R | Un_cv (fun N : nat => sum_f_R0 (tg_alt Un) N) l}

forall Un : nat -> R, Un_decreasing Un -> positivity_seq Un -> Un_cv Un 0 -> {l : R | Un_cv (fun N : nat => sum_f_R0 (tg_alt Un) N) l}
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:Un_cv Un 0

{l : R | Un_cv (fun N : nat => sum_f_R0 (tg_alt Un) N) l}
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:Un_cv Un 0
H2:Un_growing (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N)))

{l : R | Un_cv (fun N : nat => sum_f_R0 (tg_alt Un) N) l}
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:Un_cv Un 0
H2:Un_growing (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N)))
H3:has_ub (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N)))

{l : R | Un_cv (fun N : nat => sum_f_R0 (tg_alt Un) N) l}
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:Un_cv Un 0
H2:Un_growing (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N)))
H3:has_ub (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N)))
x:R
p:Un_cv (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N))) x

{l : R | Un_cv (fun N : nat => sum_f_R0 (tg_alt Un) N) l}
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:Un_cv Un 0
H2:Un_growing (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N)))
H3:has_ub (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N)))
x:R
p:Un_cv (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N))) x

Un_cv (fun N : nat => sum_f_R0 (tg_alt Un) N) x
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (Un n - 0) < eps
H2:Un_growing (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N)))
H3:has_ub (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N)))
x:R
p:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n)) - x) < eps

forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (sum_f_R0 (tg_alt Un) n - x) < eps
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (Un n - 0) < eps0
H2:Un_growing (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N)))
H3:has_ub (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2

exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (sum_f_R0 (tg_alt Un) n - x) < eps
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (Un n - 0) < eps0
H2:Un_growing (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N)))
H3:has_ub (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n : nat, (n >= N2)%nat -> Rabs (Un n - 0) < eps / 2

exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (sum_f_R0 (tg_alt Un) n - x) < eps
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (Un n - 0) < eps0
H2:Un_growing (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N)))
H3:has_ub (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n : nat, (n >= N2)%nat -> Rabs (Un n - 0) < eps / 2
N1:nat
H7:forall n : nat, (n >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n)) - x) < eps / 2

exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (sum_f_R0 (tg_alt Un) n - x) < eps
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n : nat, (n >= N0)%nat -> Rabs (Un n - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n : nat, (n >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n : nat, (n >= N2)%nat -> Rabs (Un n - 0) < eps / 2
N1:nat
H7:forall n : nat, (n >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat

exists N0 : nat, forall n : nat, (n >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n - x) < eps
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat

Rabs (sum_f_R0 (tg_alt Un) n - x) < eps
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)

Rabs (sum_f_R0 (tg_alt Un) n - x) < eps
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)

Rabs (sum_f_R0 (tg_alt Un) n - x) < eps
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)

(N1 <= P)%nat -> Rabs (sum_f_R0 (tg_alt Un) n - x) < eps
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
(N1 <= P)%nat
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:(N1 <= P)%nat
H12:n = (2 * P)%nat

Rabs (sum_f_R0 (tg_alt Un) n - x) < eps
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:(N1 <= P)%nat
H12:n = S (2 * P)
Rabs (sum_f_R0 (tg_alt Un) n - x) < eps
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
(N1 <= P)%nat
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:(N1 <= P)%nat
H12:n = (2 * P)%nat

Rabs (sum_f_R0 (tg_alt Un) (S n) - x + - tg_alt Un (S n)) < eps
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:(N1 <= P)%nat
H12:n = (2 * P)%nat
sum_f_R0 (tg_alt Un) (S n) - x + - tg_alt Un (S n) = sum_f_R0 (tg_alt Un) n - x
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:(N1 <= P)%nat
H12:n = S (2 * P)
Rabs (sum_f_R0 (tg_alt Un) n - x) < eps
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
(N1 <= P)%nat
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:(N1 <= P)%nat
H12:n = (2 * P)%nat

Rabs (sum_f_R0 (tg_alt Un) (S n) - x + - tg_alt Un (S n)) <= Rabs (sum_f_R0 (tg_alt Un) (S n) - x) + Rabs (- tg_alt Un (S n))
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:(N1 <= P)%nat
H12:n = (2 * P)%nat
Rabs (sum_f_R0 (tg_alt Un) (S n) - x) + Rabs (- tg_alt Un (S n)) < eps
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:(N1 <= P)%nat
H12:n = (2 * P)%nat
sum_f_R0 (tg_alt Un) (S n) - x + - tg_alt Un (S n) = sum_f_R0 (tg_alt Un) n - x
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:(N1 <= P)%nat
H12:n = S (2 * P)
Rabs (sum_f_R0 (tg_alt Un) n - x) < eps
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
(N1 <= P)%nat
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:(N1 <= P)%nat
H12:n = (2 * P)%nat

Rabs (sum_f_R0 (tg_alt Un) (S n) - x) + Rabs (- tg_alt Un (S n)) < eps
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:(N1 <= P)%nat
H12:n = (2 * P)%nat
sum_f_R0 (tg_alt Un) (S n) - x + - tg_alt Un (S n) = sum_f_R0 (tg_alt Un) n - x
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:(N1 <= P)%nat
H12:n = S (2 * P)
Rabs (sum_f_R0 (tg_alt Un) n - x) < eps
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
(N1 <= P)%nat
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:(N1 <= P)%nat
H12:n = (2 * P)%nat

Rabs (sum_f_R0 (tg_alt Un) (S n) - x) < eps / 2
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:(N1 <= P)%nat
H12:n = (2 * P)%nat
Rabs (- tg_alt Un (S n)) < eps / 2
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:(N1 <= P)%nat
H12:n = (2 * P)%nat
sum_f_R0 (tg_alt Un) (S n) - x + - tg_alt Un (S n) = sum_f_R0 (tg_alt Un) n - x
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:(N1 <= P)%nat
H12:n = S (2 * P)
Rabs (sum_f_R0 (tg_alt Un) n - x) < eps
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
(N1 <= P)%nat
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:(N1 <= P)%nat
H12:n = (2 * P)%nat

Rabs (- tg_alt Un (S n)) < eps / 2
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:(N1 <= P)%nat
H12:n = (2 * P)%nat
sum_f_R0 (tg_alt Un) (S n) - x + - tg_alt Un (S n) = sum_f_R0 (tg_alt Un) n - x
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:(N1 <= P)%nat
H12:n = S (2 * P)
Rabs (sum_f_R0 (tg_alt Un) n - x) < eps
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
(N1 <= P)%nat
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 + 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:(N1 <= P)%nat
H12:n = (2 * P)%nat

(S n >= N2)%nat
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:(N1 <= P)%nat
H12:n = (2 * P)%nat
sum_f_R0 (tg_alt Un) (S n) - x + - tg_alt Un (S n) = sum_f_R0 (tg_alt Un) n - x
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:(N1 <= P)%nat
H12:n = S (2 * P)
Rabs (sum_f_R0 (tg_alt Un) n - x) < eps
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
(N1 <= P)%nat
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 + 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:(N1 <= P)%nat
H12:n = (2 * P)%nat

(N2 <= n)%nat
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 + 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:(N1 <= P)%nat
H12:n = (2 * P)%nat
(n <= S n)%nat
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:(N1 <= P)%nat
H12:n = (2 * P)%nat
sum_f_R0 (tg_alt Un) (S n) - x + - tg_alt Un (S n) = sum_f_R0 (tg_alt Un) n - x
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:(N1 <= P)%nat
H12:n = S (2 * P)
Rabs (sum_f_R0 (tg_alt Un) n - x) < eps
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
(N1 <= P)%nat
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 + 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:(N1 <= P)%nat
H12:n = (2 * P)%nat

(n <= S n)%nat
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:(N1 <= P)%nat
H12:n = (2 * P)%nat
sum_f_R0 (tg_alt Un) (S n) - x + - tg_alt Un (S n) = sum_f_R0 (tg_alt Un) n - x
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:(N1 <= P)%nat
H12:n = S (2 * P)
Rabs (sum_f_R0 (tg_alt Un) n - x) < eps
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
(N1 <= P)%nat
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:(N1 <= P)%nat
H12:n = (2 * P)%nat

sum_f_R0 (tg_alt Un) (S n) - x + - tg_alt Un (S n) = sum_f_R0 (tg_alt Un) n - x
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:(N1 <= P)%nat
H12:n = S (2 * P)
Rabs (sum_f_R0 (tg_alt Un) n - x) < eps
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
(N1 <= P)%nat
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:(N1 <= P)%nat
H12:n = S (2 * P)

Rabs (sum_f_R0 (tg_alt Un) n - x) < eps
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
(N1 <= P)%nat
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:(N1 <= P)%nat
H12:n = S (2 * P)

Rabs (sum_f_R0 (tg_alt Un) (S (2 * P)) - x) < eps / 2
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:(N1 <= P)%nat
H12:n = S (2 * P)
eps / 2 < eps
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
(N1 <= P)%nat
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:(N1 <= P)%nat
H12:n = S (2 * P)

eps / 2 < eps
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
(N1 <= P)%nat
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:(N1 <= P)%nat
H12:n = S (2 * P)

0 < 2
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:(N1 <= P)%nat
H12:n = S (2 * P)
2 * (eps * / 2) < 2 * eps
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
(N1 <= P)%nat
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:(N1 <= P)%nat
H12:n = S (2 * P)

2 * (eps * / 2) < 2 * eps
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
(N1 <= P)%nat
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:(N1 <= P)%nat
H12:n = S (2 * P)

eps < 2 * eps
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
(N1 <= P)%nat
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:(N1 <= P)%nat
H12:n = S (2 * P)

eps < eps + eps
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
(N1 <= P)%nat
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)

(N1 <= P)%nat
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:n = (2 * P)%nat

(2 * N1 <= 2 * P)%nat
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:n = S (2 * P)
(2 * N1 <= 2 * P)%nat
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:n = (2 * P)%nat

(2 * N1 <= N)%nat
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:n = (2 * P)%nat
(N <= n)%nat
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:n = S (2 * P)
(2 * N1 <= 2 * P)%nat
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:n = (2 * P)%nat

(N <= n)%nat
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:n = S (2 * P)
(2 * N1 <= 2 * P)%nat
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:n = S (2 * P)

(2 * N1 <= 2 * P)%nat
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:n = S (2 * P)

(2 * N1 < S (2 * P))%nat
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:n = S (2 * P)

(2 * N1 < n)%nat
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:n = S (2 * P)

(2 * N1 < N)%nat
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:n = S (2 * P)
(N <= n)%nat
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:n = S (2 * P)

(2 * N1 < S (2 * N1))%nat
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:n = S (2 * P)
(S (2 * N1) <= Nat.max (S (2 * N1)) N2)%nat
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:n = S (2 * P)
(N <= n)%nat
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:n = S (2 * P)

(S (2 * N1) <= Nat.max (S (2 * N1)) N2)%nat
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:n = S (2 * P)
(N <= n)%nat
Un:nat -> R
H:Un_decreasing Un
H0:positivity_seq Un
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0
H2:Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
H3:has_ub (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
x:R
p:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps0
eps:R
H4:eps > 0
H5:0 < eps / 2
N2:nat
H6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2
N1:nat
H7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2
N:=Nat.max (S (2 * N1)) N2:nat
n:nat
H8:(n >= N)%nat
H9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)
P:nat
H10:n = (2 * P)%nat \/ n = S (2 * P)
H11:n = S (2 * P)

(N <= n)%nat
assumption. Qed. (*************************************************)

Convergence of alternated series


forall Un : nat -> R, Un_decreasing Un -> Un_cv Un 0 -> {l : R | Un_cv (fun N : nat => sum_f_R0 (tg_alt Un) N) l}

forall Un : nat -> R, Un_decreasing Un -> Un_cv Un 0 -> {l : R | Un_cv (fun N : nat => sum_f_R0 (tg_alt Un) N) l}
Un:nat -> R
H:Un_decreasing Un
H0:Un_cv Un 0

Un_decreasing Un
Un:nat -> R
H:Un_decreasing Un
H0:Un_cv Un 0
positivity_seq Un
Un:nat -> R
H:Un_decreasing Un
H0:Un_cv Un 0
Un_cv Un 0
Un:nat -> R
H:Un_decreasing Un
H0:Un_cv Un 0

positivity_seq Un
Un:nat -> R
H:Un_decreasing Un
H0:Un_cv Un 0
Un_cv Un 0
Un:nat -> R
H:Un_decreasing Un
H0:Un_cv Un 0

Un_cv Un 0
assumption. Qed.

forall (Un : nat -> R) (l : R) (N : nat), Un_decreasing Un -> Un_cv Un 0 -> Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) l -> sum_f_R0 (tg_alt Un) (S (2 * N)) <= l <= sum_f_R0 (tg_alt Un) (2 * N)

forall (Un : nat -> R) (l : R) (N : nat), Un_decreasing Un -> Un_cv Un 0 -> Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) l -> sum_f_R0 (tg_alt Un) (S (2 * N)) <= l <= sum_f_R0 (tg_alt Un) (2 * N)
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) l

sum_f_R0 (tg_alt Un) (S (2 * N)) <= l <= sum_f_R0 (tg_alt Un) (2 * N)
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) l

Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) l -> sum_f_R0 (tg_alt Un) (S (2 * N)) <= l <= sum_f_R0 (tg_alt Un) (2 * N)
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) l
Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) l
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) l

Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0))) l -> Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) l -> sum_f_R0 (tg_alt Un) (S (2 * N)) <= l <= sum_f_R0 (tg_alt Un) (2 * N)
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) l
Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0))) l
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) l
Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) l
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) l
H2:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0))) l
H3:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) l

sum_f_R0 (tg_alt Un) (S (2 * N)) <= l
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) l
H2:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0))) l
H3:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) l
l <= sum_f_R0 (tg_alt Un) (2 * N)
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) l
Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0))) l
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) l
Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) l
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) l
H2:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0))) l
H3:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) l

Un_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) l
H2:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0))) l
H3:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) l
Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0))) l
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) l
H2:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0))) l
H3:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) l
l <= sum_f_R0 (tg_alt Un) (2 * N)
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) l
Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0))) l
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) l
Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) l
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) l
H2:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0))) l
H3:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) l

Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0))) l
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) l
H2:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0))) l
H3:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) l
l <= sum_f_R0 (tg_alt Un) (2 * N)
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) l
Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0))) l
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) l
Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) l
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) l
H2:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0))) l
H3:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) l

l <= sum_f_R0 (tg_alt Un) (2 * N)
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) l
Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0))) l
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) l
Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) l
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) l
H2:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0))) l
H3:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) l

Un_decreasing (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0))
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) l
H2:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0))) l
H3:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) l
Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) l
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) l
Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0))) l
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) l
Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) l
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) l
H2:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0))) l
H3:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) l

Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) l
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) l
Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0))) l
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) l
Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) l
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) l

Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0))) l
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) l
Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) l
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n : nat, (n >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n - l) < eps0
eps:R
H2:eps > 0

exists N0 : nat, forall n : nat, (n >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n)) - l) < eps
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) l
Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) l
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n : nat, (n >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n - l) < eps0
eps:R
H2:eps > 0
x:nat
H3:forall n : nat, (n >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n - l) < eps

exists N0 : nat, forall n : nat, (n >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n)) - l) < eps
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) l
Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) l
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps0
eps:R
H2:eps > 0
x:nat
H3:forall n0 : nat, (n0 >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps
n:nat
H4:(n >= x)%nat

Rabs (sum_f_R0 (tg_alt Un) (S (2 * n)) - l) < eps
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) l
Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) l
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps0
eps:R
H2:eps > 0
x:nat
H3:forall n0 : nat, (n0 >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps
n:nat
H4:(n >= x)%nat

(S (2 * n) >= x)%nat
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) l
Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) l
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps0
eps:R
H2:eps > 0
x:nat
H3:forall n0 : nat, (n0 >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps
n:nat
H4:(n >= x)%nat

(x <= 2 * n)%nat
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps0
eps:R
H2:eps > 0
x:nat
H3:forall n0 : nat, (n0 >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps
n:nat
H4:(n >= x)%nat
(2 * n <= S (2 * n))%nat
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) l
Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) l
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps0
eps:R
H2:eps > 0
x:nat
H3:forall n0 : nat, (n0 >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps
n:nat
H4:(n >= x)%nat

(x <= n)%nat
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps0
eps:R
H2:eps > 0
x:nat
H3:forall n0 : nat, (n0 >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps
n:nat
H4:(n >= x)%nat
(n <= 2 * n)%nat
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps0
eps:R
H2:eps > 0
x:nat
H3:forall n0 : nat, (n0 >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps
n:nat
H4:(n >= x)%nat
(2 * n <= S (2 * n))%nat
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) l
Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) l
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps0
eps:R
H2:eps > 0
x:nat
H3:forall n0 : nat, (n0 >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps
n:nat
H4:(n >= x)%nat

(n <= 2 * n)%nat
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps0
eps:R
H2:eps > 0
x:nat
H3:forall n0 : nat, (n0 >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps
n:nat
H4:(n >= x)%nat
(2 * n <= S (2 * n))%nat
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) l
Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) l
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps0
eps:R
H2:eps > 0
x:nat
H3:forall n0 : nat, (n0 >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps
n:nat
H4:(n >= x)%nat
H5:2%nat = 0%nat \/ (n <= 2 * n)%nat

(n <= 2 * n)%nat
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps0
eps:R
H2:eps > 0
x:nat
H3:forall n0 : nat, (n0 >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps
n:nat
H4:(n >= x)%nat
(2 * n <= S (2 * n))%nat
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) l
Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) l
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps0
eps:R
H2:eps > 0
x:nat
H3:forall n0 : nat, (n0 >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps
n:nat
H4:(n >= x)%nat
H5:2%nat = 0%nat \/ (n <= 2 * n)%nat
H6:2%nat = 0%nat

(n <= 2 * n)%nat
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps0
eps:R
H2:eps > 0
x:nat
H3:forall n0 : nat, (n0 >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps
n:nat
H4:(n >= x)%nat
H5:2%nat = 0%nat \/ (n <= 2 * n)%nat
H6:(n <= 2 * n)%nat
(n <= 2 * n)%nat
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps0
eps:R
H2:eps > 0
x:nat
H3:forall n0 : nat, (n0 >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps
n:nat
H4:(n >= x)%nat
(2 * n <= S (2 * n))%nat
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) l
Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) l
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps0
eps:R
H2:eps > 0
x:nat
H3:forall n0 : nat, (n0 >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps
n:nat
H4:(n >= x)%nat
H5:2%nat = 0%nat \/ (n <= 2 * n)%nat
H6:(n <= 2 * n)%nat

(n <= 2 * n)%nat
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps0
eps:R
H2:eps > 0
x:nat
H3:forall n0 : nat, (n0 >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps
n:nat
H4:(n >= x)%nat
(2 * n <= S (2 * n))%nat
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) l
Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) l
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps0
eps:R
H2:eps > 0
x:nat
H3:forall n0 : nat, (n0 >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps
n:nat
H4:(n >= x)%nat

(2 * n <= S (2 * n))%nat
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) l
Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) l
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) l

Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) l
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n : nat, (n >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n - l) < eps0
eps:R
H2:eps > 0

exists N0 : nat, forall n : nat, (n >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (2 * n) - l) < eps
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n : nat, (n >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n - l) < eps0
eps:R
H2:eps > 0
x:nat
H3:forall n : nat, (n >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n - l) < eps

exists N0 : nat, forall n : nat, (n >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (2 * n) - l) < eps
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps0
eps:R
H2:eps > 0
x:nat
H3:forall n0 : nat, (n0 >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps
n:nat
H4:(n >= x)%nat

Rabs (sum_f_R0 (tg_alt Un) (2 * n) - l) < eps
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps0
eps:R
H2:eps > 0
x:nat
H3:forall n0 : nat, (n0 >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps
n:nat
H4:(n >= x)%nat

(2 * n >= x)%nat
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps0
eps:R
H2:eps > 0
x:nat
H3:forall n0 : nat, (n0 >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps
n:nat
H4:(n >= x)%nat

(x <= n)%nat
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps0
eps:R
H2:eps > 0
x:nat
H3:forall n0 : nat, (n0 >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps
n:nat
H4:(n >= x)%nat
(n <= 2 * n)%nat
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps0
eps:R
H2:eps > 0
x:nat
H3:forall n0 : nat, (n0 >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps
n:nat
H4:(n >= x)%nat

(n <= 2 * n)%nat
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps0
eps:R
H2:eps > 0
x:nat
H3:forall n0 : nat, (n0 >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps
n:nat
H4:(n >= x)%nat
H5:2%nat = 0%nat \/ (n <= 2 * n)%nat

(n <= 2 * n)%nat
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps0
eps:R
H2:eps > 0
x:nat
H3:forall n0 : nat, (n0 >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps
n:nat
H4:(n >= x)%nat
H5:2%nat = 0%nat \/ (n <= 2 * n)%nat
H6:2%nat = 0%nat

(n <= 2 * n)%nat
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps0
eps:R
H2:eps > 0
x:nat
H3:forall n0 : nat, (n0 >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps
n:nat
H4:(n >= x)%nat
H5:2%nat = 0%nat \/ (n <= 2 * n)%nat
H6:(n <= 2 * n)%nat
(n <= 2 * n)%nat
Un:nat -> R
l:R
N:nat
H:Un_decreasing Un
H0:Un_cv Un 0
H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps0
eps:R
H2:eps > 0
x:nat
H3:forall n0 : nat, (n0 >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps
n:nat
H4:(n >= x)%nat
H5:2%nat = 0%nat \/ (n <= 2 * n)%nat
H6:(n <= 2 * n)%nat

(n <= 2 * n)%nat
assumption. Qed. (***************************************)

Application : construction of PI

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

Definition PI_tg (n:nat) := / INR (2 * n + 1).


forall n : nat, 0 <= PI_tg n

forall n : nat, 0 <= PI_tg n
intro; unfold PI_tg; left; apply Rinv_0_lt_compat; apply lt_INR_0; replace (2 * n + 1)%nat with (S (2 * n)); [ apply lt_O_Sn | ring ]. Qed.

Un_decreasing PI_tg

Un_decreasing PI_tg
n:nat

/ INR (2 * S n + 1) <= / INR (2 * n + 1)
n:nat

0 < INR (2 * n + 1)
n:nat
INR (2 * n + 1) * / INR (2 * S n + 1) <= INR (2 * n + 1) * / INR (2 * n + 1)
n:nat

(0 < 2 * n + 1)%nat
n:nat
INR (2 * n + 1) * / INR (2 * S n + 1) <= INR (2 * n + 1) * / INR (2 * n + 1)
n:nat

INR (2 * n + 1) * / INR (2 * S n + 1) <= INR (2 * n + 1) * / INR (2 * n + 1)
n:nat

INR (2 * n + 1) * / INR (2 * S n + 1) <= 1
n:nat
INR (2 * n + 1) <> 0
n:nat

0 < INR (2 * S n + 1)
n:nat
INR (2 * S n + 1) * (INR (2 * n + 1) * / INR (2 * S n + 1)) <= INR (2 * S n + 1) * 1
n:nat
INR (2 * n + 1) <> 0
n:nat

(0 < 2 * S n + 1)%nat
n:nat
INR (2 * S n + 1) * (INR (2 * n + 1) * / INR (2 * S n + 1)) <= INR (2 * S n + 1) * 1
n:nat
INR (2 * n + 1) <> 0
n:nat

INR (2 * S n + 1) * (INR (2 * n + 1) * / INR (2 * S n + 1)) <= INR (2 * S n + 1) * 1
n:nat
INR (2 * n + 1) <> 0
n:nat

INR (2 * n + 1) * 1 <= INR (2 * S n + 1) * 1
n:nat
INR (2 * S n + 1) <> 0
n:nat
INR (2 * n + 1) <> 0
n:nat

(2 * n + 1 <= 2 * S n + 1)%nat
n:nat
INR (2 * S n + 1) <> 0
n:nat
INR (2 * n + 1) <> 0
n:nat

(2 * n + 1 <= S (S (2 * n + 1)))%nat
n:nat
S (S (2 * n + 1)) = (2 * S n + 1)%nat
n:nat
INR (2 * S n + 1) <> 0
n:nat
INR (2 * n + 1) <> 0
n:nat

S (S (2 * n + 1)) = (2 * S n + 1)%nat
n:nat
INR (2 * S n + 1) <> 0
n:nat
INR (2 * n + 1) <> 0
n:nat

INR (2 * S n + 1) <> 0
n:nat
INR (2 * n + 1) <> 0
n:nat

INR (2 * n + 1) <> 0
apply not_O_INR; replace (2 * n + 1)%nat with (S (2 * n)); [ discriminate | ring ]. Qed.

Un_cv PI_tg 0

Un_cv PI_tg 0
eps:R
H:eps > 0

exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (PI_tg n - 0) < eps
eps:R
H:eps > 0
H0:0 < 2 * eps

exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (PI_tg n - 0) < eps
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1

exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (PI_tg n - 0) < eps
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1

(0 <= up (/ (2 * eps)))%Z -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (PI_tg n - 0) < eps
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m

exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (PI_tg n - 0) < eps
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N

exists N0 : nat, forall n : nat, (n >= N0)%nat -> Rabs (PI_tg n - 0) < eps
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N

(0 < N)%nat -> exists N0 : nat, forall n : nat, (n >= N0)%nat -> Rabs (PI_tg n - 0) < eps
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
(0 < N)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat

Rabs (PI_tg n - 0) < eps
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
(0 < N)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat

(0 < n)%nat -> Rabs (PI_tg n - 0) < eps
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
(0 < n)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
(0 < N)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat

PI_tg n < eps
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
PI_tg n >= 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
(0 < n)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
(0 < N)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat

/ INR (2 * n + 1) < / INR (2 * n)
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
/ INR (2 * n) < eps
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
PI_tg n >= 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
(0 < n)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
(0 < N)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat

0 < INR (2 * n)
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
INR (2 * n) * / INR (2 * n + 1) < INR (2 * n) * / INR (2 * n)
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
/ INR (2 * n) < eps
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
PI_tg n >= 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
(0 < n)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
(0 < N)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat

(0 < 2 * n)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
INR (2 * n) * / INR (2 * n + 1) < INR (2 * n) * / INR (2 * n)
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
/ INR (2 * n) < eps
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
PI_tg n >= 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
(0 < n)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
(0 < N)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat

(0 < n + n)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
INR (2 * n) * / INR (2 * n + 1) < INR (2 * n) * / INR (2 * n)
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
/ INR (2 * n) < eps
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
PI_tg n >= 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
(0 < n)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
(0 < N)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat

(0 < n)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
(n <= n + n)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
INR (2 * n) * / INR (2 * n + 1) < INR (2 * n) * / INR (2 * n)
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
/ INR (2 * n) < eps
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
PI_tg n >= 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
(0 < n)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
(0 < N)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat

(n <= n + n)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
INR (2 * n) * / INR (2 * n + 1) < INR (2 * n) * / INR (2 * n)
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
/ INR (2 * n) < eps
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
PI_tg n >= 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
(0 < n)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
(0 < N)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat

INR (2 * n) * / INR (2 * n + 1) < INR (2 * n) * / INR (2 * n)
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
/ INR (2 * n) < eps
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
PI_tg n >= 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
(0 < n)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
(0 < N)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat

INR (2 * n) * / INR (2 * n + 1) < 1
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
INR (2 * n) <> 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
/ INR (2 * n) < eps
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
PI_tg n >= 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
(0 < n)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
(0 < N)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat

0 < INR (2 * n + 1)
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
INR (2 * n + 1) * (INR (2 * n) * / INR (2 * n + 1)) < INR (2 * n + 1) * 1
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
INR (2 * n) <> 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
/ INR (2 * n) < eps
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
PI_tg n >= 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
(0 < n)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
(0 < N)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat

(0 < 2 * n + 1)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
INR (2 * n + 1) * (INR (2 * n) * / INR (2 * n + 1)) < INR (2 * n + 1) * 1
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
INR (2 * n) <> 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
/ INR (2 * n) < eps
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
PI_tg n >= 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
(0 < n)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
(0 < N)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat

INR (2 * n + 1) * (INR (2 * n) * / INR (2 * n + 1)) < INR (2 * n + 1) * 1
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
INR (2 * n) <> 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
/ INR (2 * n) < eps
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
PI_tg n >= 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
(0 < n)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
(0 < N)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat

INR (2 * n) * / INR (2 * n + 1) * INR (2 * n + 1) < INR (2 * n + 1) * 1
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
INR (2 * n) <> 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
/ INR (2 * n) < eps
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
PI_tg n >= 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
(0 < n)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
(0 < N)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat

INR (2 * n) * 1 < INR (2 * n + 1) * 1
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
INR (2 * n + 1) <> 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
INR (2 * n) <> 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
/ INR (2 * n) < eps
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
PI_tg n >= 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
(0 < n)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
(0 < N)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat

(2 * n < 2 * n + 1)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
INR (2 * n + 1) <> 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
INR (2 * n) <> 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
/ INR (2 * n) < eps
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
PI_tg n >= 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
(0 < n)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
(0 < N)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat

INR (2 * n + 1) <> 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
INR (2 * n) <> 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
/ INR (2 * n) < eps
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
PI_tg n >= 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
(0 < n)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
(0 < N)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat

INR (2 * n) <> 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
/ INR (2 * n) < eps
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
PI_tg n >= 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
(0 < n)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
(0 < N)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat

INR (2 * S (Init.Nat.pred n)) <> 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
S (Init.Nat.pred n) = n
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
/ INR (2 * n) < eps
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
PI_tg n >= 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
(0 < n)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
(0 < N)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat

S (Init.Nat.pred n) = n
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
/ INR (2 * n) < eps
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
PI_tg n >= 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
(0 < n)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
(0 < N)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat

(0 < n)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
/ INR (2 * n) < eps
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
PI_tg n >= 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
(0 < n)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
(0 < N)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat

/ INR (2 * n) < eps
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
PI_tg n >= 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
(0 < n)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
(0 < N)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat

/ INR (2 * n) <= / INR (2 * N)
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
/ INR (2 * N) < eps
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
PI_tg n >= 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
(0 < n)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
(0 < N)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat

0 < INR (2 * N)
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
INR (2 * N) <= INR (2 * n)
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
/ INR (2 * N) < eps
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
PI_tg n >= 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
(0 < n)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
(0 < N)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat

INR (2 * N) <= INR (2 * n)
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
/ INR (2 * N) < eps
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
PI_tg n >= 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
(0 < n)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
(0 < N)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat

(2 * N <= 2 * n)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
/ INR (2 * N) < eps
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
PI_tg n >= 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
(0 < n)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
(0 < N)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat

/ INR (2 * N) < eps
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
PI_tg n >= 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
(0 < n)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
(0 < N)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat

/ (INR 2 * INR N) < eps
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
PI_tg n >= 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
(0 < n)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
(0 < N)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat

0 < INR N / eps
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
INR N / eps * / (INR 2 * INR N) < INR N / eps * eps
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
PI_tg n >= 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
(0 < n)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
(0 < N)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat

0 < INR N
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
INR N / eps * / (INR 2 * INR N) < INR N / eps * eps
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
PI_tg n >= 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
(0 < n)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
(0 < N)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat

INR N / eps * / (INR 2 * INR N) < INR N / eps * eps
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
PI_tg n >= 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
(0 < n)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
(0 < N)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat

/ (2 * eps) < INR N / eps * eps
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
/ (2 * eps) = INR N / eps * / (INR 2 * INR N)
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
PI_tg n >= 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
(0 < n)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
(0 < N)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat

/ (2 * eps) < INR N
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
INR N = INR N / eps * eps
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
/ (2 * eps) = INR N / eps * / (INR 2 * INR N)
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
PI_tg n >= 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
(0 < n)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
(0 < N)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat

/ (2 * eps) < IZR (Z.of_nat N)
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
INR N = INR N / eps * eps
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
/ (2 * eps) = INR N / eps * / (INR 2 * INR N)
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
PI_tg n >= 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
(0 < n)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
(0 < N)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat

INR N = INR N / eps * eps
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
/ (2 * eps) = INR N / eps * / (INR 2 * INR N)
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
PI_tg n >= 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
(0 < n)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
(0 < N)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat

eps <> 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
/ (2 * eps) = INR N / eps * / (INR 2 * INR N)
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
PI_tg n >= 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
(0 < n)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
(0 < N)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat

/ (2 * eps) = INR N / eps * / (INR 2 * INR N)
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
PI_tg n >= 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
(0 < n)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
(0 < N)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat

INR N <> 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
eps <> 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
PI_tg n >= 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
(0 < n)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
(0 < N)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat

eps <> 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat
PI_tg n >= 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
(0 < n)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
(0 < N)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
H7:(0 < n)%nat

PI_tg n >= 0
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat
(0 < n)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
(0 < N)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:(0 < N)%nat
n:nat
H6:(n >= N)%nat

(0 < n)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
(0 < N)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N

(0 < N)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:IZR (up (/ (2 * eps))) > / (2 * eps)

(0 < N)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:IZR (up (/ (2 * eps))) > / (2 * eps)
l:(0 < N)%nat

(0 < N)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
H4:up (/ (2 * eps)) = Z.of_nat 0
H5:IZR (up (/ (2 * eps))) > / (2 * eps)
(0 < 0)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:IZR (up (/ (2 * eps))) > / (2 * eps)
H6:(N < 0)%nat
(0 < N)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
H4:up (/ (2 * eps)) = Z.of_nat 0
H5:IZR (up (/ (2 * eps))) > / (2 * eps)

(0 < 0)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:IZR (up (/ (2 * eps))) > / (2 * eps)
H6:(N < 0)%nat
(0 < N)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
H4:up (/ (2 * eps)) = Z.of_nat 0
H5:IZR (Z.of_nat 0) > / (2 * eps)

(0 < 0)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:IZR (up (/ (2 * eps))) > / (2 * eps)
H6:(N < 0)%nat
(0 < N)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
H4:up (/ (2 * eps)) = Z.of_nat 0
H5:0 > / (2 * eps)

(0 < 0)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:IZR (up (/ (2 * eps))) > / (2 * eps)
H6:(N < 0)%nat
(0 < N)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
H4:up (/ (2 * eps)) = Z.of_nat 0
H5:0 > / (2 * eps)
H6:0 < / (2 * eps)

(0 < 0)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:IZR (up (/ (2 * eps))) > / (2 * eps)
H6:(N < 0)%nat
(0 < N)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
H2:(0 <= up (/ (2 * eps)))%Z
H3:exists m : nat, up (/ (2 * eps)) = Z.of_nat m
N:nat
H4:up (/ (2 * eps)) = Z.of_nat N
H5:IZR (up (/ (2 * eps))) > / (2 * eps)
H6:(N < 0)%nat

(0 < N)%nat
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1

(0 <= up (/ (2 * eps)))%Z
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1

0 <= IZR (up (/ (2 * eps)))
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1

0 < / (2 * eps)
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1
/ (2 * eps) < IZR (up (/ (2 * eps)))
eps:R
H:eps > 0
H0:0 < 2 * eps
H1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1

/ (2 * eps) < IZR (up (/ (2 * eps)))
elim H1; intros; assumption. Qed.

{l : R | Un_cv (fun N : nat => sum_f_R0 (tg_alt PI_tg) N) l}

{l : R | Un_cv (fun N : nat => sum_f_R0 (tg_alt PI_tg) N) l}

Un_decreasing PI_tg

Un_cv PI_tg 0

Un_cv PI_tg 0
apply PI_tg_cv. Qed.
Now, PI is defined
Definition Alt_PI : R := 4 * (let (a,_) := exist_PI in a).
We can get an approximation of PI with the following inequality

forall N : nat, sum_f_R0 (tg_alt PI_tg) (S (2 * N)) <= Alt_PI / 4 <= sum_f_R0 (tg_alt PI_tg) (2 * N)

forall N : nat, sum_f_R0 (tg_alt PI_tg) (S (2 * N)) <= Alt_PI / 4 <= sum_f_R0 (tg_alt PI_tg) (2 * N)
N:nat

Un_decreasing PI_tg
N:nat
Un_cv PI_tg 0
N:nat
Un_cv (fun N0 : nat => sum_f_R0 (tg_alt PI_tg) N0) (Alt_PI / 4)
N:nat

Un_cv PI_tg 0
N:nat
Un_cv (fun N0 : nat => sum_f_R0 (tg_alt PI_tg) N0) (Alt_PI / 4)
N:nat

Un_cv (fun N0 : nat => sum_f_R0 (tg_alt PI_tg) N0) (Alt_PI / 4)
N:nat
x:R

Un_cv (fun N0 : nat => sum_f_R0 (tg_alt PI_tg) N0) x -> Un_cv (fun N0 : nat => sum_f_R0 (tg_alt PI_tg) N0) (4 * x / 4)
N:nat
x:R

Un_cv (fun N0 : nat => sum_f_R0 (tg_alt PI_tg) N0) x -> Un_cv (fun N0 : nat => sum_f_R0 (tg_alt PI_tg) N0) x
N:nat
x:R
x = 4 * x / 4
N:nat
x:R

x = 4 * x / 4
unfold Rdiv; rewrite (Rmult_comm 4); rewrite Rmult_assoc; rewrite <- Rinv_r_sym; [ rewrite Rmult_1_r; reflexivity | discrR ]. Qed.

0 < Alt_PI

0 < Alt_PI
H:sum_f_R0 (tg_alt PI_tg) (S (2 * 0)) <= Alt_PI / 4 <= sum_f_R0 (tg_alt PI_tg) (2 * 0)

0 < Alt_PI
H:sum_f_R0 (tg_alt PI_tg) (S (2 * 0)) <= Alt_PI / 4 <= sum_f_R0 (tg_alt PI_tg) (2 * 0)

0 < / 4
H:sum_f_R0 (tg_alt PI_tg) (S (2 * 0)) <= Alt_PI / 4 <= sum_f_R0 (tg_alt PI_tg) (2 * 0)
/ 4 * 0 < / 4 * Alt_PI
H:sum_f_R0 (tg_alt PI_tg) (S (2 * 0)) <= Alt_PI / 4 <= sum_f_R0 (tg_alt PI_tg) (2 * 0)

/ 4 * 0 < / 4 * Alt_PI
H:sum_f_R0 (tg_alt PI_tg) (S (2 * 0)) <= Alt_PI / 4 <= sum_f_R0 (tg_alt PI_tg) (2 * 0)

0 < Alt_PI * / 4
H:sum_f_R0 (tg_alt PI_tg) (S (2 * 0)) <= Alt_PI / 4

0 < Alt_PI * / 4
H:sum_f_R0 (tg_alt PI_tg) (S (2 * 0)) <= Alt_PI * / 4

0 < sum_f_R0 (tg_alt PI_tg) (S (2 * 0))
H:sum_f_R0 (tg_alt PI_tg) (S (2 * 0)) <= Alt_PI * / 4
sum_f_R0 (tg_alt PI_tg) (S (2 * 0)) <= Alt_PI * / 4
H:sum_f_R0 (tg_alt PI_tg) (S (2 * 0)) <= Alt_PI * / 4

PI_tg 1 + 0 < PI_tg 1 + (PI_tg 0 + -1 * PI_tg 1)
H:sum_f_R0 (tg_alt PI_tg) (S (2 * 0)) <= Alt_PI * / 4
sum_f_R0 (tg_alt PI_tg) (S (2 * 0)) <= Alt_PI * / 4
H:sum_f_R0 (tg_alt PI_tg) (S (2 * 0)) <= Alt_PI * / 4

/ INR (2 * 1 + 1) < / INR (2 * 0 + 1)
H:sum_f_R0 (tg_alt PI_tg) (S (2 * 0)) <= Alt_PI * / 4
sum_f_R0 (tg_alt PI_tg) (S (2 * 0)) <= Alt_PI * / 4
H:sum_f_R0 (tg_alt PI_tg) (S (2 * 0)) <= Alt_PI * / 4

0 < 1 * (1 + 1 + 1)
H:sum_f_R0 (tg_alt PI_tg) (S (2 * 0)) <= Alt_PI * / 4
1 < 1 + 1 + 1
H:sum_f_R0 (tg_alt PI_tg) (S (2 * 0)) <= Alt_PI * / 4
sum_f_R0 (tg_alt PI_tg) (S (2 * 0)) <= Alt_PI * / 4
H:sum_f_R0 (tg_alt PI_tg) (S (2 * 0)) <= Alt_PI * / 4

1 < 1 + 1 + 1
H:sum_f_R0 (tg_alt PI_tg) (S (2 * 0)) <= Alt_PI * / 4
sum_f_R0 (tg_alt PI_tg) (S (2 * 0)) <= Alt_PI * / 4
H:sum_f_R0 (tg_alt PI_tg) (S (2 * 0)) <= Alt_PI * / 4

sum_f_R0 (tg_alt PI_tg) (S (2 * 0)) <= Alt_PI * / 4
assumption. Qed.