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. (**********)
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 -> RH:Un_decreasing Unn:natsum_f_R0 (tg_alt Un) (S (2 * n)) <= sum_f_R0 (tg_alt Un) (S (2 * S n))Un:nat -> RH:Un_decreasing Unn: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 -> RH:Un_decreasing Unn:nat(2 * S n)%nat = S (S (2 * n))Un:nat -> RH:Un_decreasing Unn:natH0:(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 -> RH:Un_decreasing Unn:nat(2 * S n)%nat = S (S (2 * n))Un:nat -> RH:Un_decreasing Unn:natH0:(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 -> RH:Un_decreasing Unn:nat(2 * S n)%nat = S (S (2 * n))Un:nat -> RH:Un_decreasing Unn:natH0:(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 -> RH:Un_decreasing Unn:nat(2 * S n)%nat = S (S (2 * n))Un:nat -> RH:Un_decreasing Unn:natH0:(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 -> RH:Un_decreasing Unn:nat(2 * S n)%nat = S (S (2 * n))Un:nat -> RH:Un_decreasing Unn:natH0:(2 * S n)%nat = S (S (2 * n))0 <= Un (2 * S n)%nat + -1 * Un (S (2 * S n))Un:nat -> RH:Un_decreasing Unn:nat(2 * S n)%nat = S (S (2 * n))Un:nat -> RH:Un_decreasing Unn:natH0:(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 -> RH:Un_decreasing Unn:nat(2 * S n)%nat = S (S (2 * n))Un:nat -> RH:Un_decreasing Unn:natH0:(2 * S n)%nat = S (S (2 * n))Un (S (2 * S n)) <= Un (2 * S n)%natUn:nat -> RH:Un_decreasing Unn:nat(2 * S n)%nat = S (S (2 * n))Un:nat -> RH:Un_decreasing Unn:nat(2 * S n)%nat = S (S (2 * n))rewrite (H0 n); rewrite (H0 (S (2 * n))); rewrite (H0 (2 * n)%nat); ring. Qed.Un:nat -> RH:Un_decreasing Unn:natH0:forall n0 : nat, S n0 = (n0 + 1)%nat(2 * S n)%nat = S (S (2 * n))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 -> RH:Un_decreasing Unn:natsum_f_R0 (tg_alt Un) (2 * S n) <= sum_f_R0 (tg_alt Un) (2 * n)Un:nat -> RH:Un_decreasing Unn: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 -> RH:Un_decreasing Unn:nat(2 * S n)%nat = S (S (2 * n))Un:nat -> RH:Un_decreasing Unn:natH0:(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 -> RH:Un_decreasing Unn:nat(2 * S n)%nat = S (S (2 * n))Un:nat -> RH:Un_decreasing Unn:natH0:(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) + 0Un:nat -> RH:Un_decreasing Unn:nat(2 * S n)%nat = S (S (2 * n))Un:nat -> RH:Un_decreasing Unn:natH0:(2 * S n)%nat = S (S (2 * n))tg_alt Un (S (2 * n)) + tg_alt Un (S (S (2 * n))) <= 0Un:nat -> RH:Un_decreasing Unn:nat(2 * S n)%nat = S (S (2 * n))Un:nat -> RH:Un_decreasing Unn:natH0:(2 * S n)%nat = S (S (2 * n))-1 * Un (S (2 * n)) + Un (2 * S n)%nat <= 0Un:nat -> RH:Un_decreasing Unn:nat(2 * S n)%nat = S (S (2 * n))Un:nat -> RH:Un_decreasing Unn:natH0:(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)) + 0Un:nat -> RH:Un_decreasing Unn:nat(2 * S n)%nat = S (S (2 * n))Un:nat -> RH:Un_decreasing Unn:natH0:(2 * S n)%nat = S (S (2 * n))Un (2 * S n)%nat <= Un (S (2 * n))Un:nat -> RH:Un_decreasing Unn:nat(2 * S n)%nat = S (S (2 * n))Un:nat -> RH:Un_decreasing Unn:nat(2 * S n)%nat = S (S (2 * n))rewrite (H0 n); rewrite (H0 (S (2 * n))); rewrite (H0 (2 * n)%nat); ring. Qed. (**********)Un:nat -> RH:Un_decreasing Unn:natH0:forall n0 : nat, S n0 = (n0 + 1)%nat(2 * S n)%nat = S (S (2 * n))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)) <= 0forall (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)) <= 0Un:nat -> RH:Un_decreasing UnH0:positivity_seq Unsum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * 0)) <= 0Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * S N)) <= 0Un:nat -> RH:Un_decreasing UnH0:positivity_seq Un-1 * Un 1%nat + -1 * -1 * Un 2%nat <= 0Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * S N)) <= 0Un:nat -> RH:Un_decreasing UnH0:positivity_seq Un-1 * Un 1%nat + Un 2%nat <= 0Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * S N)) <= 0Un:nat -> RH:Un_decreasing UnH0:positivity_seq UnUn 1%nat + (-1 * Un 1%nat + Un 2%nat) <= Un 1%natUn:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * S N)) <= 0Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * S N)) <= 0Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0S (2 * S N) = S (S (S (2 * N))) -> sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * S N)) <= 0Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0S (2 * S N) = S (S (S (2 * N)))Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0H1: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))))) <= 0Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0S (2 * S N) = S (S (S (2 * N)))Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0H1: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 -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0H1:S (2 * S N) = S (S (S (2 * N)))sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0S (2 * S N) = S (S (S (2 * N)))Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0H1: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)) + 0Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0H1:S (2 * S N) = S (S (S (2 * N)))sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0S (2 * S N) = S (S (S (2 * N)))Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0H1: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))))) <= 0Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0H1:S (2 * S N) = S (S (S (2 * N)))sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0S (2 * S N) = S (S (S (2 * N)))Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0H1: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))) <= 0Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0H1:S (2 * S N) = S (S (S (2 * N)))sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0S (2 * S N) = S (S (S (2 * N)))Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0H1: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))) <= 0Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0H1:S (2 * S N) = S (S (S (2 * N)))sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0S (2 * S N) = S (S (S (2 * N)))Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0H1: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))) <= 0Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0H1:S (2 * S N) = S (S (S (2 * N)))S (S (2 * S N)) = (2 * S (S N))%natUn:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0H1:S (2 * S N) = S (S (S (2 * N)))sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0S (2 * S N) = S (S (S (2 * N)))Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0H1: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))) <= 0Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0H1:S (2 * S N) = S (S (S (2 * N)))S (S (2 * S N)) = (2 * S (S N))%natUn:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0H1:S (2 * S N) = S (S (S (2 * N)))sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0S (2 * S N) = S (S (S (2 * N)))Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0H1:S (2 * S N) = S (S (S (2 * N)))H2:S (S (2 * S N)) = (2 * S (S N))%natUn (S (2 * S N)) + (-1 * Un (S (2 * S N)) + Un (S (S (2 * S N)))) <= Un (S (2 * S N)) + 0Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0H1:S (2 * S N) = S (S (S (2 * N)))S (S (2 * S N)) = (2 * S (S N))%natUn:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0H1:S (2 * S N) = S (S (S (2 * N)))sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0S (2 * S N) = S (S (S (2 * N)))Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0H1:S (2 * S N) = S (S (S (2 * N)))H2:S (S (2 * S N)) = (2 * S (S N))%natUn (S (S (2 * S N))) <= Un (S (2 * S N))Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0H1:S (2 * S N) = S (S (S (2 * N)))S (S (2 * S N)) = (2 * S (S N))%natUn:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0H1:S (2 * S N) = S (S (S (2 * N)))sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0S (2 * S N) = S (S (S (2 * N)))Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0H1:S (2 * S N) = S (S (S (2 * N)))S (S (2 * S N)) = (2 * S (S N))%natUn:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0H1:S (2 * S N) = S (S (S (2 * N)))sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0S (2 * S N) = S (S (S (2 * N)))Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0H1:S (2 * S N) = S (S (S (2 * N)))sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0S (2 * S N) = S (S (S (2 * N)))ring. Qed.Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S (2 * N)) <= 0S (2 * S N) = S (S (S (2 * N)))
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 <= 0forall (Un : nat -> R) (N : nat), Un_decreasing Un -> positivity_seq Un -> sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0Un:nat -> RH:Un_decreasing UnH0:positivity_seq Unsum_f_R0 (fun i : nat => tg_alt Un (S i)) 0 <= 0Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S N) <= 0Un:nat -> RH:Un_decreasing UnH0:positivity_seq Un-1 * Un 1%nat <= 0Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S N) <= 0Un:nat -> RH:Un_decreasing UnH0:positivity_seq UnUn 1%nat + -1 * Un 1%nat <= Un 1%nat + 0Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S N) <= 0Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S N) <= 0Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0H1:exists p : nat, N = (2 * p)%nat \/ N = S (2 * p)sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S N) <= 0Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0H1:exists p : nat, N = (2 * p)%nat \/ N = S (2 * p)x:natH2:N = (2 * x)%nat \/ N = S (2 * x)sum_f_R0 (fun i : nat => tg_alt Un (S i)) (S N) <= 0Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0H1:exists p : nat, N = (2 * p)%nat \/ N = S (2 * p)x:natH2:N = (2 * x)%nat \/ N = S (2 * x)H3:N = (2 * x)%natsum_f_R0 (fun i : nat => tg_alt Un (S i)) (S N) <= 0Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0H1:exists p : nat, N = (2 * p)%nat \/ N = S (2 * p)x:natH2: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) <= 0Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0H1:exists p : nat, N = (2 * p)%nat \/ N = S (2 * p)x:natH2: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) <= 0Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0H1:exists p : nat, N = (2 * p)%nat \/ N = S (2 * p)x:natH2: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)))) <= 0Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0H1:exists p : nat, N = (2 * p)%nat \/ N = S (2 * p)x:natH2: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 -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0H1:exists p : nat, N = (2 * p)%nat \/ N = S (2 * p)x:natH2: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)) <= 0Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0H1:exists p : nat, N = (2 * p)%nat \/ N = S (2 * p)x:natH2: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)) + 0Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0H1:exists p : nat, N = (2 * p)%nat \/ N = S (2 * p)x:natH2: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)) <= 0Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0H1:exists p : nat, N = (2 * p)%nat \/ N = S (2 * p)x:natH2:N = (2 * x)%nat \/ N = S (2 * x)H3:N = S (2 * x)tg_alt Un (S (S (S (2 * x)))) <= 0Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0H1:exists p : nat, N = (2 * p)%nat \/ N = S (2 * p)x:natH2: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)) <= 0Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0H1:exists p : nat, N = (2 * p)%nat \/ N = S (2 * p)x:natH2: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))))) <= 0Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0H1:exists p : nat, N = (2 * p)%nat \/ N = S (2 * p)x:natH2: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)) <= 0Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0H1:exists p : nat, N = (2 * p)%nat \/ N = S (2 * p)x:natH2: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)))) <= 0Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0H1:exists p : nat, N = (2 * p)%nat \/ N = S (2 * p)x:natH2: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)) <= 0Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0H1:exists p : nat, N = (2 * p)%nat \/ N = S (2 * p)x:natH2:N = (2 * x)%nat \/ N = S (2 * x)H3:N = S (2 * x)-1 * (-1 * (-1 * 1)) * Un (S (S (S (2 * x)))) <= 0Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0H1:exists p : nat, N = (2 * p)%nat \/ N = S (2 * p)x:natH2: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)) <= 0Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0H1:exists p : nat, N = (2 * p)%nat \/ N = S (2 * p)x:natH2:N = (2 * x)%nat \/ N = S (2 * x)H3:N = S (2 * x)- Un (S (S (S (2 * x)))) <= 0Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0H1:exists p : nat, N = (2 * p)%nat \/ N = S (2 * p)x:natH2: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)) <= 0Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0H1:exists p : nat, N = (2 * p)%nat \/ N = S (2 * p)x:natH2: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)))) + 0Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0H1:exists p : nat, N = (2 * p)%nat \/ N = S (2 * p)x:natH2: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)) <= 0Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0H1:exists p : nat, N = (2 * p)%nat \/ N = S (2 * p)x:natH2:N = (2 * x)%nat \/ N = S (2 * x)H3:N = S (2 * x)0 <= Un (S (S (S (2 * x))))Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0H1:exists p : nat, N = (2 * p)%nat \/ N = S (2 * p)x:natH2: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)) <= 0apply CV_ALT_step2; assumption. Qed. (**********)Un:nat -> RN:natH:Un_decreasing UnH0:positivity_seq UnHrecN:sum_f_R0 (fun i : nat => tg_alt Un (S i)) N <= 0H1:exists p : nat, N = (2 * p)%nat \/ N = S (2 * p)x:natH2: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)) <= 0forall 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 -> RH:Un_decreasing UnH0:positivity_seq Unexists m : R, is_upper_bound (EUn (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N)))) mUn:nat -> RH:Un_decreasing UnH0:positivity_seq Unis_upper_bound (EUn (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N)))) (Un 0%nat)Un:nat -> RH:Un_decreasing UnH0:positivity_seq Unx:RH1:EUn (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N))) xx0:natH2:x = sum_f_R0 (tg_alt Un) (S (2 * x0))x <= Un 0%natUn:nat -> RH:Un_decreasing UnH0:positivity_seq Unx:RH1:EUn (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N))) xx0:natH2: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%natUn:nat -> RH:Un_decreasing UnH0:positivity_seq Unx:RH1:EUn (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N))) xx0:natH2:x = sum_f_R0 (tg_alt Un) (S (2 * x0))(0 < S (2 * x0))%natUn:nat -> RH:Un_decreasing UnH0:positivity_seq Unx:RH1:EUn (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N))) xx0:natH2: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%natUn:nat -> RH:Un_decreasing UnH0:positivity_seq Unx:RH1:EUn (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N))) xx0:natH2:x = sum_f_R0 (tg_alt Un) (S (2 * x0))Un 0%nat = tg_alt Un 0Un:nat -> RH:Un_decreasing UnH0:positivity_seq Unx:RH1:EUn (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N))) xx0:natH2:x = sum_f_R0 (tg_alt Un) (S (2 * x0))(0 < S (2 * x0))%natUn:nat -> RH:Un_decreasing UnH0:positivity_seq Unx:RH1:EUn (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N))) xx0:natH2: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 + 0Un:nat -> RH:Un_decreasing UnH0:positivity_seq Unx:RH1:EUn (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N))) xx0:natH2:x = sum_f_R0 (tg_alt Un) (S (2 * x0))Un 0%nat = tg_alt Un 0Un:nat -> RH:Un_decreasing UnH0:positivity_seq Unx:RH1:EUn (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N))) xx0:natH2:x = sum_f_R0 (tg_alt Un) (S (2 * x0))(0 < S (2 * x0))%natUn:nat -> RH:Un_decreasing UnH0:positivity_seq Unx:RH1:EUn (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N))) xx0:natH2: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))) <= 0Un:nat -> RH:Un_decreasing UnH0:positivity_seq Unx:RH1:EUn (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N))) xx0:natH2:x = sum_f_R0 (tg_alt Un) (S (2 * x0))Un 0%nat = tg_alt Un 0Un:nat -> RH:Un_decreasing UnH0:positivity_seq Unx:RH1:EUn (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N))) xx0:natH2:x = sum_f_R0 (tg_alt Un) (S (2 * x0))(0 < S (2 * x0))%natUn:nat -> RH:Un_decreasing UnH0:positivity_seq Unx:RH1:EUn (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N))) xx0:natH2:x = sum_f_R0 (tg_alt Un) (S (2 * x0))Un 0%nat = tg_alt Un 0Un:nat -> RH:Un_decreasing UnH0:positivity_seq Unx:RH1:EUn (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N))) xx0:natH2:x = sum_f_R0 (tg_alt Un) (S (2 * x0))(0 < S (2 * x0))%natapply lt_O_Sn. Qed.Un:nat -> RH:Un_decreasing UnH0:positivity_seq Unx:RH1:EUn (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N))) xx0:natH2:x = sum_f_R0 (tg_alt Un) (S (2 * x0))(0 < S (2 * x0))%nat
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 -> RH:Un_decreasing UnH0:positivity_seq UnH1:Un_cv Un 0{l : R | Un_cv (fun N : nat => sum_f_R0 (tg_alt Un) N) l}Un:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:Un_cv Un 0H2: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 -> RH:Un_decreasing UnH0:positivity_seq UnH1:Un_cv Un 0H2: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 -> RH:Un_decreasing UnH0:positivity_seq UnH1:Un_cv Un 0H2: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:Rp: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 -> RH:Un_decreasing UnH0:positivity_seq UnH1:Un_cv Un 0H2: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:Rp:Un_cv (fun N : nat => sum_f_R0 (tg_alt Un) (S (2 * N))) xUn_cv (fun N : nat => sum_f_R0 (tg_alt Un) N) xUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (Un n - 0) < epsH2: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:Rp: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) < epsforall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (sum_f_R0 (tg_alt Un) n - x) < epsUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (Un n - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (sum_f_R0 (tg_alt Un) n - x) < epsUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (Un n - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n : nat, (n >= N2)%nat -> Rabs (Un n - 0) < eps / 2exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (sum_f_R0 (tg_alt Un) n - x) < epsUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (Un n - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n : nat, (n >= N2)%nat -> Rabs (Un n - 0) < eps / 2N1:natH7:forall n : nat, (n >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n)) - x) < eps / 2exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (sum_f_R0 (tg_alt Un) n - x) < epsUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n : nat, (n >= N0)%nat -> Rabs (Un n - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n : nat, (n >= N2)%nat -> Rabs (Un n - 0) < eps / 2N1:natH7:forall n : nat, (n >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natexists N0 : nat, forall n : nat, (n >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n - x) < epsUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natRabs (sum_f_R0 (tg_alt Un) n - x) < epsUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)Rabs (sum_f_R0 (tg_alt Un) n - x) < epsUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)Rabs (sum_f_R0 (tg_alt Un) n - x) < epsUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)(N1 <= P)%nat -> Rabs (sum_f_R0 (tg_alt Un) n - x) < epsUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)(N1 <= P)%natUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:(N1 <= P)%natH12:n = (2 * P)%natRabs (sum_f_R0 (tg_alt Un) n - x) < epsUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:(N1 <= P)%natH12:n = S (2 * P)Rabs (sum_f_R0 (tg_alt Un) n - x) < epsUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)(N1 <= P)%natUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:(N1 <= P)%natH12:n = (2 * P)%natRabs (sum_f_R0 (tg_alt Un) (S n) - x + - tg_alt Un (S n)) < epsUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:(N1 <= P)%natH12:n = (2 * P)%natsum_f_R0 (tg_alt Un) (S n) - x + - tg_alt Un (S n) = sum_f_R0 (tg_alt Un) n - xUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:(N1 <= P)%natH12:n = S (2 * P)Rabs (sum_f_R0 (tg_alt Un) n - x) < epsUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)(N1 <= P)%natUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:(N1 <= P)%natH12:n = (2 * P)%natRabs (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 -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:(N1 <= P)%natH12:n = (2 * P)%natRabs (sum_f_R0 (tg_alt Un) (S n) - x) + Rabs (- tg_alt Un (S n)) < epsUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:(N1 <= P)%natH12:n = (2 * P)%natsum_f_R0 (tg_alt Un) (S n) - x + - tg_alt Un (S n) = sum_f_R0 (tg_alt Un) n - xUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:(N1 <= P)%natH12:n = S (2 * P)Rabs (sum_f_R0 (tg_alt Un) n - x) < epsUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)(N1 <= P)%natUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:(N1 <= P)%natH12:n = (2 * P)%natRabs (sum_f_R0 (tg_alt Un) (S n) - x) + Rabs (- tg_alt Un (S n)) < epsUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:(N1 <= P)%natH12:n = (2 * P)%natsum_f_R0 (tg_alt Un) (S n) - x + - tg_alt Un (S n) = sum_f_R0 (tg_alt Un) n - xUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:(N1 <= P)%natH12:n = S (2 * P)Rabs (sum_f_R0 (tg_alt Un) n - x) < epsUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)(N1 <= P)%natUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:(N1 <= P)%natH12:n = (2 * P)%natRabs (sum_f_R0 (tg_alt Un) (S n) - x) < eps / 2Un:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:(N1 <= P)%natH12:n = (2 * P)%natRabs (- tg_alt Un (S n)) < eps / 2Un:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:(N1 <= P)%natH12:n = (2 * P)%natsum_f_R0 (tg_alt Un) (S n) - x + - tg_alt Un (S n) = sum_f_R0 (tg_alt Un) n - xUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:(N1 <= P)%natH12:n = S (2 * P)Rabs (sum_f_R0 (tg_alt Un) n - x) < epsUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)(N1 <= P)%natUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:(N1 <= P)%natH12:n = (2 * P)%natRabs (- tg_alt Un (S n)) < eps / 2Un:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:(N1 <= P)%natH12:n = (2 * P)%natsum_f_R0 (tg_alt Un) (S n) - x + - tg_alt Un (S n) = sum_f_R0 (tg_alt Un) n - xUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:(N1 <= P)%natH12:n = S (2 * P)Rabs (sum_f_R0 (tg_alt Un) n - x) < epsUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)(N1 <= P)%natUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 + 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:(N1 <= P)%natH12:n = (2 * P)%nat(S n >= N2)%natUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:(N1 <= P)%natH12:n = (2 * P)%natsum_f_R0 (tg_alt Un) (S n) - x + - tg_alt Un (S n) = sum_f_R0 (tg_alt Un) n - xUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:(N1 <= P)%natH12:n = S (2 * P)Rabs (sum_f_R0 (tg_alt Un) n - x) < epsUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)(N1 <= P)%natUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 + 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:(N1 <= P)%natH12:n = (2 * P)%nat(N2 <= n)%natUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 + 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:(N1 <= P)%natH12:n = (2 * P)%nat(n <= S n)%natUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:(N1 <= P)%natH12:n = (2 * P)%natsum_f_R0 (tg_alt Un) (S n) - x + - tg_alt Un (S n) = sum_f_R0 (tg_alt Un) n - xUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:(N1 <= P)%natH12:n = S (2 * P)Rabs (sum_f_R0 (tg_alt Un) n - x) < epsUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)(N1 <= P)%natUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 + 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:(N1 <= P)%natH12:n = (2 * P)%nat(n <= S n)%natUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:(N1 <= P)%natH12:n = (2 * P)%natsum_f_R0 (tg_alt Un) (S n) - x + - tg_alt Un (S n) = sum_f_R0 (tg_alt Un) n - xUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:(N1 <= P)%natH12:n = S (2 * P)Rabs (sum_f_R0 (tg_alt Un) n - x) < epsUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)(N1 <= P)%natUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:(N1 <= P)%natH12:n = (2 * P)%natsum_f_R0 (tg_alt Un) (S n) - x + - tg_alt Un (S n) = sum_f_R0 (tg_alt Un) n - xUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:(N1 <= P)%natH12:n = S (2 * P)Rabs (sum_f_R0 (tg_alt Un) n - x) < epsUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)(N1 <= P)%natUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:(N1 <= P)%natH12:n = S (2 * P)Rabs (sum_f_R0 (tg_alt Un) n - x) < epsUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)(N1 <= P)%natUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:(N1 <= P)%natH12:n = S (2 * P)Rabs (sum_f_R0 (tg_alt Un) (S (2 * P)) - x) < eps / 2Un:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:(N1 <= P)%natH12:n = S (2 * P)eps / 2 < epsUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)(N1 <= P)%natUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:(N1 <= P)%natH12:n = S (2 * P)eps / 2 < epsUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)(N1 <= P)%natUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:(N1 <= P)%natH12:n = S (2 * P)0 < 2Un:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:(N1 <= P)%natH12:n = S (2 * P)2 * (eps * / 2) < 2 * epsUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)(N1 <= P)%natUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:(N1 <= P)%natH12:n = S (2 * P)2 * (eps * / 2) < 2 * epsUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)(N1 <= P)%natUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:(N1 <= P)%natH12:n = S (2 * P)eps < 2 * epsUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)(N1 <= P)%natUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:(N1 <= P)%natH12:n = S (2 * P)eps < eps + epsUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)(N1 <= P)%natUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)(N1 <= P)%natUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:n = (2 * P)%nat(2 * N1 <= 2 * P)%natUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:n = S (2 * P)(2 * N1 <= 2 * P)%natUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:n = (2 * P)%nat(2 * N1 <= N)%natUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:n = (2 * P)%nat(N <= n)%natUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:n = S (2 * P)(2 * N1 <= 2 * P)%natUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:n = (2 * P)%nat(N <= n)%natUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:n = S (2 * P)(2 * N1 <= 2 * P)%natUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:n = S (2 * P)(2 * N1 <= 2 * P)%natUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:n = S (2 * P)(2 * N1 < S (2 * P))%natUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:n = S (2 * P)(2 * N1 < n)%natUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:n = S (2 * P)(2 * N1 < N)%natUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:n = S (2 * P)(N <= n)%natUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:n = S (2 * P)(2 * N1 < S (2 * N1))%natUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:n = S (2 * P)(S (2 * N1) <= Nat.max (S (2 * N1)) N2)%natUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:n = S (2 * P)(N <= n)%natUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:n = S (2 * P)(S (2 * N1) <= Nat.max (S (2 * N1)) N2)%natUn:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:n = S (2 * P)(N <= n)%natassumption. Qed. (*************************************************)Un:nat -> RH:Un_decreasing UnH0:positivity_seq UnH1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (Un n0 - 0) < eps0H2: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:Rp: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) < eps0eps:RH4:eps > 0H5:0 < eps / 2N2:natH6:forall n0 : nat, (n0 >= N2)%nat -> Rabs (Un n0 - 0) < eps / 2N1:natH7:forall n0 : nat, (n0 >= N1)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n0)) - x) < eps / 2N:=Nat.max (S (2 * N1)) N2:natn:natH8:(n >= N)%natH9:exists p0 : nat, n = (2 * p0)%nat \/ n = S (2 * p0)P:natH10:n = (2 * P)%nat \/ n = S (2 * P)H11:n = S (2 * P)(N <= n)%nat
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 -> RH:Un_decreasing UnH0:Un_cv Un 0Un_decreasing UnUn:nat -> RH:Un_decreasing UnH0:Un_cv Un 0positivity_seq UnUn:nat -> RH:Un_decreasing UnH0:Un_cv Un 0Un_cv Un 0Un:nat -> RH:Un_decreasing UnH0:Un_cv Un 0positivity_seq UnUn:nat -> RH:Un_decreasing UnH0:Un_cv Un 0Un_cv Un 0assumption. Qed.Un:nat -> RH:Un_decreasing UnH0:Un_cv Un 0Un_cv Un 0forall (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 -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) lsum_f_R0 (tg_alt Un) (S (2 * N)) <= l <= sum_f_R0 (tg_alt Un) (2 * N)Un:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) lUn_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 -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) lUn_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) lUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) lUn_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 -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) lUn_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0))) lUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) lUn_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) lUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) lH2:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0))) lH3:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) lsum_f_R0 (tg_alt Un) (S (2 * N)) <= lUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) lH2:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0))) lH3:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) ll <= sum_f_R0 (tg_alt Un) (2 * N)Un:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) lUn_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0))) lUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) lUn_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) lUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) lH2:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0))) lH3:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) lUn_growing (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0)))Un:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) lH2:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0))) lH3:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) lUn_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0))) lUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) lH2:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0))) lH3:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) ll <= sum_f_R0 (tg_alt Un) (2 * N)Un:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) lUn_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0))) lUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) lUn_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) lUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) lH2:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0))) lH3:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) lUn_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0))) lUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) lH2:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0))) lH3:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) ll <= sum_f_R0 (tg_alt Un) (2 * N)Un:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) lUn_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0))) lUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) lUn_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) lUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) lH2:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0))) lH3:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) ll <= sum_f_R0 (tg_alt Un) (2 * N)Un:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) lUn_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0))) lUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) lUn_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) lUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) lH2:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0))) lH3:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) lUn_decreasing (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0))Un:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) lH2:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0))) lH3:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) lUn_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) lUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) lUn_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0))) lUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) lUn_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) lUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) lH2:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0))) lH3:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) lUn_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) lUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) lUn_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0))) lUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) lUn_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) lUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) lUn_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (S (2 * N0))) lUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) lUn_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) lUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n : nat, (n >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n - l) < eps0eps:RH2:eps > 0exists N0 : nat, forall n : nat, (n >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n)) - l) < epsUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) lUn_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) lUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n : nat, (n >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n - l) < eps0eps:RH2:eps > 0x:natH3:forall n : nat, (n >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n - l) < epsexists N0 : nat, forall n : nat, (n >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (S (2 * n)) - l) < epsUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) lUn_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) lUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps0eps:RH2:eps > 0x:natH3:forall n0 : nat, (n0 >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < epsn:natH4:(n >= x)%natRabs (sum_f_R0 (tg_alt Un) (S (2 * n)) - l) < epsUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) lUn_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) lUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps0eps:RH2:eps > 0x:natH3:forall n0 : nat, (n0 >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < epsn:natH4:(n >= x)%nat(S (2 * n) >= x)%natUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) lUn_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) lUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps0eps:RH2:eps > 0x:natH3:forall n0 : nat, (n0 >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < epsn:natH4:(n >= x)%nat(x <= 2 * n)%natUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps0eps:RH2:eps > 0x:natH3:forall n0 : nat, (n0 >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < epsn:natH4:(n >= x)%nat(2 * n <= S (2 * n))%natUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) lUn_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) lUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps0eps:RH2:eps > 0x:natH3:forall n0 : nat, (n0 >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < epsn:natH4:(n >= x)%nat(x <= n)%natUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps0eps:RH2:eps > 0x:natH3:forall n0 : nat, (n0 >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < epsn:natH4:(n >= x)%nat(n <= 2 * n)%natUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps0eps:RH2:eps > 0x:natH3:forall n0 : nat, (n0 >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < epsn:natH4:(n >= x)%nat(2 * n <= S (2 * n))%natUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) lUn_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) lUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps0eps:RH2:eps > 0x:natH3:forall n0 : nat, (n0 >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < epsn:natH4:(n >= x)%nat(n <= 2 * n)%natUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps0eps:RH2:eps > 0x:natH3:forall n0 : nat, (n0 >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < epsn:natH4:(n >= x)%nat(2 * n <= S (2 * n))%natUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) lUn_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) lUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps0eps:RH2:eps > 0x:natH3:forall n0 : nat, (n0 >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < epsn:natH4:(n >= x)%natH5:2%nat = 0%nat \/ (n <= 2 * n)%nat(n <= 2 * n)%natUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps0eps:RH2:eps > 0x:natH3:forall n0 : nat, (n0 >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < epsn:natH4:(n >= x)%nat(2 * n <= S (2 * n))%natUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) lUn_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) lUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps0eps:RH2:eps > 0x:natH3:forall n0 : nat, (n0 >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < epsn:natH4:(n >= x)%natH5:2%nat = 0%nat \/ (n <= 2 * n)%natH6:2%nat = 0%nat(n <= 2 * n)%natUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps0eps:RH2:eps > 0x:natH3:forall n0 : nat, (n0 >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < epsn:natH4:(n >= x)%natH5:2%nat = 0%nat \/ (n <= 2 * n)%natH6:(n <= 2 * n)%nat(n <= 2 * n)%natUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps0eps:RH2:eps > 0x:natH3:forall n0 : nat, (n0 >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < epsn:natH4:(n >= x)%nat(2 * n <= S (2 * n))%natUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) lUn_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) lUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps0eps:RH2:eps > 0x:natH3:forall n0 : nat, (n0 >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < epsn:natH4:(n >= x)%natH5:2%nat = 0%nat \/ (n <= 2 * n)%natH6:(n <= 2 * n)%nat(n <= 2 * n)%natUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps0eps:RH2:eps > 0x:natH3:forall n0 : nat, (n0 >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < epsn:natH4:(n >= x)%nat(2 * n <= S (2 * n))%natUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) lUn_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) lUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps0eps:RH2:eps > 0x:natH3:forall n0 : nat, (n0 >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < epsn:natH4:(n >= x)%nat(2 * n <= S (2 * n))%natUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) lUn_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) lUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) N0) lUn_cv (fun N0 : nat => sum_f_R0 (tg_alt Un) (2 * N0)) lUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n : nat, (n >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n - l) < eps0eps:RH2:eps > 0exists N0 : nat, forall n : nat, (n >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (2 * n) - l) < epsUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n : nat, (n >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n - l) < eps0eps:RH2:eps > 0x:natH3:forall n : nat, (n >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n - l) < epsexists N0 : nat, forall n : nat, (n >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) (2 * n) - l) < epsUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps0eps:RH2:eps > 0x:natH3:forall n0 : nat, (n0 >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < epsn:natH4:(n >= x)%natRabs (sum_f_R0 (tg_alt Un) (2 * n) - l) < epsUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps0eps:RH2:eps > 0x:natH3:forall n0 : nat, (n0 >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < epsn:natH4:(n >= x)%nat(2 * n >= x)%natUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps0eps:RH2:eps > 0x:natH3:forall n0 : nat, (n0 >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < epsn:natH4:(n >= x)%nat(x <= n)%natUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps0eps:RH2:eps > 0x:natH3:forall n0 : nat, (n0 >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < epsn:natH4:(n >= x)%nat(n <= 2 * n)%natUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps0eps:RH2:eps > 0x:natH3:forall n0 : nat, (n0 >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < epsn:natH4:(n >= x)%nat(n <= 2 * n)%natUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps0eps:RH2:eps > 0x:natH3:forall n0 : nat, (n0 >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < epsn:natH4:(n >= x)%natH5:2%nat = 0%nat \/ (n <= 2 * n)%nat(n <= 2 * n)%natUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps0eps:RH2:eps > 0x:natH3:forall n0 : nat, (n0 >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < epsn:natH4:(n >= x)%natH5:2%nat = 0%nat \/ (n <= 2 * n)%natH6:2%nat = 0%nat(n <= 2 * n)%natUn:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps0eps:RH2:eps > 0x:natH3:forall n0 : nat, (n0 >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < epsn:natH4:(n >= x)%natH5:2%nat = 0%nat \/ (n <= 2 * n)%natH6:(n <= 2 * n)%nat(n <= 2 * n)%natassumption. Qed. (***************************************)Un:nat -> Rl:RN:natH:Un_decreasing UnH0:Un_cv Un 0H1:forall eps0 : R, eps0 > 0 -> exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < eps0eps:RH2:eps > 0x:natH3:forall n0 : nat, (n0 >= x)%nat -> Rabs (sum_f_R0 (tg_alt Un) n0 - l) < epsn:natH4:(n >= x)%natH5:2%nat = 0%nat \/ (n <= 2 * n)%natH6:(n <= 2 * n)%nat(n <= 2 * n)%nat
(***************************************) Definition PI_tg (n:nat) := / INR (2 * n + 1).forall n : nat, 0 <= PI_tg nintro; 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.forall n : nat, 0 <= PI_tg nUn_decreasing PI_tgUn_decreasing PI_tgn:nat/ INR (2 * S n + 1) <= / INR (2 * n + 1)n:nat0 < INR (2 * n + 1)n:natINR (2 * n + 1) * / INR (2 * S n + 1) <= INR (2 * n + 1) * / INR (2 * n + 1)n:nat(0 < 2 * n + 1)%natn:natINR (2 * n + 1) * / INR (2 * S n + 1) <= INR (2 * n + 1) * / INR (2 * n + 1)n:natINR (2 * n + 1) * / INR (2 * S n + 1) <= INR (2 * n + 1) * / INR (2 * n + 1)n:natINR (2 * n + 1) * / INR (2 * S n + 1) <= 1n:natINR (2 * n + 1) <> 0n:nat0 < INR (2 * S n + 1)n:natINR (2 * S n + 1) * (INR (2 * n + 1) * / INR (2 * S n + 1)) <= INR (2 * S n + 1) * 1n:natINR (2 * n + 1) <> 0n:nat(0 < 2 * S n + 1)%natn:natINR (2 * S n + 1) * (INR (2 * n + 1) * / INR (2 * S n + 1)) <= INR (2 * S n + 1) * 1n:natINR (2 * n + 1) <> 0n:natINR (2 * S n + 1) * (INR (2 * n + 1) * / INR (2 * S n + 1)) <= INR (2 * S n + 1) * 1n:natINR (2 * n + 1) <> 0n:natINR (2 * n + 1) * 1 <= INR (2 * S n + 1) * 1n:natINR (2 * S n + 1) <> 0n:natINR (2 * n + 1) <> 0n:nat(2 * n + 1 <= 2 * S n + 1)%natn:natINR (2 * S n + 1) <> 0n:natINR (2 * n + 1) <> 0n:nat(2 * n + 1 <= S (S (2 * n + 1)))%natn:natS (S (2 * n + 1)) = (2 * S n + 1)%natn:natINR (2 * S n + 1) <> 0n:natINR (2 * n + 1) <> 0n:natS (S (2 * n + 1)) = (2 * S n + 1)%natn:natINR (2 * S n + 1) <> 0n:natINR (2 * n + 1) <> 0n:natINR (2 * S n + 1) <> 0n:natINR (2 * n + 1) <> 0apply not_O_INR; replace (2 * n + 1)%nat with (S (2 * n)); [ discriminate | ring ]. Qed.n:natINR (2 * n + 1) <> 0Un_cv PI_tg 0Un_cv PI_tg 0eps:RH:eps > 0exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (PI_tg n - 0) < epseps:RH:eps > 0H0:0 < 2 * epsexists N : nat, forall n : nat, (n >= N)%nat -> Rabs (PI_tg n - 0) < epseps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1exists N : nat, forall n : nat, (n >= N)%nat -> Rabs (PI_tg n - 0) < epseps:RH:eps > 0H0:0 < 2 * epsH1: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) < epseps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mexists N : nat, forall n : nat, (n >= N)%nat -> Rabs (PI_tg n - 0) < epseps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat Nexists N0 : nat, forall n : nat, (n >= N0)%nat -> Rabs (PI_tg n - 0) < epseps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat N(0 < N)%nat -> exists N0 : nat, forall n : nat, (n >= N0)%nat -> Rabs (PI_tg n - 0) < epseps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat N(0 < N)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natRabs (PI_tg n - 0) < epseps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat N(0 < N)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%nat(0 < n)%nat -> Rabs (PI_tg n - 0) < epseps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%nat(0 < n)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat N(0 < N)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natPI_tg n < epseps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natPI_tg n >= 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%nat(0 < n)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat N(0 < N)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%nat/ INR (2 * n + 1) < / INR (2 * n)eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%nat/ INR (2 * n) < epseps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natPI_tg n >= 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%nat(0 < n)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat N(0 < N)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%nat0 < INR (2 * n)eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natINR (2 * n) * / INR (2 * n + 1) < INR (2 * n) * / INR (2 * n)eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%nat/ INR (2 * n) < epseps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natPI_tg n >= 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%nat(0 < n)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat N(0 < N)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%nat(0 < 2 * n)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natINR (2 * n) * / INR (2 * n + 1) < INR (2 * n) * / INR (2 * n)eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%nat/ INR (2 * n) < epseps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natPI_tg n >= 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%nat(0 < n)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat N(0 < N)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%nat(0 < n + n)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natINR (2 * n) * / INR (2 * n + 1) < INR (2 * n) * / INR (2 * n)eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%nat/ INR (2 * n) < epseps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natPI_tg n >= 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%nat(0 < n)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat N(0 < N)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%nat(0 < n)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%nat(n <= n + n)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natINR (2 * n) * / INR (2 * n + 1) < INR (2 * n) * / INR (2 * n)eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%nat/ INR (2 * n) < epseps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natPI_tg n >= 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%nat(0 < n)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat N(0 < N)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%nat(n <= n + n)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natINR (2 * n) * / INR (2 * n + 1) < INR (2 * n) * / INR (2 * n)eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%nat/ INR (2 * n) < epseps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natPI_tg n >= 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%nat(0 < n)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat N(0 < N)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natINR (2 * n) * / INR (2 * n + 1) < INR (2 * n) * / INR (2 * n)eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%nat/ INR (2 * n) < epseps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natPI_tg n >= 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%nat(0 < n)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat N(0 < N)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natINR (2 * n) * / INR (2 * n + 1) < 1eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natINR (2 * n) <> 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%nat/ INR (2 * n) < epseps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natPI_tg n >= 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%nat(0 < n)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat N(0 < N)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%nat0 < INR (2 * n + 1)eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natINR (2 * n + 1) * (INR (2 * n) * / INR (2 * n + 1)) < INR (2 * n + 1) * 1eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natINR (2 * n) <> 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%nat/ INR (2 * n) < epseps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natPI_tg n >= 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%nat(0 < n)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat N(0 < N)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%nat(0 < 2 * n + 1)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natINR (2 * n + 1) * (INR (2 * n) * / INR (2 * n + 1)) < INR (2 * n + 1) * 1eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natINR (2 * n) <> 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%nat/ INR (2 * n) < epseps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natPI_tg n >= 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%nat(0 < n)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat N(0 < N)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natINR (2 * n + 1) * (INR (2 * n) * / INR (2 * n + 1)) < INR (2 * n + 1) * 1eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natINR (2 * n) <> 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%nat/ INR (2 * n) < epseps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natPI_tg n >= 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%nat(0 < n)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat N(0 < N)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natINR (2 * n) * / INR (2 * n + 1) * INR (2 * n + 1) < INR (2 * n + 1) * 1eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natINR (2 * n) <> 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%nat/ INR (2 * n) < epseps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natPI_tg n >= 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%nat(0 < n)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat N(0 < N)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natINR (2 * n) * 1 < INR (2 * n + 1) * 1eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natINR (2 * n + 1) <> 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natINR (2 * n) <> 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%nat/ INR (2 * n) < epseps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natPI_tg n >= 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%nat(0 < n)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat N(0 < N)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%nat(2 * n < 2 * n + 1)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natINR (2 * n + 1) <> 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natINR (2 * n) <> 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%nat/ INR (2 * n) < epseps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natPI_tg n >= 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%nat(0 < n)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat N(0 < N)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natINR (2 * n + 1) <> 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natINR (2 * n) <> 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%nat/ INR (2 * n) < epseps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natPI_tg n >= 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%nat(0 < n)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat N(0 < N)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natINR (2 * n) <> 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%nat/ INR (2 * n) < epseps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natPI_tg n >= 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%nat(0 < n)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat N(0 < N)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natINR (2 * S (Init.Nat.pred n)) <> 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natS (Init.Nat.pred n) = neps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%nat/ INR (2 * n) < epseps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natPI_tg n >= 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%nat(0 < n)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat N(0 < N)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natS (Init.Nat.pred n) = neps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%nat/ INR (2 * n) < epseps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natPI_tg n >= 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%nat(0 < n)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat N(0 < N)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%nat(0 < n)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%nat/ INR (2 * n) < epseps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natPI_tg n >= 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%nat(0 < n)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat N(0 < N)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%nat/ INR (2 * n) < epseps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natPI_tg n >= 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%nat(0 < n)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat N(0 < N)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%nat/ INR (2 * n) <= / INR (2 * N)eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%nat/ INR (2 * N) < epseps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natPI_tg n >= 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%nat(0 < n)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat N(0 < N)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%nat0 < INR (2 * N)eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natINR (2 * N) <= INR (2 * n)eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%nat/ INR (2 * N) < epseps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natPI_tg n >= 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%nat(0 < n)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat N(0 < N)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natINR (2 * N) <= INR (2 * n)eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%nat/ INR (2 * N) < epseps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natPI_tg n >= 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%nat(0 < n)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat N(0 < N)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%nat(2 * N <= 2 * n)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%nat/ INR (2 * N) < epseps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natPI_tg n >= 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%nat(0 < n)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat N(0 < N)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%nat/ INR (2 * N) < epseps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natPI_tg n >= 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%nat(0 < n)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat N(0 < N)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%nat/ (INR 2 * INR N) < epseps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natPI_tg n >= 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%nat(0 < n)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat N(0 < N)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%nat0 < INR N / epseps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natINR N / eps * / (INR 2 * INR N) < INR N / eps * epseps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natPI_tg n >= 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%nat(0 < n)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat N(0 < N)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%nat0 < INR Neps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natINR N / eps * / (INR 2 * INR N) < INR N / eps * epseps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natPI_tg n >= 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%nat(0 < n)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat N(0 < N)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natINR N / eps * / (INR 2 * INR N) < INR N / eps * epseps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natPI_tg n >= 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%nat(0 < n)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat N(0 < N)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%nat/ (2 * eps) < INR N / eps * epseps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%nat/ (2 * eps) = INR N / eps * / (INR 2 * INR N)eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natPI_tg n >= 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%nat(0 < n)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat N(0 < N)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%nat/ (2 * eps) < INR Neps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natINR N = INR N / eps * epseps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%nat/ (2 * eps) = INR N / eps * / (INR 2 * INR N)eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natPI_tg n >= 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%nat(0 < n)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat N(0 < N)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%nat/ (2 * eps) < IZR (Z.of_nat N)eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natINR N = INR N / eps * epseps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%nat/ (2 * eps) = INR N / eps * / (INR 2 * INR N)eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natPI_tg n >= 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%nat(0 < n)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat N(0 < N)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natINR N = INR N / eps * epseps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%nat/ (2 * eps) = INR N / eps * / (INR 2 * INR N)eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natPI_tg n >= 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%nat(0 < n)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat N(0 < N)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%nateps <> 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%nat/ (2 * eps) = INR N / eps * / (INR 2 * INR N)eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natPI_tg n >= 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%nat(0 < n)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat N(0 < N)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%nat/ (2 * eps) = INR N / eps * / (INR 2 * INR N)eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natPI_tg n >= 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%nat(0 < n)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat N(0 < N)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natINR N <> 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%nateps <> 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natPI_tg n >= 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%nat(0 < n)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat N(0 < N)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%nateps <> 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natPI_tg n >= 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%nat(0 < n)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat N(0 < N)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%natH7:(0 < n)%natPI_tg n >= 0eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%nat(0 < n)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat N(0 < N)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:(0 < N)%natn:natH6:(n >= N)%nat(0 < n)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat N(0 < N)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat N(0 < N)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:IZR (up (/ (2 * eps))) > / (2 * eps)(0 < N)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:IZR (up (/ (2 * eps))) > / (2 * eps)l:(0 < N)%nat(0 < N)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mH4:up (/ (2 * eps)) = Z.of_nat 0H5:IZR (up (/ (2 * eps))) > / (2 * eps)(0 < 0)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:IZR (up (/ (2 * eps))) > / (2 * eps)H6:(N < 0)%nat(0 < N)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mH4:up (/ (2 * eps)) = Z.of_nat 0H5:IZR (up (/ (2 * eps))) > / (2 * eps)(0 < 0)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:IZR (up (/ (2 * eps))) > / (2 * eps)H6:(N < 0)%nat(0 < N)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mH4:up (/ (2 * eps)) = Z.of_nat 0H5:IZR (Z.of_nat 0) > / (2 * eps)(0 < 0)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:IZR (up (/ (2 * eps))) > / (2 * eps)H6:(N < 0)%nat(0 < N)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mH4:up (/ (2 * eps)) = Z.of_nat 0H5:0 > / (2 * eps)(0 < 0)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:IZR (up (/ (2 * eps))) > / (2 * eps)H6:(N < 0)%nat(0 < N)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mH4:up (/ (2 * eps)) = Z.of_nat 0H5:0 > / (2 * eps)H6:0 < / (2 * eps)(0 < 0)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:IZR (up (/ (2 * eps))) > / (2 * eps)H6:(N < 0)%nat(0 < N)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1H2:(0 <= up (/ (2 * eps)))%ZH3:exists m : nat, up (/ (2 * eps)) = Z.of_nat mN:natH4:up (/ (2 * eps)) = Z.of_nat NH5:IZR (up (/ (2 * eps))) > / (2 * eps)H6:(N < 0)%nat(0 < N)%nateps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1(0 <= up (/ (2 * eps)))%Zeps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 10 <= IZR (up (/ (2 * eps)))eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 10 < / (2 * eps)eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1/ (2 * eps) < IZR (up (/ (2 * eps)))elim H1; intros; assumption. Qed.eps:RH:eps > 0H0:0 < 2 * epsH1:IZR (up (/ (2 * eps))) > / (2 * eps) /\ IZR (up (/ (2 * eps))) - / (2 * eps) <= 1/ (2 * eps) < IZR (up (/ (2 * eps))){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_tgUn_cv PI_tg 0apply PI_tg_cv. Qed.Un_cv PI_tg 0
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:natUn_decreasing PI_tgN:natUn_cv PI_tg 0N:natUn_cv (fun N0 : nat => sum_f_R0 (tg_alt PI_tg) N0) (Alt_PI / 4)N:natUn_cv PI_tg 0N:natUn_cv (fun N0 : nat => sum_f_R0 (tg_alt PI_tg) N0) (Alt_PI / 4)N:natUn_cv (fun N0 : nat => sum_f_R0 (tg_alt PI_tg) N0) (Alt_PI / 4)N:natx:RUn_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:natx:RUn_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) xN:natx:Rx = 4 * x / 4unfold Rdiv; rewrite (Rmult_comm 4); rewrite Rmult_assoc; rewrite <- Rinv_r_sym; [ rewrite Rmult_1_r; reflexivity | discrR ]. Qed.N:natx:Rx = 4 * x / 40 < Alt_PI0 < Alt_PIH:sum_f_R0 (tg_alt PI_tg) (S (2 * 0)) <= Alt_PI / 4 <= sum_f_R0 (tg_alt PI_tg) (2 * 0)0 < Alt_PIH:sum_f_R0 (tg_alt PI_tg) (S (2 * 0)) <= Alt_PI / 4 <= sum_f_R0 (tg_alt PI_tg) (2 * 0)0 < / 4H: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_PIH: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_PIH: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 * / 4H:sum_f_R0 (tg_alt PI_tg) (S (2 * 0)) <= Alt_PI / 40 < Alt_PI * / 4H:sum_f_R0 (tg_alt PI_tg) (S (2 * 0)) <= Alt_PI * / 40 < sum_f_R0 (tg_alt PI_tg) (S (2 * 0))H:sum_f_R0 (tg_alt PI_tg) (S (2 * 0)) <= Alt_PI * / 4sum_f_R0 (tg_alt PI_tg) (S (2 * 0)) <= Alt_PI * / 4H:sum_f_R0 (tg_alt PI_tg) (S (2 * 0)) <= Alt_PI * / 4PI_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 * / 4sum_f_R0 (tg_alt PI_tg) (S (2 * 0)) <= Alt_PI * / 4H: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 * / 4sum_f_R0 (tg_alt PI_tg) (S (2 * 0)) <= Alt_PI * / 4H:sum_f_R0 (tg_alt PI_tg) (S (2 * 0)) <= Alt_PI * / 40 < 1 * (1 + 1 + 1)H:sum_f_R0 (tg_alt PI_tg) (S (2 * 0)) <= Alt_PI * / 41 < 1 + 1 + 1H:sum_f_R0 (tg_alt PI_tg) (S (2 * 0)) <= Alt_PI * / 4sum_f_R0 (tg_alt PI_tg) (S (2 * 0)) <= Alt_PI * / 4H:sum_f_R0 (tg_alt PI_tg) (S (2 * 0)) <= Alt_PI * / 41 < 1 + 1 + 1H:sum_f_R0 (tg_alt PI_tg) (S (2 * 0)) <= Alt_PI * / 4sum_f_R0 (tg_alt PI_tg) (S (2 * 0)) <= Alt_PI * / 4assumption. Qed.H:sum_f_R0 (tg_alt PI_tg) (S (2 * 0)) <= Alt_PI * / 4sum_f_R0 (tg_alt PI_tg) (S (2 * 0)) <= Alt_PI * / 4