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 Lra.
Require Import Rbase.
Require Import PSeries_reg.
Require Import Rtrigo1.
Require Import Ranalysis_reg.
Require Import Rfunctions.
Require Import AltSeries.
Require Import Rseries.
Require Import SeqProp.
Require Import Ranalysis5.
Require Import SeqSeries.
Require Import PartSum.
Require Import Omega.

Local Open Scope R_scope.
Tools

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

forall x y : R, - x / y = - (x / y)
intros x y; unfold Rdiv; rewrite <-Ropp_mult_distr_l_reverse; reflexivity. Qed.

0 < / 2

0 < / 2
lra. Qed. Definition pos_half := mkposreal (/2) pos_half_prf.

forall x : R, Boule (/ 2) pos_half x -> 0 <= x <= 1

forall x : R, Boule (/ 2) pos_half x -> 0 <= x <= 1

forall x : R, Rabs (x - / 2) < / 2 -> 0 <= x <= 1
intros x b; apply Rabs_def2 in b; destruct b; split; lra. Qed.

forall (c : R) (r : posreal) (x : R), Boule c r x -> Rabs x < Rabs c + r

forall (c : R) (r : posreal) (x : R), Boule c r x -> Rabs x < Rabs c + r
c:R
r:posreal
x:R
h:Rabs (x - c) < r

Rabs x < Rabs c + r
apply Rabs_def2 in h; destruct h; apply Rabs_def1; (destruct (Rle_lt_dec 0 c);[rewrite Rabs_pos_eq; lra | rewrite <- Rabs_Ropp, Rabs_pos_eq; lra]). Qed. (* The following lemma does not belong here. *)

forall un vn : nat -> R, (forall n : nat, un n = vn n) -> forall l : R, Un_cv un l -> Un_cv vn l

forall un vn : nat -> R, (forall n : nat, un n = vn n) -> forall l : R, Un_cv un l -> Un_cv vn l
un, vn:nat -> R
quv:forall n : nat, un n = vn n
l:R
P:Un_cv un l
eps:R
ep:eps > 0
N:nat
Pn:forall n : nat, (n >= N)%nat -> R_dist (un n) l < eps

forall n : nat, (n >= N)%nat -> R_dist (vn n) l < eps
intro n; rewrite <- quv; apply Pn. Qed. (* The following two lemmas are general purposes about alternated series. They do not belong here. *)

forall (f : nat -> R) (l : R) (N n : nat), Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (N <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f N

forall (f : nat -> R) (l : R) (N n : nat), Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (N <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f N
f:nat -> R
l:R

forall N n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (N <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f N
f:nat -> R
l:R

forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
forall N n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (N <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f N

forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
forall N n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (N <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f N
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n

forall N n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (N <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f N
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n

forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
Npos:(0 < 0)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(0 <= n)%nat

R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S N)
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat

R_dist (sum_f_R0 (tg_alt f) n) l <= f (S N)
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat

Un_decreasing (fun i : nat => f (S N + i)%nat)
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S N)
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k0 : nat, (0 < k0)%nat -> P k0) -> ((forall k0 : nat, (0 < k0)%nat -> P k0) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
k:nat

f (S (S N + k)) <= f (S N + k)%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S N)
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)

R_dist (sum_f_R0 (tg_alt f) n) l <= f (S N)
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)

Un_cv (fun i : nat => f (S N + i)%nat) 0
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S N)
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
eps:R
ep:eps > 0
M:nat
PM:forall n0 : nat, (n0 >= M)%nat -> R_dist (f n0) 0 < eps

exists N0 : nat, forall n0 : nat, (n0 >= N0)%nat -> R_dist (f (S N + n0)%nat) 0 < eps
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S N)
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0

R_dist (sum_f_R0 (tg_alt f) n) l <= f (S N)
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0

Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S N)
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
eps:R
ep:eps > 0
M:nat
PM:forall n0 : nat, (n0 >= M)%nat -> R_dist (sum_f_R0 (tg_alt f) n0) l < eps

forall n0 : nat, (n0 >= M)%nat -> R_dist (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat)) n0) (l - sum_f_R0 (tg_alt f) N) < eps
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S N)
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
eps:R
ep:eps > 0
M:nat
PM:forall n0 : nat, (n0 >= M)%nat -> R_dist (sum_f_R0 (tg_alt f) n0) l < eps
n':nat
nM:(n' >= M)%nat

R_dist (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat)) n') (l - sum_f_R0 (tg_alt f) N) < eps
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S N)
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
eps:R
ep:eps > 0
M:nat
PM:forall n0 : nat, (n0 >= M)%nat -> R_dist (sum_f_R0 (tg_alt f) n0) l < eps
n':nat
nM:(n' >= M)%nat
U:=R_dist (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat)) n') (l - sum_f_R0 (tg_alt f) N) < eps:Prop

U
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S N)
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
eps:R
ep:eps > 0
M:nat
PM:forall n0 : nat, (n0 >= M)%nat -> R_dist (sum_f_R0 (tg_alt f) n0) l < eps
n':nat
nM:(n' >= M)%nat
U:=R_dist (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat)) n') (l - sum_f_R0 (tg_alt f) N) < eps:Prop
nM':(n' + S N >= M)%nat

U
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S N)
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
eps:R
ep:eps > 0
M:nat
PM:forall n0 : nat, (n0 >= M)%nat -> R_dist (sum_f_R0 (tg_alt f) n0) l < eps
n':nat
nM:(n' >= M)%nat
U:=R_dist (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat)) n') (l - sum_f_R0 (tg_alt f) N) < eps:Prop
nM':(n' + S N >= M)%nat

Rabs (sum_f_R0 (tg_alt f) (n' + S N) - l) < eps -> U
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S N)
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
eps:R
ep:eps > 0
M:nat
PM:forall n0 : nat, (n0 >= M)%nat -> R_dist (sum_f_R0 (tg_alt f) n0) l < eps
n':nat
nM:(n' >= M)%nat
U:=R_dist (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat)) n') (l - sum_f_R0 (tg_alt f) N) < eps:Prop
nM':(n' + S N >= M)%nat

Rabs (sum_f_R0 (tg_alt f) N + sum_f_R0 (fun i : nat => tg_alt f (S N + i)) (n' + S N - S N) - l) < eps -> U
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
eps:R
ep:eps > 0
M:nat
PM:forall n0 : nat, (n0 >= M)%nat -> R_dist (sum_f_R0 (tg_alt f) n0) l < eps
n':nat
nM:(n' >= M)%nat
U:=R_dist (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat)) n') (l - sum_f_R0 (tg_alt f) N) < eps:Prop
nM':(n' + S N >= M)%nat
(N < n' + S N)%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S N)
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
eps:R
ep:eps > 0
M:nat
PM:forall n0 : nat, (n0 >= M)%nat -> R_dist (sum_f_R0 (tg_alt f) n0) l < eps
n':nat
nM:(n' >= M)%nat
U:=R_dist (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat)) n') (l - sum_f_R0 (tg_alt f) N) < eps:Prop
nM':(n' + S N >= M)%nat
t:forall a b c : R, a + b - c = b - (c - a)

Rabs (sum_f_R0 (tg_alt f) N + sum_f_R0 (fun i : nat => tg_alt f (S N + i)) (n' + S N - S N) - l) < eps -> U
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
eps:R
ep:eps > 0
M:nat
PM:forall n0 : nat, (n0 >= M)%nat -> R_dist (sum_f_R0 (tg_alt f) n0) l < eps
n':nat
nM:(n' >= M)%nat
U:=R_dist (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat)) n') (l - sum_f_R0 (tg_alt f) N) < eps:Prop
nM':(n' + S N >= M)%nat
(N < n' + S N)%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S N)
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
eps:R
ep:eps > 0
M:nat
PM:forall n0 : nat, (n0 >= M)%nat -> R_dist (sum_f_R0 (tg_alt f) n0) l < eps
n':nat
nM:(n' >= M)%nat
nM':(n' + S N >= M)%nat

Rabs (sum_f_R0 (fun i : nat => tg_alt f (S N + i)) (n' + S N - S N) - (l - sum_f_R0 (tg_alt f) N)) < eps -> Rabs (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat)) n' - (l - sum_f_R0 (tg_alt f) N)) < eps
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
eps:R
ep:eps > 0
M:nat
PM:forall n0 : nat, (n0 >= M)%nat -> R_dist (sum_f_R0 (tg_alt f) n0) l < eps
n':nat
nM:(n' >= M)%nat
U:=R_dist (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat)) n') (l - sum_f_R0 (tg_alt f) N) < eps:Prop
nM':(n' + S N >= M)%nat
(N < n' + S N)%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S N)
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
eps:R
ep:eps > 0
M:nat
PM:forall n0 : nat, (n0 >= M)%nat -> R_dist (sum_f_R0 (tg_alt f) n0) l < eps
n':nat
nM:(n' >= M)%nat
nM':(n' + S N >= M)%nat

Rabs (sum_f_R0 (fun i : nat => tg_alt f (S N + i)) n' - (l - sum_f_R0 (tg_alt f) N)) < eps -> Rabs (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat)) n' - (l - sum_f_R0 (tg_alt f) N)) < eps
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
eps:R
ep:eps > 0
M:nat
PM:forall n0 : nat, (n0 >= M)%nat -> R_dist (sum_f_R0 (tg_alt f) n0) l < eps
n':nat
nM:(n' >= M)%nat
U:=R_dist (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat)) n') (l - sum_f_R0 (tg_alt f) N) < eps:Prop
nM':(n' + S N >= M)%nat
(N < n' + S N)%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S N)
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
eps:R
ep:eps > 0
M:nat
PM:forall n0 : nat, (n0 >= M)%nat -> R_dist (sum_f_R0 (tg_alt f) n0) l < eps
n':nat
nM:(n' >= M)%nat
nM':(n' + S N >= M)%nat

Rabs (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat)) n' - (l - sum_f_R0 (tg_alt f) N)) < eps -> Rabs (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat)) n' - (l - sum_f_R0 (tg_alt f) N)) < eps
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
eps:R
ep:eps > 0
M:nat
PM:forall n0 : nat, (n0 >= M)%nat -> R_dist (sum_f_R0 (tg_alt f) n0) l < eps
n':nat
nM:(n' >= M)%nat
nM':(n' + S N >= M)%nat
forall i : nat, (i <= n')%nat -> tg_alt (fun i0 : nat => (-1) ^ S N * f (S N + i0)%nat) i = tg_alt f (S N + i)
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
eps:R
ep:eps > 0
M:nat
PM:forall n0 : nat, (n0 >= M)%nat -> R_dist (sum_f_R0 (tg_alt f) n0) l < eps
n':nat
nM:(n' >= M)%nat
U:=R_dist (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat)) n') (l - sum_f_R0 (tg_alt f) N) < eps:Prop
nM':(n' + S N >= M)%nat
(N < n' + S N)%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S N)
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
eps:R
ep:eps > 0
M:nat
PM:forall n0 : nat, (n0 >= M)%nat -> R_dist (sum_f_R0 (tg_alt f) n0) l < eps
n':nat
nM:(n' >= M)%nat
nM':(n' + S N >= M)%nat

forall i : nat, (i <= n')%nat -> tg_alt (fun i0 : nat => (-1) ^ S N * f (S N + i0)%nat) i = tg_alt f (S N + i)
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
eps:R
ep:eps > 0
M:nat
PM:forall n0 : nat, (n0 >= M)%nat -> R_dist (sum_f_R0 (tg_alt f) n0) l < eps
n':nat
nM:(n' >= M)%nat
U:=R_dist (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat)) n') (l - sum_f_R0 (tg_alt f) N) < eps:Prop
nM':(n' + S N >= M)%nat
(N < n' + S N)%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S N)
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i0 : nat => f (S N + i0)%nat)
to':Un_cv (fun i0 : nat => f (S N + i0)%nat) 0
eps:R
ep:eps > 0
M:nat
PM:forall n0 : nat, (n0 >= M)%nat -> R_dist (sum_f_R0 (tg_alt f) n0) l < eps
n':nat
nM:(n' >= M)%nat
nM':(n' + S N >= M)%nat
i:nat

(-1) ^ i * ((-1) ^ S N * f (S N + i)%nat) = (-1) ^ (S N + i) * f (S N + i)%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
eps:R
ep:eps > 0
M:nat
PM:forall n0 : nat, (n0 >= M)%nat -> R_dist (sum_f_R0 (tg_alt f) n0) l < eps
n':nat
nM:(n' >= M)%nat
U:=R_dist (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat)) n') (l - sum_f_R0 (tg_alt f) N) < eps:Prop
nM':(n' + S N >= M)%nat
(N < n' + S N)%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S N)
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
eps:R
ep:eps > 0
M:nat
PM:forall n0 : nat, (n0 >= M)%nat -> R_dist (sum_f_R0 (tg_alt f) n0) l < eps
n':nat
nM:(n' >= M)%nat
U:=R_dist (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat)) n') (l - sum_f_R0 (tg_alt f) N) < eps:Prop
nM':(n' + S N >= M)%nat

(N < n' + S N)%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S N)
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)

R_dist (sum_f_R0 (tg_alt f) n) l <= f (S N)
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)

Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S N)
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)

forall n0 : nat, (-1) ^ S N * sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat)) n0 = sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat)) n0
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
Un_cv (fun n0 : nat => (-1) ^ S N * sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat)) n0) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S N)
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n1 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n1
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i0 : nat => f (S N + i0)%nat)
to':Un_cv (fun i0 : nat => f (S N + i0)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i0 : nat => (-1) ^ S N * f (S N + i0)%nat))) (l - sum_f_R0 (tg_alt f) N)
n0, i:nat

tg_alt (fun i0 : nat => (-1) ^ S N * f (S N + i0)%nat) i * (-1) ^ S N = tg_alt (fun i0 : nat => f (S N + i0)%nat) i
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
Un_cv (fun n0 : nat => (-1) ^ S N * sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat)) n0) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S N)
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n1 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n1
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i0 : nat => f (S N + i0)%nat)
to':Un_cv (fun i0 : nat => f (S N + i0)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i0 : nat => (-1) ^ S N * f (S N + i0)%nat))) (l - sum_f_R0 (tg_alt f) N)
n0, i:nat

(-1) ^ i * 1 * f (S N + i)%nat = (-1) ^ i * f (S N + i)%nat
f:nat -> R
l:R
WLOG:forall (n1 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n1
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i0 : nat => f (S N + i0)%nat)
to':Un_cv (fun i0 : nat => f (S N + i0)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i0 : nat => (-1) ^ S N * f (S N + i0)%nat))) (l - sum_f_R0 (tg_alt f) N)
n0, i:nat
1 = ((-1) ^ S N) ^ 2
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
Un_cv (fun n0 : nat => (-1) ^ S N * sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat)) n0) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S N)
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n1 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n1
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i0 : nat => f (S N + i0)%nat)
to':Un_cv (fun i0 : nat => f (S N + i0)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i0 : nat => (-1) ^ S N * f (S N + i0)%nat))) (l - sum_f_R0 (tg_alt f) N)
n0, i:nat

1 = ((-1) ^ S N) ^ 2
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
Un_cv (fun n0 : nat => (-1) ^ S N * sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat)) n0) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S N)
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n1 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n1
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i0 : nat => f (S N + i0)%nat)
to':Un_cv (fun i0 : nat => f (S N + i0)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i0 : nat => (-1) ^ S N * f (S N + i0)%nat))) (l - sum_f_R0 (tg_alt f) N)
n0, i:nat

1 = 1 ^ S N
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
Un_cv (fun n0 : nat => (-1) ^ S N * sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat)) n0) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S N)
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)

Un_cv (fun n0 : nat => (-1) ^ S N * sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat)) n0) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S N)
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)

Un_cv (fun _ : nat => (-1) ^ S N) ((-1) ^ S N)
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S N)
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)

Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S N)
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))

R_dist (sum_f_R0 (tg_alt f) n) l <= f (S N)
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Neven:N = (2 * p)%nat

R_dist (sum_f_R0 (tg_alt f) n) l <= f (S N)
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Nodd:N = S (2 * p)
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S N)
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Neven:N = (2 * p)%nat
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * p)

R_dist (sum_f_R0 (tg_alt f) n) l <= f (S (2 * p))
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Nodd:N = S (2 * p)
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S N)
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Neven:N = (2 * p)%nat
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * p)
p':nat
neven:n = (2 * p')%nat

R_dist (sum_f_R0 (tg_alt f) n) l <= f (S (2 * p))
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Neven:N = (2 * p)%nat
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * p)
p':nat
nodd:n = S (2 * p')
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S (2 * p))
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Nodd:N = S (2 * p)
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S N)
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Neven:N = (2 * p)%nat
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * p)
p':nat
neven:n = (2 * p')%nat

R_dist (sum_f_R0 (tg_alt f) (2 * p')) l <= f (S (2 * p))
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Neven:N = (2 * p)%nat
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * p)
p':nat
nodd:n = S (2 * p')
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S (2 * p))
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Nodd:N = S (2 * p)
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S N)
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Neven:N = (2 * p)%nat
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * p)
p':nat
neven:n = (2 * p')%nat
D:sum_f_R0 (tg_alt f) (S (2 * p')) <= l
E:l <= sum_f_R0 (tg_alt f) (2 * p')

R_dist (sum_f_R0 (tg_alt f) (2 * p')) l <= f (S (2 * p))
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Neven:N = (2 * p)%nat
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * p)
p':nat
nodd:n = S (2 * p')
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S (2 * p))
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Nodd:N = S (2 * p)
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S N)
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Neven:N = (2 * p)%nat
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * p)
p':nat
neven:n = (2 * p')%nat
D:sum_f_R0 (tg_alt f) (S (2 * p')) <= l
E:l <= sum_f_R0 (tg_alt f) (2 * p')

sum_f_R0 (tg_alt f) (2 * p') - l <= f (S (2 * p))
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Neven:N = (2 * p)%nat
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * p)
p':nat
nodd:n = S (2 * p')
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S (2 * p))
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Nodd:N = S (2 * p)
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S N)
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Neven:N = (2 * p)%nat
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * p)
p':nat
neven:n = (2 * p')%nat
D:sum_f_R0 (tg_alt f) (S (2 * p')) <= l
E:l <= sum_f_R0 (tg_alt f) (2 * p')
dist:(p <= p')%nat

sum_f_R0 (tg_alt f) (2 * p') - l <= f (S (2 * p))
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Neven:N = (2 * p)%nat
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * p)
p':nat
nodd:n = S (2 * p')
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S (2 * p))
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Nodd:N = S (2 * p)
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S N)
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Neven:N = (2 * p)%nat
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * p)
p':nat
neven:n = (2 * p')%nat
D:sum_f_R0 (tg_alt f) (S (2 * p')) <= l
E:l <= sum_f_R0 (tg_alt f) (2 * p')
dist:(p <= p')%nat
t:(fun N0 : nat => sum_f_R0 (tg_alt f) (2 * N0)) p' <= (fun N0 : nat => sum_f_R0 (tg_alt f) (2 * N0)) p

sum_f_R0 (tg_alt f) (2 * p') - l <= f (S (2 * p))
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Neven:N = (2 * p)%nat
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * p)
p':nat
nodd:n = S (2 * p')
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S (2 * p))
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Nodd:N = S (2 * p)
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S N)
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Neven:N = (2 * p)%nat
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * p)
p':nat
neven:n = (2 * p')%nat
D:sum_f_R0 (tg_alt f) (S (2 * p')) <= l
E:l <= sum_f_R0 (tg_alt f) (2 * p')
dist:(p <= p')%nat
t:(fun N0 : nat => sum_f_R0 (tg_alt f) (2 * N0)) p' <= (fun N0 : nat => sum_f_R0 (tg_alt f) (2 * N0)) p

sum_f_R0 (tg_alt f) (2 * p') - l <= sum_f_R0 (tg_alt f) (2 * p) - l
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Neven:N = (2 * p)%nat
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * p)
p':nat
neven:n = (2 * p')%nat
D:sum_f_R0 (tg_alt f) (S (2 * p')) <= l
E:l <= sum_f_R0 (tg_alt f) (2 * p')
dist:(p <= p')%nat
t:(fun N0 : nat => sum_f_R0 (tg_alt f) (2 * N0)) p' <= (fun N0 : nat => sum_f_R0 (tg_alt f) (2 * N0)) p
sum_f_R0 (tg_alt f) (2 * p) - l <= f (S (2 * p))
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Neven:N = (2 * p)%nat
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * p)
p':nat
nodd:n = S (2 * p')
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S (2 * p))
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Nodd:N = S (2 * p)
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S N)
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Neven:N = (2 * p)%nat
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * p)
p':nat
neven:n = (2 * p')%nat
D:sum_f_R0 (tg_alt f) (S (2 * p')) <= l
E:l <= sum_f_R0 (tg_alt f) (2 * p')
dist:(p <= p')%nat
t:(fun N0 : nat => sum_f_R0 (tg_alt f) (2 * N0)) p' <= (fun N0 : nat => sum_f_R0 (tg_alt f) (2 * N0)) p

sum_f_R0 (tg_alt f) (2 * p) - l <= f (S (2 * p))
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Neven:N = (2 * p)%nat
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * p)
p':nat
nodd:n = S (2 * p')
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S (2 * p))
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Nodd:N = S (2 * p)
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S N)
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Neven:N = (2 * p)%nat
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * p)
p':nat
nodd:n = S (2 * p')

R_dist (sum_f_R0 (tg_alt f) n) l <= f (S (2 * p))
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Nodd:N = S (2 * p)
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S N)
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Neven:N = (2 * p)%nat
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * p)
p':nat
nodd:n = S (2 * p')
D:sum_f_R0 (tg_alt f) (S (2 * p')) <= l
E:l <= sum_f_R0 (tg_alt f) (2 * p')

R_dist (sum_f_R0 (tg_alt f) (S (2 * p'))) l <= f (S (2 * p))
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Nodd:N = S (2 * p)
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S N)
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Neven:N = (2 * p)%nat
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * p)
p':nat
nodd:n = S (2 * p')
D:sum_f_R0 (tg_alt f) (S (2 * p')) <= l
E:l <= sum_f_R0 (tg_alt f) (2 * p')

l - sum_f_R0 (tg_alt f) (S (2 * p')) <= f (S (2 * p))
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Nodd:N = S (2 * p)
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S N)
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Neven:N = (2 * p)%nat
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * p)
p':nat
nodd:n = S (2 * p')
D:sum_f_R0 (tg_alt f) (S (2 * p')) <= l
E:l <= sum_f_R0 (tg_alt f) (2 * p')
dist:(p <= p')%nat

l - sum_f_R0 (tg_alt f) (S (2 * p')) <= f (S (2 * p))
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Nodd:N = S (2 * p)
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S N)
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Neven:N = (2 * p)%nat
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * p)
p':nat
nodd:n = S (2 * p')
D:sum_f_R0 (tg_alt f) (S (2 * p')) <= l
E:l <= sum_f_R0 (tg_alt f) (2 * p')
dist:(p <= p')%nat

l - sum_f_R0 (tg_alt f) (S (2 * p')) <= l - sum_f_R0 (tg_alt f) (S (2 * p))
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Neven:N = (2 * p)%nat
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * p)
p':nat
nodd:n = S (2 * p')
D:sum_f_R0 (tg_alt f) (S (2 * p')) <= l
E:l <= sum_f_R0 (tg_alt f) (2 * p')
dist:(p <= p')%nat
l - sum_f_R0 (tg_alt f) (S (2 * p)) <= f (S (2 * p))
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Nodd:N = S (2 * p)
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S N)
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Neven:N = (2 * p)%nat
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * p)
p':nat
nodd:n = S (2 * p')
D:sum_f_R0 (tg_alt f) (S (2 * p')) <= l
E:l <= sum_f_R0 (tg_alt f) (2 * p')
dist:(p <= p')%nat

sum_f_R0 (tg_alt f) (S (2 * p)) <= sum_f_R0 (tg_alt f) (S (2 * p'))
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Neven:N = (2 * p)%nat
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * p)
p':nat
nodd:n = S (2 * p')
D:sum_f_R0 (tg_alt f) (S (2 * p')) <= l
E:l <= sum_f_R0 (tg_alt f) (2 * p')
dist:(p <= p')%nat
l - sum_f_R0 (tg_alt f) (S (2 * p)) <= f (S (2 * p))
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Nodd:N = S (2 * p)
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S N)
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Neven:N = (2 * p)%nat
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * p)
p':nat
nodd:n = S (2 * p')
D:sum_f_R0 (tg_alt f) (S (2 * p')) <= l
E:l <= sum_f_R0 (tg_alt f) (2 * p')
dist:(p <= p')%nat

l - sum_f_R0 (tg_alt f) (S (2 * p)) <= f (S (2 * p))
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Nodd:N = S (2 * p)
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S N)
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Neven:N = (2 * p)%nat
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * p)
p':nat
nodd:n = S (2 * p')
D:sum_f_R0 (tg_alt f) (S (2 * p')) <= l
E:l <= sum_f_R0 (tg_alt f) (2 * p')
dist:(p <= p')%nat

l + - sum_f_R0 (tg_alt f) (2 * p) + - tg_alt f (S (2 * p)) <= f (S (2 * p))
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Nodd:N = S (2 * p)
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S N)
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Nodd:N = S (2 * p)

R_dist (sum_f_R0 (tg_alt f) n) l <= f (S N)
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Nodd:N = S (2 * p)
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l

R_dist (sum_f_R0 (tg_alt f) n) l <= f (S (S (2 * p)))
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Nodd:N = S (2 * p)
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * S p)

R_dist (sum_f_R0 (tg_alt f) n) l <= f (S (S (2 * p)))
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Nodd:N = S (2 * p)
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * S p)
keep:(2 * S p)%nat = S (S (2 * p))

R_dist (sum_f_R0 (tg_alt f) n) l <= f (S (S (2 * p)))
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Nodd:N = S (2 * p)
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * S p)
keep:(2 * S p)%nat = S (S (2 * p))
p':nat
neven:n = (2 * p')%nat

R_dist (sum_f_R0 (tg_alt f) n) l <= f (S (S (2 * p)))
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Nodd:N = S (2 * p)
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * S p)
keep:(2 * S p)%nat = S (S (2 * p))
p':nat
nodd:n = S (2 * p')
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S (S (2 * p)))
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Nodd:N = S (2 * p)
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * S p)
keep:(2 * S p)%nat = S (S (2 * p))
p':nat
neven:n = (2 * p')%nat
D:sum_f_R0 (tg_alt f) (S (2 * p')) <= l
E:l <= sum_f_R0 (tg_alt f) (2 * p')

R_dist (sum_f_R0 (tg_alt f) (2 * p')) l <= f (S (S (2 * p)))
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Nodd:N = S (2 * p)
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * S p)
keep:(2 * S p)%nat = S (S (2 * p))
p':nat
nodd:n = S (2 * p')
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S (S (2 * p)))
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Nodd:N = S (2 * p)
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * S p)
keep:(2 * S p)%nat = S (S (2 * p))
p':nat
neven:n = (2 * p')%nat
D:sum_f_R0 (tg_alt f) (S (2 * p')) <= l
E:l <= sum_f_R0 (tg_alt f) (2 * p')

sum_f_R0 (tg_alt f) (2 * p') - l <= f (S (S (2 * p)))
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Nodd:N = S (2 * p)
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * S p)
keep:(2 * S p)%nat = S (S (2 * p))
p':nat
nodd:n = S (2 * p')
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S (S (2 * p)))
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Nodd:N = S (2 * p)
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * S p)
keep:(2 * S p)%nat = S (S (2 * p))
p':nat
neven:n = (2 * p')%nat
D:sum_f_R0 (tg_alt f) (S (2 * p')) <= l
E:l <= sum_f_R0 (tg_alt f) (2 * p')
dist:(S p < S p')%nat

sum_f_R0 (tg_alt f) (2 * p') - l <= f (S (S (2 * p)))
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Nodd:N = S (2 * p)
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * S p)
keep:(2 * S p)%nat = S (S (2 * p))
p':nat
nodd:n = S (2 * p')
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S (S (2 * p)))
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Nodd:N = S (2 * p)
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * S p)
keep:(2 * S p)%nat = S (S (2 * p))
p':nat
neven:n = (2 * p')%nat
D:sum_f_R0 (tg_alt f) (S (2 * p')) <= l
E:l <= sum_f_R0 (tg_alt f) (2 * p')
dist:(S p < S p')%nat

sum_f_R0 (tg_alt f) (2 * p') - l <= sum_f_R0 (tg_alt f) (2 * S p) - l
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Nodd:N = S (2 * p)
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * S p)
keep:(2 * S p)%nat = S (S (2 * p))
p':nat
neven:n = (2 * p')%nat
D:sum_f_R0 (tg_alt f) (S (2 * p')) <= l
E:l <= sum_f_R0 (tg_alt f) (2 * p')
dist:(S p < S p')%nat
sum_f_R0 (tg_alt f) (2 * S p) - l <= f (S (S (2 * p)))
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Nodd:N = S (2 * p)
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * S p)
keep:(2 * S p)%nat = S (S (2 * p))
p':nat
nodd:n = S (2 * p')
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S (S (2 * p)))
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Nodd:N = S (2 * p)
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * S p)
keep:(2 * S p)%nat = S (S (2 * p))
p':nat
neven:n = (2 * p')%nat
D:sum_f_R0 (tg_alt f) (S (2 * p')) <= l
E:l <= sum_f_R0 (tg_alt f) (2 * p')
dist:(S p < S p')%nat

(S p <= p')%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Nodd:N = S (2 * p)
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * S p)
keep:(2 * S p)%nat = S (S (2 * p))
p':nat
neven:n = (2 * p')%nat
D:sum_f_R0 (tg_alt f) (S (2 * p')) <= l
E:l <= sum_f_R0 (tg_alt f) (2 * p')
dist:(S p < S p')%nat
sum_f_R0 (tg_alt f) (2 * S p) - l <= f (S (S (2 * p)))
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Nodd:N = S (2 * p)
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * S p)
keep:(2 * S p)%nat = S (S (2 * p))
p':nat
nodd:n = S (2 * p')
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S (S (2 * p)))
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Nodd:N = S (2 * p)
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * S p)
keep:(2 * S p)%nat = S (S (2 * p))
p':nat
neven:n = (2 * p')%nat
D:sum_f_R0 (tg_alt f) (S (2 * p')) <= l
E:l <= sum_f_R0 (tg_alt f) (2 * p')
dist:(S p < S p')%nat

sum_f_R0 (tg_alt f) (2 * S p) - l <= f (S (S (2 * p)))
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Nodd:N = S (2 * p)
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * S p)
keep:(2 * S p)%nat = S (S (2 * p))
p':nat
nodd:n = S (2 * p')
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S (S (2 * p)))
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Nodd:N = S (2 * p)
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * S p)
keep:(2 * S p)%nat = S (S (2 * p))
p':nat
neven:n = (2 * p')%nat
D:sum_f_R0 (tg_alt f) (S (2 * p')) <= l
E:l <= sum_f_R0 (tg_alt f) (2 * p')
dist:(S p < S p')%nat

sum_f_R0 (tg_alt f) (S (2 * p)) + 1 * f (2 * S p)%nat - l <= f (2 * S p)%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Nodd:N = S (2 * p)
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * S p)
keep:(2 * S p)%nat = S (S (2 * p))
p':nat
nodd:n = S (2 * p')
R_dist (sum_f_R0 (tg_alt f) n) l <= f (S (S (2 * p)))
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Nodd:N = S (2 * p)
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * S p)
keep:(2 * S p)%nat = S (S (2 * p))
p':nat
nodd:n = S (2 * p')

R_dist (sum_f_R0 (tg_alt f) n) l <= f (S (S (2 * p)))
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Nodd:N = S (2 * p)
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * S p)
keep:(2 * S p)%nat = S (S (2 * p))
p':nat
nodd:n = S (2 * p')
D:sum_f_R0 (tg_alt f) (S (2 * p')) <= l
E:l <= sum_f_R0 (tg_alt f) (2 * p')

R_dist (sum_f_R0 (tg_alt f) (S (2 * p'))) l <= f (S (S (2 * p)))
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Nodd:N = S (2 * p)
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * S p)
keep:(2 * S p)%nat = S (S (2 * p))
p':nat
nodd:n = S (2 * p')
D:sum_f_R0 (tg_alt f) (S (2 * p')) <= l
E:l <= sum_f_R0 (tg_alt f) (2 * p')

- (sum_f_R0 (tg_alt f) (S (2 * p')) - l) <= f (S (S (2 * p)))
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Nodd:N = S (2 * p)
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * S p)
keep:(2 * S p)%nat = S (S (2 * p))
p':nat
nodd:n = S (2 * p')
D:sum_f_R0 (tg_alt f) (S (2 * p')) <= l
E:l <= sum_f_R0 (tg_alt f) (2 * p')

l - sum_f_R0 (tg_alt f) (S (2 * p')) <= f (S (S (2 * p)))
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Nodd:N = S (2 * p)
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * S p)
keep:(2 * S p)%nat = S (S (2 * p))
p':nat
nodd:n = S (2 * p')
D:sum_f_R0 (tg_alt f) (S (2 * p')) <= l
E:l <= sum_f_R0 (tg_alt f) (2 * p')

l - sum_f_R0 (tg_alt f) (S (2 * p')) <= l - sum_f_R0 (tg_alt f) (S (2 * p))
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Nodd:N = S (2 * p)
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * S p)
keep:(2 * S p)%nat = S (S (2 * p))
p':nat
nodd:n = S (2 * p')
D:sum_f_R0 (tg_alt f) (S (2 * p')) <= l
E:l <= sum_f_R0 (tg_alt f) (2 * p')
l - sum_f_R0 (tg_alt f) (S (2 * p)) <= f (S (S (2 * p)))
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Nodd:N = S (2 * p)
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * S p)
keep:(2 * S p)%nat = S (S (2 * p))
p':nat
nodd:n = S (2 * p')
D:sum_f_R0 (tg_alt f) (S (2 * p')) <= l
E:l <= sum_f_R0 (tg_alt f) (2 * p')

l - sum_f_R0 (tg_alt f) (S (2 * p)) <= f (S (S (2 * p)))
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Nodd:N = S (2 * p)
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * S p)
keep:(2 * S p)%nat = S (S (2 * p))
p':nat
nodd:n = S (2 * p')
D:sum_f_R0 (tg_alt f) (S (2 * p')) <= l
E:l <= sum_f_R0 (tg_alt f) (2 * p')

l <= sum_f_R0 (fun i : nat => (-1) ^ i * f i) (S (2 * p)) + (-1) ^ S (S (2 * p)) * f (S (S (2 * p))) -> l - sum_f_R0 (fun i : nat => (-1) ^ i * f i) (S (2 * p)) <= f (S (S (2 * p)))
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Nodd:N = S (2 * p)
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * S p)
keep:(2 * S p)%nat = S (S (2 * p))
p':nat
nodd:n = S (2 * p')
D:sum_f_R0 (tg_alt f) (S (2 * p')) <= l
E:l <= sum_f_R0 (tg_alt f) (2 * p')

l <= sum_f_R0 (fun i : nat => (-1) ^ i * f i) (S (2 * p)) + 1 * f (2 * S p)%nat -> l - sum_f_R0 (fun i : nat => (-1) ^ i * f i) (S (2 * p)) <= f (2 * S p)%nat
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n0 : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n0
N:nat
Npos:(0 < S N)%nat
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
nN:(S N <= n)%nat
decr':Un_decreasing (fun i : nat => f (S N + i)%nat)
to':Un_cv (fun i : nat => f (S N + i)%nat) 0
cv':Un_cv (sum_f_R0 (tg_alt (fun i : nat => (-1) ^ S N * f (S N + i)%nat))) (l - sum_f_R0 (tg_alt f) N)
cv'':Un_cv (sum_f_R0 (tg_alt (fun i : nat => f (S N + i)%nat))) ((-1) ^ S N * (l - sum_f_R0 (tg_alt f) N))
p:nat
Nodd:N = S (2 * p)
B:sum_f_R0 (tg_alt f) (S (2 * p)) <= l
C:l <= sum_f_R0 (tg_alt f) (2 * S p)
keep:(2 * S p)%nat = S (S (2 * p))
p':nat
nodd:n = S (2 * p')
D:sum_f_R0 (tg_alt f) (S (2 * p')) <= l
E:l <= sum_f_R0 (tg_alt f) (2 * p')
t:forall a b c : R, a <= b + 1 * c -> a - b <= c

l <= sum_f_R0 (fun i : nat => (-1) ^ i * f i) (S (2 * p)) + 1 * f (2 * S p)%nat -> l - sum_f_R0 (fun i : nat => (-1) ^ i * f i) (S (2 * p)) <= f (2 * S p)%nat
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n
(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
WLOG:forall (n : nat) (P : nat -> Type), (forall k : nat, (0 < k)%nat -> P k) -> ((forall k : nat, (0 < k)%nat -> P k) -> P 0%nat) -> P n

(forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k) -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (0 <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f 0%nat
f:nat -> R
l:R
Hyp:forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l

R_dist (sum_f_R0 (tg_alt f) 0) l <= f 0%nat
f:nat -> R
l:R
Hyp:forall k : nat, (0 < k)%nat -> forall n0 : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n0)%nat -> R_dist (sum_f_R0 (tg_alt f) n0) l <= f k
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
R_dist (sum_f_R0 (tg_alt f) (S n)) l <= f 0%nat
f:nat -> R
l:R
Hyp:forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l

sum_f_R0 (tg_alt f) (S (2 * 0)) <= l <= sum_f_R0 (tg_alt f) (2 * 0) -> R_dist (sum_f_R0 (tg_alt f) 0) l <= f 0%nat
f:nat -> R
l:R
Hyp:forall k : nat, (0 < k)%nat -> forall n0 : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n0)%nat -> R_dist (sum_f_R0 (tg_alt f) n0) l <= f k
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
R_dist (sum_f_R0 (tg_alt f) (S n)) l <= f 0%nat
f:nat -> R
l:R
Hyp:forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l

f 0%nat + -1 * f 1%nat <= l <= f 0%nat -> Rabs (f 0%nat - l) <= f 0%nat
f:nat -> R
l:R
Hyp:forall k : nat, (0 < k)%nat -> forall n0 : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n0)%nat -> R_dist (sum_f_R0 (tg_alt f) n0) l <= f k
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
R_dist (sum_f_R0 (tg_alt f) (S n)) l <= f 0%nat
f:nat -> R
l:R
Hyp:forall k : nat, (0 < k)%nat -> forall n : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n)%nat -> R_dist (sum_f_R0 (tg_alt f) n) l <= f k
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
H:f 1%nat <= f 0%nat

f 0%nat + -1 * f 1%nat <= l <= f 0%nat -> Rabs (f 0%nat - l) <= f 0%nat
f:nat -> R
l:R
Hyp:forall k : nat, (0 < k)%nat -> forall n0 : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n0)%nat -> R_dist (sum_f_R0 (tg_alt f) n0) l <= f k
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
R_dist (sum_f_R0 (tg_alt f) (S n)) l <= f 0%nat
f:nat -> R
l:R
Hyp:forall k : nat, (0 < k)%nat -> forall n0 : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n0)%nat -> R_dist (sum_f_R0 (tg_alt f) n0) l <= f k
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l

R_dist (sum_f_R0 (tg_alt f) (S n)) l <= f 0%nat
f:nat -> R
l:R
Hyp:forall k : nat, (0 < k)%nat -> forall n0 : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n0)%nat -> R_dist (sum_f_R0 (tg_alt f) n0) l <= f k
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l

R_dist (sum_f_R0 (tg_alt f) (S n)) l <= f 1%nat
f:nat -> R
l:R
Hyp:forall k : nat, (0 < k)%nat -> forall n0 : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n0)%nat -> R_dist (sum_f_R0 (tg_alt f) n0) l <= f k
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
f 1%nat <= f 0%nat
f:nat -> R
l:R
Hyp:forall k : nat, (0 < k)%nat -> forall n0 : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n0)%nat -> R_dist (sum_f_R0 (tg_alt f) n0) l <= f k
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l

(1 <= S n)%nat
f:nat -> R
l:R
Hyp:forall k : nat, (0 < k)%nat -> forall n0 : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n0)%nat -> R_dist (sum_f_R0 (tg_alt f) n0) l <= f k
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l
f 1%nat <= f 0%nat
f:nat -> R
l:R
Hyp:forall k : nat, (0 < k)%nat -> forall n0 : nat, Un_decreasing f -> Un_cv f 0 -> Un_cv (sum_f_R0 (tg_alt f)) l -> (k <= n0)%nat -> R_dist (sum_f_R0 (tg_alt f) n0) l <= f k
n:nat
decr:Un_decreasing f
to0:Un_cv f 0
cv:Un_cv (sum_f_R0 (tg_alt f)) l

f 1%nat <= f 0%nat
solve[apply decr]. Qed.

forall (f : nat -> R -> R) (g : R -> R) (h : nat -> R) (c : R) (r : posreal), (forall x : R, Boule c r x -> Un_decreasing (fun n : nat => f n x)) -> (forall x : R, Boule c r x -> Un_cv (fun n : nat => f n x) 0) -> (forall x : R, Boule c r x -> Un_cv (sum_f_R0 (tg_alt (fun i : nat => f i x))) (g x)) -> (forall (x : R) (n : nat), Boule c r x -> f n x <= h n) -> Un_cv h 0 -> CVU (fun (N : nat) (x : R) => sum_f_R0 (tg_alt (fun i : nat => f i x)) N) g c r

forall (f : nat -> R -> R) (g : R -> R) (h : nat -> R) (c : R) (r : posreal), (forall x : R, Boule c r x -> Un_decreasing (fun n : nat => f n x)) -> (forall x : R, Boule c r x -> Un_cv (fun n : nat => f n x) 0) -> (forall x : R, Boule c r x -> Un_cv (sum_f_R0 (tg_alt (fun i : nat => f i x))) (g x)) -> (forall (x : R) (n : nat), Boule c r x -> f n x <= h n) -> Un_cv h 0 -> CVU (fun (N : nat) (x : R) => sum_f_R0 (tg_alt (fun i : nat => f i x)) N) g c r
f:nat -> R -> R
g:R -> R
h:nat -> R
c:R
r:posreal
decr:forall x : R, Boule c r x -> Un_decreasing (fun n : nat => f n x)
to0:forall x : R, Boule c r x -> Un_cv (fun n : nat => f n x) 0
to_g:forall x : R, Boule c r x -> Un_cv (sum_f_R0 (tg_alt (fun i : nat => f i x))) (g x)
bound:forall (x : R) (n : nat), Boule c r x -> f n x <= h n
bound0:Un_cv h 0
eps:R
ep:0 < eps

exists N : nat, forall (n : nat) (y : R), (N <= n)%nat -> Boule c r y -> Rabs (g y - sum_f_R0 (tg_alt (fun i : nat => f i y)) n) < eps
f:nat -> R -> R
g:R -> R
h:nat -> R
c:R
r:posreal
decr:forall x : R, Boule c r x -> Un_decreasing (fun n : nat => f n x)
to0:forall x : R, Boule c r x -> Un_cv (fun n : nat => f n x) 0
to_g:forall x : R, Boule c r x -> Un_cv (sum_f_R0 (tg_alt (fun i : nat => f i x))) (g x)
bound:forall (x : R) (n : nat), Boule c r x -> f n x <= h n
bound0:Un_cv h 0
eps:R
ep:0 < eps
ep':0 < eps / 2

exists N : nat, forall (n : nat) (y : R), (N <= n)%nat -> Boule c r y -> Rabs (g y - sum_f_R0 (tg_alt (fun i : nat => f i y)) n) < eps
f:nat -> R -> R
g:R -> R
h:nat -> R
c:R
r:posreal
decr:forall x : R, Boule c r x -> Un_decreasing (fun n : nat => f n x)
to0:forall x : R, Boule c r x -> Un_cv (fun n : nat => f n x) 0
to_g:forall x : R, Boule c r x -> Un_cv (sum_f_R0 (tg_alt (fun i : nat => f i x))) (g x)
bound:forall (x : R) (n : nat), Boule c r x -> f n x <= h n
bound0:Un_cv h 0
eps:R
ep:0 < eps
ep':0 < eps / 2
N:nat
Pn:forall n : nat, (n >= N)%nat -> R_dist (h n) 0 < eps

forall (n : nat) (y : R), (N <= n)%nat -> Boule c r y -> Rabs (g y - sum_f_R0 (tg_alt (fun i : nat => f i y)) n) < eps
f:nat -> R -> R
g:R -> R
h:nat -> R
c:R
r:posreal
decr:forall x : R, Boule c r x -> Un_decreasing (fun n0 : nat => f n0 x)
to0:forall x : R, Boule c r x -> Un_cv (fun n0 : nat => f n0 x) 0
to_g:forall x : R, Boule c r x -> Un_cv (sum_f_R0 (tg_alt (fun i : nat => f i x))) (g x)
bound:forall (x : R) (n0 : nat), Boule c r x -> f n0 x <= h n0
bound0:Un_cv h 0
eps:R
ep:0 < eps
ep':0 < eps / 2
N:nat
Pn:forall n0 : nat, (n0 >= N)%nat -> R_dist (h n0) 0 < eps
n:nat
y:R
nN:(N <= n)%nat
dy:Boule c r y

Rabs (g y - sum_f_R0 (tg_alt (fun i : nat => f i y)) n) < eps
f:nat -> R -> R
g:R -> R
h:nat -> R
c:R
r:posreal
decr:forall x : R, Boule c r x -> Un_decreasing (fun n0 : nat => f n0 x)
to0:forall x : R, Boule c r x -> Un_cv (fun n0 : nat => f n0 x) 0
to_g:forall x : R, Boule c r x -> Un_cv (sum_f_R0 (tg_alt (fun i : nat => f i x))) (g x)
bound:forall (x : R) (n0 : nat), Boule c r x -> f n0 x <= h n0
bound0:Un_cv h 0
eps:R
ep:0 < eps
ep':0 < eps / 2
N:nat
Pn:forall n0 : nat, (n0 >= N)%nat -> R_dist (h n0) 0 < eps
n:nat
y:R
nN:(N <= n)%nat
dy:Boule c r y

Rabs (sum_f_R0 (tg_alt (fun i : nat => f i y)) n - g y) <= f n y
f:nat -> R -> R
g:R -> R
h:nat -> R
c:R
r:posreal
decr:forall x : R, Boule c r x -> Un_decreasing (fun n0 : nat => f n0 x)
to0:forall x : R, Boule c r x -> Un_cv (fun n0 : nat => f n0 x) 0
to_g:forall x : R, Boule c r x -> Un_cv (sum_f_R0 (tg_alt (fun i : nat => f i x))) (g x)
bound:forall (x : R) (n0 : nat), Boule c r x -> f n0 x <= h n0
bound0:Un_cv h 0
eps:R
ep:0 < eps
ep':0 < eps / 2
N:nat
Pn:forall n0 : nat, (n0 >= N)%nat -> R_dist (h n0) 0 < eps
n:nat
y:R
nN:(N <= n)%nat
dy:Boule c r y
f n y < eps
f:nat -> R -> R
g:R -> R
h:nat -> R
c:R
r:posreal
decr:forall x : R, Boule c r x -> Un_decreasing (fun n0 : nat => f n0 x)
to0:forall x : R, Boule c r x -> Un_cv (fun n0 : nat => f n0 x) 0
to_g:forall x : R, Boule c r x -> Un_cv (sum_f_R0 (tg_alt (fun i : nat => f i x))) (g x)
bound:forall (x : R) (n0 : nat), Boule c r x -> f n0 x <= h n0
bound0:Un_cv h 0
eps:R
ep:0 < eps
ep':0 < eps / 2
N:nat
Pn:forall n0 : nat, (n0 >= N)%nat -> R_dist (h n0) 0 < eps
n:nat
y:R
nN:(N <= n)%nat
dy:Boule c r y

f n y < eps
f:nat -> R -> R
g:R -> R
h:nat -> R
c:R
r:posreal
decr:forall x : R, Boule c r x -> Un_decreasing (fun n0 : nat => f n0 x)
to0:forall x : R, Boule c r x -> Un_cv (fun n0 : nat => f n0 x) 0
to_g:forall x : R, Boule c r x -> Un_cv (sum_f_R0 (tg_alt (fun i : nat => f i x))) (g x)
bound:forall (x : R) (n0 : nat), Boule c r x -> f n0 x <= h n0
bound0:Un_cv h 0
eps:R
ep:0 < eps
ep':0 < eps / 2
N:nat
Pn:forall n0 : nat, (n0 >= N)%nat -> R_dist (h n0) 0 < eps
n:nat
y:R
nN:(N <= n)%nat
dy:Boule c r y

f n y <= h n
f:nat -> R -> R
g:R -> R
h:nat -> R
c:R
r:posreal
decr:forall x : R, Boule c r x -> Un_decreasing (fun n0 : nat => f n0 x)
to0:forall x : R, Boule c r x -> Un_cv (fun n0 : nat => f n0 x) 0
to_g:forall x : R, Boule c r x -> Un_cv (sum_f_R0 (tg_alt (fun i : nat => f i x))) (g x)
bound:forall (x : R) (n0 : nat), Boule c r x -> f n0 x <= h n0
bound0:Un_cv h 0
eps:R
ep:0 < eps
ep':0 < eps / 2
N:nat
Pn:forall n0 : nat, (n0 >= N)%nat -> R_dist (h n0) 0 < eps
n:nat
y:R
nN:(N <= n)%nat
dy:Boule c r y
h n < eps
f:nat -> R -> R
g:R -> R
h:nat -> R
c:R
r:posreal
decr:forall x : R, Boule c r x -> Un_decreasing (fun n0 : nat => f n0 x)
to0:forall x : R, Boule c r x -> Un_cv (fun n0 : nat => f n0 x) 0
to_g:forall x : R, Boule c r x -> Un_cv (sum_f_R0 (tg_alt (fun i : nat => f i x))) (g x)
bound:forall (x : R) (n0 : nat), Boule c r x -> f n0 x <= h n0
bound0:Un_cv h 0
eps:R
ep:0 < eps
ep':0 < eps / 2
N:nat
Pn:forall n0 : nat, (n0 >= N)%nat -> R_dist (h n0) 0 < eps
n:nat
y:R
nN:(N <= n)%nat
dy:Boule c r y

h n < eps
h:nat -> R
eps:R
N:nat
Pn:forall n0 : nat, (n0 >= N)%nat -> R_dist (h n0) 0 < eps
n:nat
nN:(N <= n)%nat

h n < eps
h:nat -> R
eps:R
N:nat
Pn:forall n0 : nat, (n0 >= N)%nat -> R_dist (h n0) 0 < eps
n:nat
nN:(N <= n)%nat
t:Rabs (h n) < eps

h n < eps
apply Rabs_def2 in t; tauto. Qed. (* The following lemmas are general purpose lemmas about squares. They do not belong here *)

forall x : R, 0 <= x ^ 2

forall x : R, 0 <= x ^ 2
x:R
r:0 <= x

0 <= x ^ 2
x:R
r:x < 0
0 <= x ^ 2
x:R
r:0 <= x

0 <= x * x
x:R
r:x < 0
0 <= x ^ 2
x:R
r:x < 0

0 <= x ^ 2
x:R
r:x < 0

0 <= - x * - x
apply Rmult_le_pos; lra. Qed.

forall x : R, Rabs x ^ 2 = x ^ 2

forall x : R, Rabs x ^ 2 = x ^ 2
x:R
r:0 <= x

Rabs x ^ 2 = x ^ 2
x:R
r:x < 0
Rabs x ^ 2 = x ^ 2
x:R
r:x < 0

Rabs x ^ 2 = x ^ 2
rewrite <- Rabs_Ropp, Rabs_pos_eq;[field | lra]. Qed.

Properties of tangent


forall x : R, - PI / 2 < x < PI / 2 -> derivable_pt tan x

forall x : R, - PI / 2 < x < PI / 2 -> derivable_pt tan x
x:R
xint:- PI / 2 < x < PI / 2

derivable_pt tan x
x:R
xint:- PI / 2 < x < PI / 2

{l : R | derivable_pt_abs (fun x0 : R => sin x0 / cos x0) x l}
x:R
xint:- PI / 2 < x < PI / 2

cos x <> 0
x:R
xint:- PI / 2 < x < PI / 2

cos x > 0
unfold Rgt ; apply cos_gt_0; [unfold Rdiv; rewrite <- Ropp_mult_distr_l_reverse; fold (-PI/2) |];tauto. Qed.

forall (x : R) (Pr1 : - PI / 2 < x < PI / 2), derive_pt tan x (derivable_pt_tan x Pr1) = 1 + tan x ^ 2

forall (x : R) (Pr1 : - PI / 2 < x < PI / 2), derive_pt tan x (derivable_pt_tan x Pr1) = 1 + tan x ^ 2
x:R
pr:- PI / 2 < x < PI / 2

derive_pt tan x (derivable_pt_tan x pr) = 1 + tan x ^ 2
x:R
pr:- PI / 2 < x < PI / 2

cos x <> 0
x:R
pr:- PI / 2 < x < PI / 2
H:cos x <> 0
derive_pt tan x (derivable_pt_tan x pr) = 1 + tan x ^ 2
x:R
pr:- PI / 2 < x < PI / 2
H:cos x <> 0

derive_pt tan x (derivable_pt_tan x pr) = 1 + tan x ^ 2
unfold tan; reg; unfold pow, Rsqr; field; assumption. Qed.
Proof that tangent is a bijection
(* to be removed? *)


forall (a b : R) (f : R -> R), a < b -> forall pr : forall x : R, a < x < b -> derivable_pt f x, (forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)) -> forall x y : R, a < x < b -> a < y < b -> x < y -> f x < f y

forall (a b : R) (f : R -> R), a < b -> forall pr : forall x : R, a < x < b -> derivable_pt f x, (forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)) -> forall x y : R, a < x < b -> a < y < b -> x < y -> f x < f y
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y

f x < f y
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y

forall c : R, x < c < y -> derivable_pt id c
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c : R, x < c < y -> derivable_pt id c
f x < f y
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c : R, x < c < y -> derivable_pt id c

f x < f y
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c : R, x < c < y -> derivable_pt id c

forall c : R, x < c < y -> derivable_pt f c
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c : R, x < c < y -> derivable_pt id c
derivable_f_interv:forall c : R, x < c < y -> derivable_pt f c
f x < f y
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
c:R
c_encad:x < c < y

derivable_pt f c
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c : R, x < c < y -> derivable_pt id c
derivable_f_interv:forall c : R, x < c < y -> derivable_pt f c
f x < f y
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
c:R
c_encad:x < c < y

a < c < b
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c : R, x < c < y -> derivable_pt id c
derivable_f_interv:forall c : R, x < c < y -> derivable_pt f c
f x < f y
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
c:R
c_encad:x < c < y

a < c
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
c:R
c_encad:x < c < y
c < b
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c : R, x < c < y -> derivable_pt id c
derivable_f_interv:forall c : R, x < c < y -> derivable_pt f c
f x < f y
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
c:R
c_encad:x < c < y

c < b
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c : R, x < c < y -> derivable_pt id c
derivable_f_interv:forall c : R, x < c < y -> derivable_pt f c
f x < f y
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c : R, x < c < y -> derivable_pt id c
derivable_f_interv:forall c : R, x < c < y -> derivable_pt f c

f x < f y
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c : R, x < c < y -> derivable_pt id c
derivable_f_interv:forall c : R, x < c < y -> derivable_pt f c

forall c : R, x <= c <= y -> continuity_pt f c
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c : R, x < c < y -> derivable_pt id c
derivable_f_interv:forall c : R, x < c < y -> derivable_pt f c
f_cont_interv:forall c : R, x <= c <= y -> continuity_pt f c
f x < f y
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
c:R
c_encad:x <= c <= y

a < c < b
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c : R, x < c < y -> derivable_pt id c
derivable_f_interv:forall c : R, x < c < y -> derivable_pt f c
f_cont_interv:forall c : R, x <= c <= y -> continuity_pt f c
f x < f y
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
c:R
c_encad:x <= c <= y

a < c
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
c:R
c_encad:x <= c <= y
c < b
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c : R, x < c < y -> derivable_pt id c
derivable_f_interv:forall c : R, x < c < y -> derivable_pt f c
f_cont_interv:forall c : R, x <= c <= y -> continuity_pt f c
f x < f y
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
c:R
c_encad:x <= c <= y

c < b
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c : R, x < c < y -> derivable_pt id c
derivable_f_interv:forall c : R, x < c < y -> derivable_pt f c
f_cont_interv:forall c : R, x <= c <= y -> continuity_pt f c
f x < f y
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c : R, x < c < y -> derivable_pt id c
derivable_f_interv:forall c : R, x < c < y -> derivable_pt f c
f_cont_interv:forall c : R, x <= c <= y -> continuity_pt f c

f x < f y
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c : R, x < c < y -> derivable_pt id c
derivable_f_interv:forall c : R, x < c < y -> derivable_pt f c
f_cont_interv:forall c : R, x <= c <= y -> continuity_pt f c

forall c : R, x <= c <= y -> continuity_pt id c
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c : R, x < c < y -> derivable_pt id c
derivable_f_interv:forall c : R, x < c < y -> derivable_pt f c
f_cont_interv:forall c : R, x <= c <= y -> continuity_pt f c
id_cont_interv:forall c : R, x <= c <= y -> continuity_pt id c
f x < f y
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c : R, x < c < y -> derivable_pt id c
derivable_f_interv:forall c : R, x < c < y -> derivable_pt f c
f_cont_interv:forall c : R, x <= c <= y -> continuity_pt f c
id_cont_interv:forall c : R, x <= c <= y -> continuity_pt id c

f x < f y
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c : R, x < c < y -> derivable_pt id c
derivable_f_interv:forall c : R, x < c < y -> derivable_pt f c
f_cont_interv:forall c : R, x <= c <= y -> continuity_pt f c
id_cont_interv:forall c : R, x <= c <= y -> continuity_pt id c

forall x0 : R, (exists P : x < x0 < y, (id y - id x) * derive_pt f x0 (derivable_f_interv x0 P) = (f y - f x) * derive_pt id x0 (derivable_id_interv x0 P)) -> f x < f y
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(id y - id x) * derive_pt f c (derivable_f_interv c Pr) = (f y - f x) * derive_pt id c (derivable_id_interv c Pr)

f x < f y
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = (f y - f x) * derive_pt id c (derivable_id_interv c Pr)

f x < f y
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = (f y - f x) * 1

f x < f y
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = (f y - f x) * derive_pt id c (derivable_id_interv c Pr)
1 = derive_pt id c (derivable_id_interv c Pr)
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = (f y - f x) * 1

f y - f x > 0
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = (f y - f x) * 1
Hyp:f y - f x > 0
f x < f y
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = (f y - f x) * derive_pt id c (derivable_id_interv c Pr)
1 = derive_pt id c (derivable_id_interv c Pr)
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = f y - f x

f y - f x > 0
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = (f y - f x) * 1
Hyp:f y - f x > 0
f x < f y
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = (f y - f x) * derive_pt id c (derivable_id_interv c Pr)
1 = derive_pt id c (derivable_id_interv c Pr)
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = f y - f x

(y - x) * derive_pt f c (derivable_f_interv c Pr) > 0
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = (f y - f x) * 1
Hyp:f y - f x > 0
f x < f y
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = (f y - f x) * derive_pt id c (derivable_id_interv c Pr)
1 = derive_pt id c (derivable_id_interv c Pr)
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = f y - f x

y - x > 0
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = f y - f x
derive_pt f c (derivable_f_interv c Pr) > 0
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = (f y - f x) * 1
Hyp:f y - f x > 0
f x < f y
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = (f y - f x) * derive_pt id c (derivable_id_interv c Pr)
1 = derive_pt id c (derivable_id_interv c Pr)
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = f y - f x

derive_pt f c (derivable_f_interv c Pr) > 0
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = (f y - f x) * 1
Hyp:f y - f x > 0
f x < f y
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = (f y - f x) * derive_pt id c (derivable_id_interv c Pr)
1 = derive_pt id c (derivable_id_interv c Pr)
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = f y - f x

a <= c < b
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = f y - f x
c_encad2:a <= c < b
derive_pt f c (derivable_f_interv c Pr) > 0
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = (f y - f x) * 1
Hyp:f y - f x > 0
f x < f y
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = (f y - f x) * derive_pt id c (derivable_id_interv c Pr)
1 = derive_pt id c (derivable_id_interv c Pr)
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = f y - f x

a <= c
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = f y - f x
c < b
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = f y - f x
c_encad2:a <= c < b
derive_pt f c (derivable_f_interv c Pr) > 0
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = (f y - f x) * 1
Hyp:f y - f x > 0
f x < f y
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = (f y - f x) * derive_pt id c (derivable_id_interv c Pr)
1 = derive_pt id c (derivable_id_interv c Pr)
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = f y - f x

c < b
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = f y - f x
c_encad2:a <= c < b
derive_pt f c (derivable_f_interv c Pr) > 0
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = (f y - f x) * 1
Hyp:f y - f x > 0
f x < f y
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = (f y - f x) * derive_pt id c (derivable_id_interv c Pr)
1 = derive_pt id c (derivable_id_interv c Pr)
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = f y - f x
c_encad2:a <= c < b

derive_pt f c (derivable_f_interv c Pr) > 0
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = (f y - f x) * 1
Hyp:f y - f x > 0
f x < f y
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = (f y - f x) * derive_pt id c (derivable_id_interv c Pr)
1 = derive_pt id c (derivable_id_interv c Pr)
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = f y - f x
c_encad2:a <= c < b

a < c < b
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = f y - f x
c_encad2:a <= c < b
c_encad:a < c < b
derive_pt f c (derivable_f_interv c Pr) > 0
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = (f y - f x) * 1
Hyp:f y - f x > 0
f x < f y
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = (f y - f x) * derive_pt id c (derivable_id_interv c Pr)
1 = derive_pt id c (derivable_id_interv c Pr)
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = f y - f x
c_encad2:a <= c < b

a < c
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = f y - f x
c_encad2:a <= c < b
c < b
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = f y - f x
c_encad2:a <= c < b
c_encad:a < c < b
derive_pt f c (derivable_f_interv c Pr) > 0
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = (f y - f x) * 1
Hyp:f y - f x > 0
f x < f y
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = (f y - f x) * derive_pt id c (derivable_id_interv c Pr)
1 = derive_pt id c (derivable_id_interv c Pr)
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = f y - f x
c_encad2:a <= c < b

c < b
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = f y - f x
c_encad2:a <= c < b
c_encad:a < c < b
derive_pt f c (derivable_f_interv c Pr) > 0
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = (f y - f x) * 1
Hyp:f y - f x > 0
f x < f y
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = (f y - f x) * derive_pt id c (derivable_id_interv c Pr)
1 = derive_pt id c (derivable_id_interv c Pr)
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = f y - f x
c_encad2:a <= c < b
c_encad:a < c < b

derive_pt f c (derivable_f_interv c Pr) > 0
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = (f y - f x) * 1
Hyp:f y - f x > 0
f x < f y
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = (f y - f x) * derive_pt id c (derivable_id_interv c Pr)
1 = derive_pt id c (derivable_id_interv c Pr)
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = f y - f x
c_encad2:a <= c < b
c_encad:a < c < b
Temp:0 < derive_pt f c (pr c c_encad)

derive_pt f c (derivable_f_interv c Pr) > 0
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = (f y - f x) * 1
Hyp:f y - f x > 0
f x < f y
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = (f y - f x) * derive_pt id c (derivable_id_interv c Pr)
1 = derive_pt id c (derivable_id_interv c Pr)
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = f y - f x
c_encad2:a <= c < b
c_encad:a < c < b
Temp:0 < derive_pt f c (pr c c_encad)
Temp2:derive_pt f c (derivable_f_interv c Pr) = derive_pt f c (pr c c_encad)

derive_pt f c (derivable_f_interv c Pr) > 0
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = (f y - f x) * 1
Hyp:f y - f x > 0
f x < f y
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = (f y - f x) * derive_pt id c (derivable_id_interv c Pr)
1 = derive_pt id c (derivable_id_interv c Pr)
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = (f y - f x) * 1
Hyp:f y - f x > 0

f x < f y
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = (f y - f x) * derive_pt id c (derivable_id_interv c Pr)
1 = derive_pt id c (derivable_id_interv c Pr)
a, b:R
f:R -> R
a_lt_b:a < b
pr:forall x0 : R, a < x0 < b -> derivable_pt f x0
Df_gt_0:forall (t : R) (t_encad : a < t < b), 0 < derive_pt f t (pr t t_encad)
x, y:R
x_encad:a < x < b
y_encad:a < y < b
x_lt_y:x < y
derivable_id_interv:forall c0 : R, x < c0 < y -> derivable_pt id c0
derivable_f_interv:forall c0 : R, x < c0 < y -> derivable_pt f c0
f_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt f c0
id_cont_interv:forall c0 : R, x <= c0 <= y -> continuity_pt id c0
c:R
Pr:x < c < y
eq:(y - x) * derive_pt f c (derivable_f_interv c Pr) = (f y - f x) * derive_pt id c (derivable_id_interv c Pr)

1 = derive_pt id c (derivable_id_interv c Pr)
symmetry ; rewrite derive_pt_eq ; apply derivable_pt_lim_id. Qed. (* begin hide *)

forall x : R, 1 + x ^ 2 > 0

forall x : R, 1 + x ^ 2 > 0
m:R

1 + m ^ 2 > 0
m:R

1 + m ^ 2 > 0 + 0
m:R

1 > 0
m:R
m ^ 2 >= 0
m:R

m ^ 2 >= 0
m:R
s':{m < 0} + {m = 0}

m ^ 2 >= 0
m:R
s':m > 0
m ^ 2 >= 0
m:R
s':{m < 0} + {m = 0}

m < 0 -> m ^ 2 >= 0
m:R
s':{m < 0} + {m = 0}
m = 0 -> m ^ 2 >= 0
m:R
s':m > 0
m ^ 2 >= 0
m:R
s':{m < 0} + {m = 0}
m_cond:m < 0

m ^ 2 >= 0
m:R
s':{m < 0} + {m = 0}
m = 0 -> m ^ 2 >= 0
m:R
s':m > 0
m ^ 2 >= 0
m:R
s':{m < 0} + {m = 0}
m_cond:m < 0

m ^ 2 >= 0 * 0
m:R
s':{m < 0} + {m = 0}
m = 0 -> m ^ 2 >= 0
m:R
s':m > 0
m ^ 2 >= 0
m:R
s':{m < 0} + {m = 0}
m_cond:m < 0

(- m) ^ 2 >= 0 * 0
m:R
s':{m < 0} + {m = 0}
m_cond:m < 0
(- m) ^ 2 = m ^ 2
m:R
s':{m < 0} + {m = 0}
m = 0 -> m ^ 2 >= 0
m:R
s':m > 0
m ^ 2 >= 0
m:R
s':{m < 0} + {m = 0}
m_cond:m < 0

(- m) ^ 2 = m ^ 2
m:R
s':{m < 0} + {m = 0}
m = 0 -> m ^ 2 >= 0
m:R
s':m > 0
m ^ 2 >= 0
m:R
s':{m < 0} + {m = 0}

m = 0 -> m ^ 2 >= 0
m:R
s':m > 0
m ^ 2 >= 0
m:R
s':m > 0

m ^ 2 >= 0
m:R
s':m > 0

m ^ 2 > 0
intuition. Qed. (* end hide *) (* The following lemmas about PI should probably be in Rtrigo. *)

forall x : R, 0 < x < 2 -> 0 < cos x -> x < PI / 2

forall x : R, 0 < x < 2 -> 0 < cos x -> x < PI / 2
x:R
xp:0 < x
xlt2:x < 2
cx:0 < cos x

x < PI / 2
x:R
xp:0 < x
xlt2:x < 2
cx:0 < cos x
xltpi2:x < PI / 2

x < PI / 2
x:R
xp:0 < x
xlt2:x < 2
cx:0 < cos x
xeqpi2:x = PI / 2
x < PI / 2
x:R
xp:0 < x
xlt2:x < 2
cx:0 < cos x
xgtpi2:x > PI / 2
x < PI / 2
x:R
xp:0 < x
xlt2:x < 2
cx:0 < cos x
xeqpi2:x = PI / 2

x < PI / 2
x:R
xp:0 < x
xlt2:x < 2
cx:0 < cos x
xgtpi2:x > PI / 2
x < PI / 2
x:R
xp:0 < x
xlt2:x < 2
cx:0 < cos x
xgtpi2:x > PI / 2

x < PI / 2
x:R
xp:0 < x
xlt2:x < 2
cx:0 < cos x
xgtpi2:x > PI / 2
c:R
Pc:cos x - cos (PI / 2) = derive_pt cos c (derivable_cos c) * (x - PI / 2)
cint1:PI / 2 < c
cint2:c < x

x < PI / 2
x:R
xp:0 < x
xlt2:x < 2
cx:0 < cos x
xgtpi2:x > PI / 2
c:R
cint1:PI / 2 < c
cint2:c < x

cos x = derive_pt cos c (derivable_cos c) * (x - PI / 2) -> x < PI / 2
x:R
xp:0 < x
xlt2:x < 2
cx:0 < cos x
xgtpi2:x > PI / 2
c:R
cint1:PI / 2 < c
cint2:c < x

cos x = - sin c * (x - PI / 2) -> x < PI / 2
x:R
xp:0 < x
xlt2:x < 2
cx:0 < cos x
xgtpi2:x > PI / 2
c:R
cint1:PI / 2 < c
cint2:c < x
H:0 < c < 2

cos x = - sin c * (x - PI / 2) -> x < PI / 2
x:R
xp:0 < x
xlt2:x < 2
cx:0 < cos x
xgtpi2:x > PI / 2
c:R
cint1:PI / 2 < c
cint2:c < x
H:0 < c < 2
H0:0 < sin c

cos x = - sin c * (x - PI / 2) -> x < PI / 2
x:R
xp:0 < x
xlt2:x < 2
cx:0 < cos x
xgtpi2:x > PI / 2
c:R
cint1:PI / 2 < c
cint2:c < x
H:0 < c < 2
H0:0 < sin c
Pc:cos x = - sin c * (x - PI / 2)

x < PI / 2
x:R
xp:0 < x
xlt2:x < 2
cx:0 < cos x
xgtpi2:x > PI / 2
c:R
cint1:PI / 2 < c
cint2:c < x
H:0 < c < 2
H0:0 < sin c
Pc:cos x = - sin c * (x - PI / 2)

cos x <= 0
x:R
xp:0 < x
xlt2:x < 2
cx:0 < cos x
xgtpi2:x > PI / 2
c:R
cint1:PI / 2 < c
cint2:c < x
H:0 < c < 2
H0:0 < sin c
Pc:cos x = - sin c * (x - PI / 2)

0 + - (sin c * (x - PI / 2)) <= 0
apply Rle_minus, Rmult_le_pos;[apply Rlt_le; assumption | lra ]. Qed.

3 / 2 < PI / 2

3 / 2 < PI / 2

0 < cos (3 / 2)
t:cos_approx (3 / 2) (2 * 1 + 1) <= cos (3 / 2)

0 < cos (3 / 2)

0 < cos_approx (3 / 2) (2 * 1 + 1)

0 < (-1) ^ 0 * ((3 / 2) ^ (2 * 0) / INR (fact (2 * 0))) + (-1) ^ 1 * ((3 / 2) ^ (2 * 1) / INR (fact (2 * 1))) + (-1) ^ 2 * ((3 / 2) ^ (2 * 2) / INR (fact (2 * 2))) + (-1) ^ 3 * ((3 / 2) ^ (2 * 3) / INR (fact (2 * 3)))

0 < (-1) ^ 0 * ((3 / 2) ^ (2 * 0) / IZR (Z.of_nat (fact (2 * 0)))) + (-1) ^ 1 * ((3 / 2) ^ (2 * 1) / IZR (Z.of_nat (fact (2 * 1)))) + (-1) ^ 2 * ((3 / 2) ^ (2 * 2) / IZR (Z.of_nat (fact (2 * 2)))) + (-1) ^ 3 * ((3 / 2) ^ (2 * 3) / IZR (Z.of_nat (fact (2 * 3))))

0 < 1 * (1 / 1) + -1 * 1 * (3 / 2 * (3 / 2 * 1) / 2) + -1 * (-1 * 1) * (3 / 2 * (3 / 2 * (3 / 2 * (3 / 2 * 1))) / 24) + -1 * (-1 * (-1 * 1)) * (3 / 2 * (3 / 2 * (3 / 2 * (3 / 2 * (3 / 2 * (3 / 2 * 1))))) / 720)

0 < 9925632 / 141557760
apply Rdiv_lt_0_compat ; now apply IZR_lt. Qed.

1 < PI / 2

1 < PI / 2
assert (t := PI2_3_2); lra. Qed.

forall x y : R, - PI / 2 < x -> x < y -> y < PI / 2 -> tan x < tan y

forall x y : R, - PI / 2 < x -> x < y -> y < PI / 2 -> tan x < tan y
x, y:R
Z_le_x:- PI / 2 < x
x_lt_y:x < y
y_le_1:y < PI / 2

tan x < tan y
x, y:R
Z_le_x:- PI / 2 < x
x_lt_y:x < y
y_le_1:y < PI / 2

- PI / 2 < x < PI / 2
x, y:R
Z_le_x:- PI / 2 < x
x_lt_y:x < y
y_le_1:y < PI / 2
x_encad:- PI / 2 < x < PI / 2
tan x < tan y
x, y:R
Z_le_x:- PI / 2 < x
x_lt_y:x < y
y_le_1:y < PI / 2
x_encad:- PI / 2 < x < PI / 2

tan x < tan y
x, y:R
Z_le_x:- PI / 2 < x
x_lt_y:x < y
y_le_1:y < PI / 2
x_encad:- PI / 2 < x < PI / 2

- PI / 2 < y < PI / 2
x, y:R
Z_le_x:- PI / 2 < x
x_lt_y:x < y
y_le_1:y < PI / 2
x_encad:- PI / 2 < x < PI / 2
y_encad:- PI / 2 < y < PI / 2
tan x < tan y
x, y:R
Z_le_x:- PI / 2 < x
x_lt_y:x < y
y_le_1:y < PI / 2
x_encad:- PI / 2 < x < PI / 2
y_encad:- PI / 2 < y < PI / 2

tan x < tan y
x, y:R
Z_le_x:- PI / 2 < x
x_lt_y:x < y
y_le_1:y < PI / 2
x_encad:- PI / 2 < x < PI / 2
y_encad:- PI / 2 < y < PI / 2

forall x0 : R, - PI / 2 < x0 < PI / 2 -> derivable_pt tan x0
x, y:R
Z_le_x:- PI / 2 < x
x_lt_y:x < y
y_le_1:y < PI / 2
x_encad:- PI / 2 < x < PI / 2
y_encad:- PI / 2 < y < PI / 2
local_derivable_pt_tan:forall x0 : R, - PI / 2 < x0 < PI / 2 -> derivable_pt tan x0
tan x < tan y
x, y:R
Z_le_x:- PI / 2 < x
x_lt_y:x < y
y_le_1:y < PI / 2
x_encad:- PI / 2 < x < PI / 2
y_encad:- PI / 2 < y < PI / 2
local_derivable_pt_tan:forall x0 : R, - PI / 2 < x0 < PI / 2 -> derivable_pt tan x0

tan x < tan y
x, y:R
Z_le_x:- PI / 2 < x
x_lt_y:x < y
y_le_1:y < PI / 2
local_derivable_pt_tan:forall x0 : R, - PI / 2 < x0 < PI / 2 -> derivable_pt tan x0
H:- PI / 2 < x
H0:x < PI / 2
H1:- PI / 2 < y
H2:y < PI / 2

- PI / 2 < PI / 2
x, y:R
Z_le_x:- PI / 2 < x
x_lt_y:x < y
y_le_1:y < PI / 2
local_derivable_pt_tan:forall x0 : R, - PI / 2 < x0 < PI / 2 -> derivable_pt tan x0
H:- PI / 2 < x
H0:x < PI / 2
H1:- PI / 2 < y
H2:y < PI / 2
t:R
t_encad:- PI / 2 < t < PI / 2
0 < derive_pt tan t (local_derivable_pt_tan t t_encad)
x, y:R
Z_le_x:- PI / 2 < x
x_lt_y:x < y
y_le_1:y < PI / 2
local_derivable_pt_tan:forall x0 : R, - PI / 2 < x0 < PI / 2 -> derivable_pt tan x0
H:- PI / 2 < x
H0:x < PI / 2
H1:- PI / 2 < y
H2:y < PI / 2
t:R
t_encad:- PI / 2 < t < PI / 2

0 < derive_pt tan t (local_derivable_pt_tan t t_encad)
x, y:R
Z_le_x:- PI / 2 < x
x_lt_y:x < y
y_le_1:y < PI / 2
local_derivable_pt_tan:forall x0 : R, - PI / 2 < x0 < PI / 2 -> derivable_pt tan x0
H:- PI / 2 < x
H0:x < PI / 2
H1:- PI / 2 < y
H2:y < PI / 2
t:R
t_encad:- PI / 2 < t < PI / 2

0 < derive_pt tan t (derivable_pt_tan t t_encad)
x, y:R
Z_le_x:- PI / 2 < x
x_lt_y:x < y
y_le_1:y < PI / 2
local_derivable_pt_tan:forall x0 : R, - PI / 2 < x0 < PI / 2 -> derivable_pt tan x0
H:- PI / 2 < x
H0:x < PI / 2
H1:- PI / 2 < y
H2:y < PI / 2
t:R
t_encad:- PI / 2 < t < PI / 2

0 < 1 + tan t ^ 2
apply plus_Rsqr_gt_0. Qed.

forall x y : R, - PI / 2 < x < PI / 2 -> - PI / 2 < y < PI / 2 -> tan x = tan y -> x = y

forall x y : R, - PI / 2 < x < PI / 2 -> - PI / 2 < y < PI / 2 -> tan x = tan y -> x = y
a, b:R
a_encad:- PI / 2 < a < PI / 2
b_encad:- PI / 2 < b < PI / 2
fa_eq_fb:tan a = tan b

a = b
a, b:R
a_encad:- PI / 2 < a < PI / 2
b_encad:- PI / 2 < b < PI / 2
fa_eq_fb:tan a = tan b

{a < b} + {a = b} -> a = b
a, b:R
a_encad:- PI / 2 < a < PI / 2
b_encad:- PI / 2 < b < PI / 2
fa_eq_fb:tan a = tan b
a > b -> a = b
a, b:R
a_encad:- PI / 2 < a < PI / 2
b_encad:- PI / 2 < b < PI / 2
fa_eq_fb:tan a = tan b

a < b -> a = b
a, b:R
a_encad:- PI / 2 < a < PI / 2
b_encad:- PI / 2 < b < PI / 2
fa_eq_fb:tan a = tan b
a = b -> a = b
a, b:R
a_encad:- PI / 2 < a < PI / 2
b_encad:- PI / 2 < b < PI / 2
fa_eq_fb:tan a = tan b
a > b -> a = b
a, b:R
a_encad:- PI / 2 < a < PI / 2
b_encad:- PI / 2 < b < PI / 2
fa_eq_fb:tan a = tan b
Hf:a < b

a = b
a, b:R
a_encad:- PI / 2 < a < PI / 2
b_encad:- PI / 2 < b < PI / 2
fa_eq_fb:tan a = tan b
a = b -> a = b
a, b:R
a_encad:- PI / 2 < a < PI / 2
b_encad:- PI / 2 < b < PI / 2
fa_eq_fb:tan a = tan b
a > b -> a = b
a, b:R
a_encad:- PI / 2 < a < PI / 2
b_encad:- PI / 2 < b < PI / 2
fa_eq_fb:tan a = tan b
Hf:a < b
Hfalse:tan a < tan b

a = b
a, b:R
a_encad:- PI / 2 < a < PI / 2
b_encad:- PI / 2 < b < PI / 2
fa_eq_fb:tan a = tan b
a = b -> a = b
a, b:R
a_encad:- PI / 2 < a < PI / 2
b_encad:- PI / 2 < b < PI / 2
fa_eq_fb:tan a = tan b
a > b -> a = b
a, b:R
a_encad:- PI / 2 < a < PI / 2
b_encad:- PI / 2 < b < PI / 2
fa_eq_fb:tan a = tan b

a = b -> a = b
a, b:R
a_encad:- PI / 2 < a < PI / 2
b_encad:- PI / 2 < b < PI / 2
fa_eq_fb:tan a = tan b
a > b -> a = b
a, b:R
a_encad:- PI / 2 < a < PI / 2
b_encad:- PI / 2 < b < PI / 2
fa_eq_fb:tan a = tan b

a > b -> a = b
a, b:R
a_encad:- PI / 2 < a < PI / 2
b_encad:- PI / 2 < b < PI / 2
fa_eq_fb:tan a = tan b
Hf:a > b

a = b
a, b:R
a_encad:- PI / 2 < a < PI / 2
b_encad:- PI / 2 < b < PI / 2
fa_eq_fb:tan a = tan b
Hf:a > b
Hfalse:tan b < tan a

a = b
case (Rlt_not_eq (tan b) (tan a)) ; [|symmetry] ; assumption. Qed.

forall lb ub y : R, lb < ub -> - PI / 2 < lb -> ub < PI / 2 -> tan lb < y < tan ub -> {x : R | lb < x < ub /\ tan x = y}

forall lb ub y : R, lb < ub -> - PI / 2 < lb -> ub < PI / 2 -> tan lb < y < tan ub -> {x : R | lb < x < ub /\ tan x = y}
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad:tan lb < y < tan ub

{x : R | lb < x < ub /\ tan x = y}
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad:tan lb < y < tan ub
y_encad1:tan lb < y
y_encad2:y < tan ub

{x : R | lb < x < ub /\ tan x = y}
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad:tan lb < y < tan ub
y_encad1:tan lb < y
y_encad2:y < tan ub

forall a : R, lb <= a <= ub -> continuity_pt tan a
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad:tan lb < y < tan ub
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
{x : R | lb < x < ub /\ tan x = y}
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad:tan lb < y < tan ub
y_encad1:tan lb < y
y_encad2:y < tan ub
a:R
a_encad:lb <= a <= ub

continuity_pt tan a
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad:tan lb < y < tan ub
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
{x : R | lb < x < ub /\ tan x = y}
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad:tan lb < y < tan ub
y_encad1:tan lb < y
y_encad2:y < tan ub
a:R
a_encad:lb <= a <= ub

- PI / 2 < a < PI / 2
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad:tan lb < y < tan ub
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
{x : R | lb < x < ub /\ tan x = y}
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad:tan lb < y < tan ub
y_encad1:tan lb < y
y_encad2:y < tan ub
a:R
a_encad:lb <= a <= ub

- PI / 2 < a
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad:tan lb < y < tan ub
y_encad1:tan lb < y
y_encad2:y < tan ub
a:R
a_encad:lb <= a <= ub
a < PI / 2
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad:tan lb < y < tan ub
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
{x : R | lb < x < ub /\ tan x = y}
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad:tan lb < y < tan ub
y_encad1:tan lb < y
y_encad2:y < tan ub
a:R
a_encad:lb <= a <= ub

a < PI / 2
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad:tan lb < y < tan ub
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
{x : R | lb < x < ub /\ tan x = y}
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad:tan lb < y < tan ub
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a

{x : R | lb < x < ub /\ tan x = y}
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad:tan lb < y < tan ub
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a

forall a : R, lb <= a <= ub -> continuity_pt (fun x : R => tan x - y) a
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad:tan lb < y < tan ub
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x : R => tan x - y) a
{x : R | lb < x < ub /\ tan x = y}
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad:tan lb < y < tan ub
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a0 : R, lb <= a0 <= ub -> continuity_pt tan a0
a:R
a_encad:lb <= a <= ub

continuity_pt (fun x : R => tan x - y) a
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad:tan lb < y < tan ub
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x : R => tan x - y) a
{x : R | lb < x < ub /\ tan x = y}
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad:tan lb < y < tan ub
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a0 : R, lb <= a0 <= ub -> continuity_pt tan a0
a:R
a_encad:lb <= a <= ub

forall eps : R, eps > 0 -> exists alp : R, alp > 0 /\ (forall x : R, D_x no_cond a x /\ Rabs (x - a) < alp -> Rabs (tan x - y - (tan a - y)) < eps)
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad:tan lb < y < tan ub
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x : R => tan x - y) a
{x : R | lb < x < ub /\ tan x = y}
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad:tan lb < y < tan ub
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a0 : R, lb <= a0 <= ub -> continuity_pt tan a0
a:R
a_encad:lb <= a <= ub
eps:R
eps_pos:eps > 0

exists alp : R, alp > 0 /\ (forall x : R, D_x no_cond a x /\ Rabs (x - a) < alp -> Rabs (tan x - y - (tan a - y)) < eps)
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad:tan lb < y < tan ub
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x : R => tan x - y) a
{x : R | lb < x < ub /\ tan x = y}
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad:tan lb < y < tan ub
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a0 : R, lb <= a0 <= ub -> continuity_pt tan a0
a:R
a_encad:lb <= a <= ub
eps:R
eps_pos:eps > 0

forall x : R, x > 0 /\ (forall x0 : Base R_met, D_x no_cond a x0 /\ dist R_met x0 a < x -> dist R_met (tan x0) (tan a) < eps) -> exists alp : R, alp > 0 /\ (forall x0 : R, D_x no_cond a x0 /\ Rabs (x0 - a) < alp -> Rabs (tan x0 - y - (tan a - y)) < eps)
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad:tan lb < y < tan ub
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x : R => tan x - y) a
{x : R | lb < x < ub /\ tan x = y}
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad:tan lb < y < tan ub
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a0 : R, lb <= a0 <= ub -> continuity_pt tan a0
a:R
a_encad:lb <= a <= ub
eps:R
eps_pos:eps > 0
alpha:R
alpha_pos:alpha > 0 /\ (forall x : Base R_met, D_x no_cond a x /\ dist R_met x a < alpha -> dist R_met (tan x) (tan a) < eps)

exists alp : R, alp > 0 /\ (forall x : R, D_x no_cond a x /\ Rabs (x - a) < alp -> Rabs (tan x - y - (tan a - y)) < eps)
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad:tan lb < y < tan ub
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x : R => tan x - y) a
{x : R | lb < x < ub /\ tan x = y}
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad:tan lb < y < tan ub
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a0 : R, lb <= a0 <= ub -> continuity_pt tan a0
a:R
a_encad:lb <= a <= ub
eps:R
eps_pos:eps > 0
alpha:R
alpha_pos:alpha > 0
Temp:forall x : Base R_met, D_x no_cond a x /\ dist R_met x a < alpha -> dist R_met (tan x) (tan a) < eps

exists alp : R, alp > 0 /\ (forall x : R, D_x no_cond a x /\ Rabs (x - a) < alp -> Rabs (tan x - y - (tan a - y)) < eps)
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad:tan lb < y < tan ub
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x : R => tan x - y) a
{x : R | lb < x < ub /\ tan x = y}
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad:tan lb < y < tan ub
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a0 : R, lb <= a0 <= ub -> continuity_pt tan a0
a:R
a_encad:lb <= a <= ub
eps:R
eps_pos:eps > 0
alpha:R
alpha_pos:alpha > 0
Temp:forall x : Base R_met, D_x no_cond a x /\ dist R_met x a < alpha -> dist R_met (tan x) (tan a) < eps

alpha > 0 /\ (forall x : R, D_x no_cond a x /\ Rabs (x - a) < alpha -> Rabs (tan x - y - (tan a - y)) < eps)
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad:tan lb < y < tan ub
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x : R => tan x - y) a
{x : R | lb < x < ub /\ tan x = y}
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad:tan lb < y < tan ub
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a0 : R, lb <= a0 <= ub -> continuity_pt tan a0
a:R
a_encad:lb <= a <= ub
eps:R
eps_pos:eps > 0
alpha:R
alpha_pos:alpha > 0
Temp:forall x : Base R_met, D_x no_cond a x /\ dist R_met x a < alpha -> dist R_met (tan x) (tan a) < eps

alpha > 0
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad:tan lb < y < tan ub
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a0 : R, lb <= a0 <= ub -> continuity_pt tan a0
a:R
a_encad:lb <= a <= ub
eps:R
eps_pos:eps > 0
alpha:R
alpha_pos:alpha > 0
Temp:forall x : Base R_met, D_x no_cond a x /\ dist R_met x a < alpha -> dist R_met (tan x) (tan a) < eps
forall x : R, D_x no_cond a x /\ Rabs (x - a) < alpha -> Rabs (tan x - y - (tan a - y)) < eps
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad:tan lb < y < tan ub
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x : R => tan x - y) a
{x : R | lb < x < ub /\ tan x = y}
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad:tan lb < y < tan ub
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a0 : R, lb <= a0 <= ub -> continuity_pt tan a0
a:R
a_encad:lb <= a <= ub
eps:R
eps_pos:eps > 0
alpha:R
alpha_pos:alpha > 0
Temp:forall x : Base R_met, D_x no_cond a x /\ dist R_met x a < alpha -> dist R_met (tan x) (tan a) < eps

forall x : R, D_x no_cond a x /\ Rabs (x - a) < alpha -> Rabs (tan x - y - (tan a - y)) < eps
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad:tan lb < y < tan ub
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x : R => tan x - y) a
{x : R | lb < x < ub /\ tan x = y}
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad:tan lb < y < tan ub
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a0 : R, lb <= a0 <= ub -> continuity_pt tan a0
a:R
a_encad:lb <= a <= ub
eps:R
eps_pos:eps > 0
alpha:R
alpha_pos:alpha > 0
Temp:forall x0 : Base R_met, D_x no_cond a x0 /\ dist R_met x0 a < alpha -> dist R_met (tan x0) (tan a) < eps
x:R
x_cond:D_x no_cond a x /\ Rabs (x - a) < alpha

Rabs (tan x - y - (tan a - y)) < eps
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad:tan lb < y < tan ub
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x : R => tan x - y) a
{x : R | lb < x < ub /\ tan x = y}
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad:tan lb < y < tan ub
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a0 : R, lb <= a0 <= ub -> continuity_pt tan a0
a:R
a_encad:lb <= a <= ub
eps:R
eps_pos:eps > 0
alpha:R
alpha_pos:alpha > 0
Temp:forall x0 : Base R_met, D_x no_cond a x0 /\ dist R_met x0 a < alpha -> dist R_met (tan x0) (tan a) < eps
x:R
x_cond:D_x no_cond a x /\ Rabs (x - a) < alpha

Rabs (tan x - tan a) < eps
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad:tan lb < y < tan ub
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x : R => tan x - y) a
{x : R | lb < x < ub /\ tan x = y}
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad:tan lb < y < tan ub
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x : R => tan x - y) a

{x : R | lb < x < ub /\ tan x = y}
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad:tan lb < y < tan ub
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x : R => tan x - y) a

tan lb - y < 0
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad:tan lb < y < tan ub
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x : R => tan x - y) a
H1:(fun x : R => tan x - y) lb < 0
{x : R | lb < x < ub /\ tan x = y}
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad:tan lb < y < tan ub
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x : R => tan x - y) a

tan lb < y
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad:tan lb < y < tan ub
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x : R => tan x - y) a
H1:(fun x : R => tan x - y) lb < 0
{x : R | lb < x < ub /\ tan x = y}
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad:tan lb < y < tan ub
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x : R => tan x - y) a
H1:(fun x : R => tan x - y) lb < 0

{x : R | lb < x < ub /\ tan x = y}
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad:tan lb < y < tan ub
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x : R => tan x - y) a
H1:(fun x : R => tan x - y) lb < 0

0 < tan ub - y
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad:tan lb < y < tan ub
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x : R => tan x - y) a
H1:(fun x : R => tan x - y) lb < 0
H2:0 < (fun x : R => tan x - y) ub
{x : R | lb < x < ub /\ tan x = y}
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad:tan lb < y < tan ub
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x : R => tan x - y) a
H1:(fun x : R => tan x - y) lb < 0

tan ub > y
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad:tan lb < y < tan ub
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x : R => tan x - y) a
H1:(fun x : R => tan x - y) lb < 0
H2:0 < (fun x : R => tan x - y) ub
{x : R | lb < x < ub /\ tan x = y}
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad:tan lb < y < tan ub
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x : R => tan x - y) a
H1:(fun x : R => tan x - y) lb < 0
H2:0 < (fun x : R => tan x - y) ub

{x : R | lb < x < ub /\ tan x = y}
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad:tan lb < y < tan ub
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x0 : R => tan x0 - y) a
H1:(fun x0 : R => tan x0 - y) lb < 0
H2:0 < (fun x0 : R => tan x0 - y) ub
x:R
Hx:lb <= x <= ub /\ tan x - y = 0

{x0 : R | lb < x0 < ub /\ tan x0 = y}
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad:tan lb < y < tan ub
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x0 : R => tan x0 - y) a
H1:(fun x0 : R => tan x0 - y) lb < 0
H2:0 < (fun x0 : R => tan x0 - y) ub
x:R
Hx:lb <= x <= ub /\ tan x - y = 0

lb < x < ub /\ tan x = y
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad:tan lb < y < tan ub
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x0 : R => tan x0 - y) a
H1:(fun x0 : R => tan x0 - y) lb < 0
H2:0 < (fun x0 : R => tan x0 - y) ub
x:R
Hyp:lb <= x <= ub
Result:tan x - y = 0

lb < x < ub /\ tan x = y
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x0 : R => tan x0 - y) a
H1:tan lb - y < 0
H2:0 < tan ub - y
x:R
Result:tan x - y = 0
H:tan lb < y
H0:y < tan ub
H3:lb <= x
H4:x <= ub

lb < x
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x0 : R => tan x0 - y) a
H1:tan lb - y < 0
H2:0 < tan ub - y
x:R
Result:tan x - y = 0
H:tan lb < y
H0:y < tan ub
H3:lb <= x
H4:x <= ub
x < ub
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x0 : R => tan x0 - y) a
H1:tan lb - y < 0
H2:0 < tan ub - y
x:R
Result:tan x - y = 0
H:tan lb < y
H0:y < tan ub
H3:lb <= x
H4:x <= ub

x <> lb
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x0 : R => tan x0 - y) a
H1:tan lb - y < 0
H2:0 < tan ub - y
x:R
Result:tan x - y = 0
H:tan lb < y
H0:y < tan ub
H3:lb <= x
H4:x <= ub
Temp2:x <> lb
lb < x
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x0 : R => tan x0 - y) a
H1:tan lb - y < 0
H2:0 < tan ub - y
x:R
Result:tan x - y = 0
H:tan lb < y
H0:y < tan ub
H3:lb <= x
H4:x <= ub
x < ub
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x0 : R => tan x0 - y) a
H1:tan lb - y < 0
H2:0 < tan ub - y
x:R
Result:tan x - y = 0
H:tan lb < y
H0:y < tan ub
H3:lb <= x
H4:x <= ub
Hfalse:x = lb

False
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x0 : R => tan x0 - y) a
H1:tan lb - y < 0
H2:0 < tan ub - y
x:R
Result:tan x - y = 0
H:tan lb < y
H0:y < tan ub
H3:lb <= x
H4:x <= ub
Temp2:x <> lb
lb < x
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x0 : R => tan x0 - y) a
H1:tan lb - y < 0
H2:0 < tan ub - y
x:R
Result:tan x - y = 0
H:tan lb < y
H0:y < tan ub
H3:lb <= x
H4:x <= ub
x < ub
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x0 : R => tan x0 - y) a
H1:tan lb - y < 0
H2:0 < tan ub - y
x:R
Result:tan lb - y = 0
H:tan lb < y
H0:y < tan ub
H3:lb <= x
H4:x <= ub
Hfalse:x = lb

False
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x0 : R => tan x0 - y) a
H1:tan lb - y < 0
H2:0 < tan ub - y
x:R
Result:tan x - y = 0
H:tan lb < y
H0:y < tan ub
H3:lb <= x
H4:x <= ub
Temp2:x <> lb
lb < x
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x0 : R => tan x0 - y) a
H1:tan lb - y < 0
H2:0 < tan ub - y
x:R
Result:tan x - y = 0
H:tan lb < y
H0:y < tan ub
H3:lb <= x
H4:x <= ub
x < ub
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x0 : R => tan x0 - y) a
H1:tan lb - y < 0
H2:0 < tan ub - y
x:R
Result:tan lb - y = 0
H:tan lb < y
H0:y < tan ub
H3:lb <= x
H4:x <= ub
Hfalse:x = lb

y <> tan lb
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x0 : R => tan x0 - y) a
H1:tan lb - y < 0
H2:0 < tan ub - y
x:R
Result:tan lb - y = 0
H:tan lb < y
H0:y < tan ub
H3:lb <= x
H4:x <= ub
Hfalse:x = lb
Temp2:y <> tan lb
False
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x0 : R => tan x0 - y) a
H1:tan lb - y < 0
H2:0 < tan ub - y
x:R
Result:tan x - y = 0
H:tan lb < y
H0:y < tan ub
H3:lb <= x
H4:x <= ub
Temp2:x <> lb
lb < x
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x0 : R => tan x0 - y) a
H1:tan lb - y < 0
H2:0 < tan ub - y
x:R
Result:tan x - y = 0
H:tan lb < y
H0:y < tan ub
H3:lb <= x
H4:x <= ub
x < ub
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x0 : R => tan x0 - y) a
H1:tan lb - y < 0
H2:0 < tan ub - y
x:R
Result:tan lb - y = 0
H:tan lb < y
H0:y < tan ub
H3:lb <= x
H4:x <= ub
Hfalse:x = lb
Temp2:y <> tan lb

False
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x0 : R => tan x0 - y) a
H1:tan lb - y < 0
H2:0 < tan ub - y
x:R
Result:tan x - y = 0
H:tan lb < y
H0:y < tan ub
H3:lb <= x
H4:x <= ub
Temp2:x <> lb
lb < x
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x0 : R => tan x0 - y) a
H1:tan lb - y < 0
H2:0 < tan ub - y
x:R
Result:tan x - y = 0
H:tan lb < y
H0:y < tan ub
H3:lb <= x
H4:x <= ub
x < ub
lb, y:R
Result:tan lb - y = 0
Temp2:y <> tan lb

False
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x0 : R => tan x0 - y) a
H1:tan lb - y < 0
H2:0 < tan ub - y
x:R
Result:tan x - y = 0
H:tan lb < y
H0:y < tan ub
H3:lb <= x
H4:x <= ub
Temp2:x <> lb
lb < x
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x0 : R => tan x0 - y) a
H1:tan lb - y < 0
H2:0 < tan ub - y
x:R
Result:tan x - y = 0
H:tan lb < y
H0:y < tan ub
H3:lb <= x
H4:x <= ub
x < ub
lb, y:R
Result:tan lb - y = 0
Temp2:y <> tan lb

y = tan lb
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x0 : R => tan x0 - y) a
H1:tan lb - y < 0
H2:0 < tan ub - y
x:R
Result:tan x - y = 0
H:tan lb < y
H0:y < tan ub
H3:lb <= x
H4:x <= ub
Temp2:x <> lb
lb < x
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x0 : R => tan x0 - y) a
H1:tan lb - y < 0
H2:0 < tan ub - y
x:R
Result:tan x - y = 0
H:tan lb < y
H0:y < tan ub
H3:lb <= x
H4:x <= ub
x < ub
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x0 : R => tan x0 - y) a
H1:tan lb - y < 0
H2:0 < tan ub - y
x:R
Result:tan x - y = 0
H:tan lb < y
H0:y < tan ub
H3:lb <= x
H4:x <= ub
Temp2:x <> lb

lb < x
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x0 : R => tan x0 - y) a
H1:tan lb - y < 0
H2:0 < tan ub - y
x:R
Result:tan x - y = 0
H:tan lb < y
H0:y < tan ub
H3:lb <= x
H4:x <= ub
x < ub
lb, x:R
H3:lb <= x
Temp2:x <> lb

lb < x
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x0 : R => tan x0 - y) a
H1:tan lb - y < 0
H2:0 < tan ub - y
x:R
Result:tan x - y = 0
H:tan lb < y
H0:y < tan ub
H3:lb <= x
H4:x <= ub
x < ub
lb, x:R
H3:lb <= x
Temp2:x = lb -> False
H:lb = x

lb < x
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x0 : R => tan x0 - y) a
H1:tan lb - y < 0
H2:0 < tan ub - y
x:R
Result:tan x - y = 0
H:tan lb < y
H0:y < tan ub
H3:lb <= x
H4:x <= ub
x < ub
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x0 : R => tan x0 - y) a
H1:tan lb - y < 0
H2:0 < tan ub - y
x:R
Result:tan x - y = 0
H:tan lb < y
H0:y < tan ub
H3:lb <= x
H4:x <= ub

x < ub
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x0 : R => tan x0 - y) a
H1:tan lb - y < 0
H2:0 < tan ub - y
x:R
Result:tan x - y = 0
H:tan lb < y
H0:y < tan ub
H3:lb <= x
H4:x <= ub

x <> ub
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x0 : R => tan x0 - y) a
H1:tan lb - y < 0
H2:0 < tan ub - y
x:R
Result:tan x - y = 0
H:tan lb < y
H0:y < tan ub
H3:lb <= x
H4:x <= ub
Temp:x <> ub
x < ub
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x0 : R => tan x0 - y) a
H1:tan lb - y < 0
H2:0 < tan ub - y
x:R
Result:tan x - y = 0
H:tan lb < y
H0:y < tan ub
H3:lb <= x
H4:x <= ub
Hfalse:x = ub

False
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x0 : R => tan x0 - y) a
H1:tan lb - y < 0
H2:0 < tan ub - y
x:R
Result:tan x - y = 0
H:tan lb < y
H0:y < tan ub
H3:lb <= x
H4:x <= ub
Temp:x <> ub
x < ub
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x0 : R => tan x0 - y) a
H1:tan lb - y < 0
H2:0 < tan ub - y
x:R
Result:tan ub - y = 0
H:tan lb < y
H0:y < tan ub
H3:lb <= x
H4:x <= ub
Hfalse:x = ub

False
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x0 : R => tan x0 - y) a
H1:tan lb - y < 0
H2:0 < tan ub - y
x:R
Result:tan x - y = 0
H:tan lb < y
H0:y < tan ub
H3:lb <= x
H4:x <= ub
Temp:x <> ub
x < ub
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x0 : R => tan x0 - y) a
H1:tan lb - y < 0
H2:0 < tan ub - y
x:R
Result:tan ub - y = 0
H:tan lb < y
H0:y < tan ub
H3:lb <= x
H4:x <= ub
Hfalse:x = ub

y <> tan ub
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x0 : R => tan x0 - y) a
H1:tan lb - y < 0
H2:0 < tan ub - y
x:R
Result:tan ub - y = 0
H:tan lb < y
H0:y < tan ub
H3:lb <= x
H4:x <= ub
Hfalse:x = ub
Temp2:y <> tan ub
False
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x0 : R => tan x0 - y) a
H1:tan lb - y < 0
H2:0 < tan ub - y
x:R
Result:tan x - y = 0
H:tan lb < y
H0:y < tan ub
H3:lb <= x
H4:x <= ub
Temp:x <> ub
x < ub
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x0 : R => tan x0 - y) a
H1:tan lb - y < 0
H2:0 < tan ub - y
x:R
Result:tan ub - y = 0
H:tan lb < y
H0:y < tan ub
H3:lb <= x
H4:x <= ub
Hfalse:x = ub
Temp2:y <> tan ub

False
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x0 : R => tan x0 - y) a
H1:tan lb - y < 0
H2:0 < tan ub - y
x:R
Result:tan x - y = 0
H:tan lb < y
H0:y < tan ub
H3:lb <= x
H4:x <= ub
Temp:x <> ub
x < ub
ub, y:R
Result:tan ub - y = 0
Temp2:y <> tan ub

False
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x0 : R => tan x0 - y) a
H1:tan lb - y < 0
H2:0 < tan ub - y
x:R
Result:tan x - y = 0
H:tan lb < y
H0:y < tan ub
H3:lb <= x
H4:x <= ub
Temp:x <> ub
x < ub
ub, y:R
Result:tan ub - y = 0
Temp2:y <> tan ub

y = tan ub
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x0 : R => tan x0 - y) a
H1:tan lb - y < 0
H2:0 < tan ub - y
x:R
Result:tan x - y = 0
H:tan lb < y
H0:y < tan ub
H3:lb <= x
H4:x <= ub
Temp:x <> ub
x < ub
lb, ub, y:R
lb_lt_ub:lb < ub
lb_cond:- PI / 2 < lb
ub_cond:ub < PI / 2
y_encad1:tan lb < y
y_encad2:y < tan ub
f_cont:forall a : R, lb <= a <= ub -> continuity_pt tan a
Cont:forall a : R, lb <= a <= ub -> continuity_pt (fun x0 : R => tan x0 - y) a
H1:tan lb - y < 0
H2:0 < tan ub - y
x:R
Result:tan x - y = 0
H:tan lb < y
H0:y < tan ub
H3:lb <= x
H4:x <= ub
Temp:x <> ub

x < ub
case H4 ; intuition. Qed.

Definition of arctangent as the reciprocal function of tangent and proof of this status


tan 1 > 1

tan 1 > 1
H:0 < cos 1

tan 1 > 1
H:0 < cos 1

cos 1 <= 1 - 1 / 2 + 1 / 24
H:0 < cos 1
t1:cos 1 <= 1 - 1 / 2 + 1 / 24
tan 1 > 1
H:0 < cos 1

cos 1 <= cos_approx 1 (2 * (0 + 1)) -> cos 1 <= 1 - 1 / 2 + 1 / 24
H:0 < cos 1
t1:cos 1 <= 1 - 1 / 2 + 1 / 24
tan 1 > 1
H:0 < cos 1
t:cos 1 <= 1 * (1 / 1) + -1 * 1 * (1 * (1 * 1) / (1 + 1)) + -1 * (-1 * 1) * (1 * (1 * (1 * (1 * 1))) / (... + 1 + 1 + 1 + 1 + 1 + 1 + 1 + 1 + 1 + 1 + 1 + 1 + 1 + 1 + 1 + 1 + 1 + 1 + 1 + 1 + 1))

1 * (1 / 1) + -1 * 1 * (1 * (1 * 1) / (1 + 1)) + -1 * (-1 * 1) * (1 * (1 * (1 * (1 * 1))) / (... + 1 + 1 + 1 + 1 + 1 + 1 + 1 + 1 + 1 + 1 + 1 + 1 + 1 + 1 + 1 + 1 + 1 + 1 + 1 + 1 + 1)) <= 1 - 1 / 2 + 1 / 24
H:0 < cos 1
t1:cos 1 <= 1 - 1 / 2 + 1 / 24
tan 1 > 1
H:0 < cos 1
t1:cos 1 <= 1 - 1 / 2 + 1 / 24

tan 1 > 1
H:0 < cos 1
t1:cos 1 <= 1 - 1 / 2 + 1 / 24

1 - 1 / 6 <= sin 1
H:0 < cos 1
t1:cos 1 <= 1 - 1 / 2 + 1 / 24
t2:1 - 1 / 6 <= sin 1
tan 1 > 1
H:0 < cos 1
t1:cos 1 <= 1 - 1 / 2 + 1 / 24

sin_approx 1 (2 * 0 + 1) <= sin 1 -> 1 - 1 / 6 <= sin 1
H:0 < cos 1
t1:cos 1 <= 1 - 1 / 2 + 1 / 24
t2:1 - 1 / 6 <= sin 1
tan 1 > 1
H:0 < cos 1
t1:cos 1 <= 1 - 1 / 2 + 1 / 24
t:1 * (1 * 1 / 1) + -1 * 1 * (1 * (1 * (1 * 1)) / (1 + 1 + 1 + 1 + 1 + 1)) <= sin 1

1 - 1 / 6 <= 1 * (1 * 1 / 1) + -1 * 1 * (1 * (1 * (1 * 1)) / (1 + 1 + 1 + 1 + 1 + 1))
H:0 < cos 1
t1:cos 1 <= 1 - 1 / 2 + 1 / 24
t2:1 - 1 / 6 <= sin 1
tan 1 > 1
H:0 < cos 1
t1:cos 1 <= 1 - 1 / 2 + 1 / 24
t2:1 - 1 / 6 <= sin 1

tan 1 > 1
H:0 < cos 1
t1:cos 1 <= 1 - 1 / 2 + 1 / 24
t2:1 - 1 / 6 <= sin 1

tan 1 > cos 1 / cos 1
H:0 < cos 1
t1:cos 1 <= 1 - 1 / 2 + 1 / 24
t2:1 - 1 / 6 <= sin 1

0 < / cos 1
H:0 < cos 1
t1:cos 1 <= 1 - 1 / 2 + 1 / 24
t2:1 - 1 / 6 <= sin 1
cos 1 < sin 1
H:0 < cos 1
t1:cos 1 <= 1 - 1 / 2 + 1 / 24
t2:1 - 1 / 6 <= sin 1

cos 1 < sin 1
H:0 < cos 1
t1:cos 1 <= 1 - 1 / 2 + 1 / 24
t2:1 - 1 / 6 <= sin 1

1 - 1 / 2 + 1 / 24 < 1 - 1 / 6
lra. Qed.
y:R

{x : R | 0 < x < PI / 2 /\ Rabs y < tan x}
y:R

{x : R | 0 < x < PI / 2 /\ Rabs y < tan x}
y:R
Hs:{Rabs y < 1} + {Rabs y = 1}

{x : R | 0 < x < PI / 2 /\ Rabs y < tan x}
y:R
Hgt:Rabs y > 1
{x : R | 0 < x < PI / 2 /\ Rabs y < tan x}
y:R
Hs:{Rabs y < 1} + {Rabs y = 1}
yle1:Rabs y <= 1

{x : R | 0 < x < PI / 2 /\ Rabs y < tan x}
y:R
Hgt:Rabs y > 1
{x : R | 0 < x < PI / 2 /\ Rabs y < tan x}
y:R
yle1:Rabs y <= 1

Rabs y < tan 1
y:R
Hgt:Rabs y > 1
{x : R | 0 < x < PI / 2 /\ Rabs y < tan x}
y:R
Hgt:Rabs y > 1

{x : R | 0 < x < PI / 2 /\ Rabs y < tan x}
y:R
Hgt:Rabs y > 1

0 < / (Rabs y + 1)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
{x : R | 0 < x < PI / 2 /\ Rabs y < tan x}
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)

{x : R | 0 < x < PI / 2 /\ Rabs y < tan x}
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R

{x : R | 0 < x < PI / 2 /\ Rabs y < tan x}
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R

0 < u
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
{x : R | 0 < x < PI / 2 /\ Rabs y < tan x}
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u

{x : R | 0 < x < PI / 2 /\ Rabs y < tan x}
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u

/ (Rabs y + 1) < 1
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
{x : R | 0 < x < PI / 2 /\ Rabs y < tan x}
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u

0 < Rabs y + 1
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
/ (Rabs y + 1) * (Rabs y + 1) < 1 * (Rabs y + 1)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
{x : R | 0 < x < PI / 2 /\ Rabs y < tan x}
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u

/ (Rabs y + 1) * (Rabs y + 1) < 1 * (Rabs y + 1)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
{x : R | 0 < x < PI / 2 /\ Rabs y < tan x}
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1

{x : R | 0 < x < PI / 2 /\ Rabs y < tan x}
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1

u < 1
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
{x : R | 0 < x < PI / 2 /\ Rabs y < tan x}
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1

u < / (Rabs y + 1)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
/ (Rabs y + 1) < 1
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
{x : R | 0 < x < PI / 2 /\ Rabs y < tan x}
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1

u < / (Rabs y + 1) / 2 + / (Rabs y + 1) / 2
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
/ (Rabs y + 1) < 1
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
{x : R | 0 < x < PI / 2 /\ Rabs y < tan x}
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
t:forall x : R, 0 < x -> x < x + x

u < / (Rabs y + 1) / 2 + / (Rabs y + 1) / 2
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
/ (Rabs y + 1) < 1
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
{x : R | 0 < x < PI / 2 /\ Rabs y < tan x}
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
t:forall x : R, 0 < x -> x < x + x

0 < / (Rabs y + 1) / 2
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
/ (Rabs y + 1) < 1
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
{x : R | 0 < x < PI / 2 /\ Rabs y < tan x}
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1

/ (Rabs y + 1) < 1
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
{x : R | 0 < x < PI / 2 /\ Rabs y < tan x}
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1

{x : R | 0 < x < PI / 2 /\ Rabs y < tan x}
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1

0 < PI / 2 - u < PI / 2
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
{x : R | 0 < x < PI / 2 /\ Rabs y < tan x}
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1

0 < PI / 2 - u
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
PI / 2 - u < PI / 2
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
{x : R | 0 < x < PI / 2 /\ Rabs y < tan x}
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1

PI / 2 - u < PI / 2
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
{x : R | 0 < x < PI / 2 /\ Rabs y < tan x}
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
dumb:forall x y0 : R, 0 < y0 -> x - y0 < x

PI / 2 - u < PI / 2
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
{x : R | 0 < x < PI / 2 /\ Rabs y < tan x}
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2

{x : R | 0 < x < PI / 2 /\ Rabs y < tan x}
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2

0 < PI / 2 - u < PI / 2 /\ Rabs y < tan (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2

forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
0 < PI / 2 - u < PI / 2 /\ Rabs y < tan (PI / 2 - u)
x, y:R
x0:0 < x
y1:y < 1

x * y < x * 1
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
0 < PI / 2 - u < PI / 2 /\ Rabs y < tan (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x

0 < PI / 2 - u < PI / 2 /\ Rabs y < tan (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x

0 < sin u
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
0 < PI / 2 - u < PI / 2 /\ Rabs y < tan (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x

u < PI
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
0 < PI / 2 - u < PI / 2 /\ Rabs y < tan (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
t:PI / 2 < PI
t':1 < PI / 2

u < PI
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
0 < PI / 2 - u < PI / 2 /\ Rabs y < tan (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u

0 < PI / 2 - u < PI / 2 /\ Rabs y < tan (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u

0 < PI / 2 - u < PI / 2
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
Rabs y < tan (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u

Rabs y < tan (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u

Rabs y < / 2 * / cos (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
/ 2 * / cos (PI / 2 - u) < tan (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u

Rabs y < / 2 * / sin u
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
/ 2 * / cos (PI / 2 - u) < tan (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u

sin u < u
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u
Rabs y < / 2 * / sin u
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
/ 2 * / cos (PI / 2 - u) < tan (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
t1:0 <= u

sin u < u
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u
Rabs y < / 2 * / sin u
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
/ 2 * / cos (PI / 2 - u) < tan (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
t1:0 <= u
t2:u <= 4

sin u < u
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u
Rabs y < / 2 * / sin u
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
/ 2 * / cos (PI / 2 - u) < tan (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
t1:0 <= u
t2:u <= 4
t:sin u <= sin_approx u (2 * (0 + 1))

sin u < u
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u
Rabs y < / 2 * / sin u
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
/ 2 * / cos (PI / 2 - u) < tan (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u

sin_approx u (2 * (0 + 1)) < u
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u
Rabs y < / 2 * / sin u
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
/ 2 * / cos (PI / 2 - u) < tan (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u

u ^ 1 / 1 + -1 * (u ^ 3 / INR (fact 3)) + u ^ 5 / INR (fact 5) < u
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u
Rabs y < / 2 * / sin u
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
/ 2 * / cos (PI / 2 - u) < tan (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u

u * u ^ 0 / 1 + -1 * (u * u ^ 2 / INR (fact 3)) + u * u ^ 4 / INR (fact 5) < u
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u
Rabs y < / 2 * / sin u
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
/ 2 * / cos (PI / 2 - u) < tan (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u

u * u ^ 0 / 1 + u * u ^ 2 / INR (fact 3) * -1 + u * u ^ 4 / INR (fact 5) < u
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u
Rabs y < / 2 * / sin u
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
/ 2 * / cos (PI / 2 - u) < tan (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u

u * (u ^ 0 * 1 + u ^ 2 * (/ INR (fact 3) * -1) + u ^ 4 * / INR (fact 5)) < u
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u
Rabs y < / 2 * / sin u
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
/ 2 * / cos (PI / 2 - u) < tan (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u

u ^ 0 * 1 + u ^ 2 * (/ INR (fact 3) * -1) + u ^ 4 * / INR (fact 5) < 1
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u
Rabs y < / 2 * / sin u
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
/ 2 * / cos (PI / 2 - u) < tan (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u

1 + (u ^ 2 * (/ INR (fact 3) * -1) + u ^ 4 * / INR (fact 5)) < 1 + 0
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u
Rabs y < / 2 * / sin u
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
/ 2 * / cos (PI / 2 - u) < tan (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u

u ^ 2 * (/ INR (fact 3) * -1) + u ^ 4 * / INR (fact 5) < 0
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u
Rabs y < / 2 * / sin u
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
/ 2 * / cos (PI / 2 - u) < tan (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u

u ^ 2 * / INR (fact 3) * -1 + u ^ 4 * / INR (fact 5) < 0
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u
Rabs y < / 2 * / sin u
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
/ 2 * / cos (PI / 2 - u) < tan (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u

- (u ^ 2 * / INR (fact 3)) + u ^ 4 * / INR (fact 5) < - (u ^ 2 * / INR (fact 3)) + u ^ 2 * / INR (fact 3)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u
Rabs y < / 2 * / sin u
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
/ 2 * / cos (PI / 2 - u) < tan (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u

u ^ 4 * / INR (fact 5) < u ^ 2 * / INR (fact 3)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u
Rabs y < / 2 * / sin u
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
/ 2 * / cos (PI / 2 - u) < tan (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:0 < u ^ 2

u ^ 4 * / INR (fact 5) < u ^ 2 * / INR (fact 3)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u
Rabs y < / 2 * / sin u
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
/ 2 * / cos (PI / 2 - u) < tan (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:0 < u ^ 2

u ^ 2 * u ^ 2 * / INR (fact 5) < u ^ 2 * / INR (fact 3)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u
Rabs y < / 2 * / sin u
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
/ 2 * / cos (PI / 2 - u) < tan (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:0 < u ^ 2

u ^ 2 * / INR (fact 5) < / INR (fact 3)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u
Rabs y < / 2 * / sin u
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
/ 2 * / cos (PI / 2 - u) < tan (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:0 < u ^ 2

u ^ 2 * / INR (fact 5) < u ^ 2 * / INR (fact 3)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:0 < u ^ 2
u ^ 2 * / INR (fact 3) < / INR (fact 3)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u
Rabs y < / 2 * / sin u
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
/ 2 * / cos (PI / 2 - u) < tan (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:0 < u ^ 2

/ INR (fact 5) < / INR (fact 3)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:0 < u ^ 2
u ^ 2 * / INR (fact 3) < / INR (fact 3)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u
Rabs y < / 2 * / sin u
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
/ 2 * / cos (PI / 2 - u) < tan (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:0 < u ^ 2

0 < INR (fact 3) * INR (fact 5)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:0 < u ^ 2
INR (fact 3) < INR (fact 5)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:0 < u ^ 2
u ^ 2 * / INR (fact 3) < / INR (fact 3)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u
Rabs y < / 2 * / sin u
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
/ 2 * / cos (PI / 2 - u) < tan (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:0 < u ^ 2

INR (fact 3) < INR (fact 5)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:0 < u ^ 2
u ^ 2 * / INR (fact 3) < / INR (fact 3)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u
Rabs y < / 2 * / sin u
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
/ 2 * / cos (PI / 2 - u) < tan (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:0 < u ^ 2

u ^ 2 * / INR (fact 3) < / INR (fact 3)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u
Rabs y < / 2 * / sin u
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
/ 2 * / cos (PI / 2 - u) < tan (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:0 < u ^ 2

0 < / INR (fact 3)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:0 < u ^ 2
u ^ 2 < 1
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u
Rabs y < / 2 * / sin u
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
/ 2 * / cos (PI / 2 - u) < tan (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:0 < u ^ 2

u ^ 2 < 1
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u
Rabs y < / 2 * / sin u
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
/ 2 * / cos (PI / 2 - u) < tan (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:0 < u ^ 2

u ^ 2 < u
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u
Rabs y < / 2 * / sin u
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
/ 2 * / cos (PI / 2 - u) < tan (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u

Rabs y < / 2 * / sin u
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
/ 2 * / cos (PI / 2 - u) < tan (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u

Rabs y + 1 < / 2 * / sin u
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
/ 2 * / cos (PI / 2 - u) < tan (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u

/ / (Rabs y + 1) < / 2 * / sin u
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
/ 2 * / cos (PI / 2 - u) < tan (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u

/ / (Rabs y + 1) < / (2 * sin u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u
2 <> 0
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u
sin u <> 0
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
/ 2 * / cos (PI / 2 - u) < tan (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u

0 < 2 * sin u * / (Rabs y + 1)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u
2 * sin u < / (Rabs y + 1)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u
2 <> 0
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u
sin u <> 0
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
/ 2 * / cos (PI / 2 - u) < tan (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u

0 < 2 * sin u
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u
0 < / (Rabs y + 1)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u
2 * sin u < / (Rabs y + 1)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u
2 <> 0
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u
sin u <> 0
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
/ 2 * / cos (PI / 2 - u) < tan (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u

0 < / (Rabs y + 1)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u
2 * sin u < / (Rabs y + 1)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u
2 <> 0
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u
sin u <> 0
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
/ 2 * / cos (PI / 2 - u) < tan (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u

2 * sin u < / (Rabs y + 1)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u
2 <> 0
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u
sin u <> 0
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
/ 2 * / cos (PI / 2 - u) < tan (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u

2 * sin u < 2 * u
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u
2 * u = / (Rabs y + 1)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u
2 <> 0
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u
sin u <> 0
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
/ 2 * / cos (PI / 2 - u) < tan (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u

2 * u = / (Rabs y + 1)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u
2 <> 0
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u
sin u <> 0
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
/ 2 * / cos (PI / 2 - u) < tan (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u

2 <> 0
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u
sin u <> 0
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
/ 2 * / cos (PI / 2 - u) < tan (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
H2:sin u < u

sin u <> 0
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
/ 2 * / cos (PI / 2 - u) < tan (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u

/ 2 * / cos (PI / 2 - u) < tan (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u

/ 2 * / cos (PI / 2 - u) < sin (PI / 2 - u) / cos (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
u':=PI / 2:R

0 < / cos (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
u':=PI / 2:R
/ 2 < sin (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
u':=PI / 2:R

0 < cos (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
u':=PI / 2:R
/ 2 < sin (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
u':=PI / 2:R

/ 2 < sin (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
u':=PI / 2:R

u < / 4
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
u':=PI / 2:R
vlt3:u < / 4
/ 2 < sin (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
u':=PI / 2:R

u < / 2 * / 2
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
u':=PI / 2:R
vlt3:u < / 4
/ 2 < sin (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
u':=PI / 2:R

/ (Rabs y + 1) < / 2
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
u':=PI / 2:R
vlt3:u < / 4
/ 2 < sin (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
u':=PI / 2:R

0 < 2 * (Rabs y + 1)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
u':=PI / 2:R
2 < Rabs y + 1
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
u':=PI / 2:R
vlt3:u < / 4
/ 2 < sin (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
u':=PI / 2:R

2 < Rabs y + 1
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
u':=PI / 2:R
vlt3:u < / 4
/ 2 < sin (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
u':=PI / 2:R
vlt3:u < / 4

/ 2 < sin (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
u':=PI / 2:R
vlt3:u < / 4
H2:1 < PI / 2 - u

/ 2 < sin (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
u':=PI / 2:R
vlt3:u < / 4
H2:1 < PI / 2 - u

/ 2 < sin 1
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
u':=PI / 2:R
vlt3:u < / 4
H2:1 < PI / 2 - u
sin 1 < sin (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
u':=PI / 2:R
vlt3:u < / 4
H2:1 < PI / 2 - u
t':1 <= 4

/ 2 < sin 1
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
u':=PI / 2:R
vlt3:u < / 4
H2:1 < PI / 2 - u
sin 1 < sin (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
u':=PI / 2:R
vlt3:u < / 4
H2:1 < PI / 2 - u
t':1 <= 4
t:sin_approx 1 (2 * 0 + 1) <= sin 1

/ 2 < sin 1
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
u':=PI / 2:R
vlt3:u < / 4
H2:1 < PI / 2 - u
sin 1 < sin (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
u':=PI / 2:R
vlt3:u < / 4
H2:1 < PI / 2 - u
t':1 <= 4

/ 2 < sin_approx 1 (2 * 0 + 1)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
u':=PI / 2:R
vlt3:u < / 4
H2:1 < PI / 2 - u
sin 1 < sin (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
u':=PI / 2:R
vlt3:u < / 4
H2:1 < PI / 2 - u
t':1 <= 4

5 / 6 = sin_approx 1 1
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
u':=PI / 2:R
vlt3:u < / 4
H2:1 < PI / 2 - u
sin 1 < sin (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
u':=PI / 2:R
vlt3:u < / 4
H2:1 < PI / 2 - u

sin 1 < sin (PI / 2 - u)
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
u':=PI / 2:R
vlt3:u < / 4
H2:1 < PI / 2 - u

- (PI / 2) <= 1
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
u':=PI / 2:R
vlt3:u < / 4
H2:1 < PI / 2 - u
1 <= PI / 2
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
u':=PI / 2:R
vlt3:u < / 4
H2:1 < PI / 2 - u
- (PI / 2) <= PI / 2 - u
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
u':=PI / 2:R
vlt3:u < / 4
H2:1 < PI / 2 - u
PI / 2 - u <= PI / 2
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
u':=PI / 2:R
vlt3:u < / 4
H2:1 < PI / 2 - u
1 < PI / 2 - u
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
u':=PI / 2:R
vlt3:u < / 4
H2:1 < PI / 2 - u

1 <= PI / 2
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
u':=PI / 2:R
vlt3:u < / 4
H2:1 < PI / 2 - u
- (PI / 2) <= PI / 2 - u
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
u':=PI / 2:R
vlt3:u < / 4
H2:1 < PI / 2 - u
PI / 2 - u <= PI / 2
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
u':=PI / 2:R
vlt3:u < / 4
H2:1 < PI / 2 - u
1 < PI / 2 - u
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
u':=PI / 2:R
vlt3:u < / 4
H2:1 < PI / 2 - u

- (PI / 2) <= PI / 2 - u
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
u':=PI / 2:R
vlt3:u < / 4
H2:1 < PI / 2 - u
PI / 2 - u <= PI / 2
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
u':=PI / 2:R
vlt3:u < / 4
H2:1 < PI / 2 - u
1 < PI / 2 - u
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
u':=PI / 2:R
vlt3:u < / 4
H2:1 < PI / 2 - u

PI / 2 - u <= PI / 2
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
u':=PI / 2:R
vlt3:u < / 4
H2:1 < PI / 2 - u
1 < PI / 2 - u
y:R
Hgt:Rabs y > 1
H:0 < / (Rabs y + 1)
u:=/ 2 * / (Rabs y + 1):R
H0:0 < u
vlt1:/ (Rabs y + 1) < 1
vlt2:u < 1
int:0 < PI / 2 - u < PI / 2
tmp:forall x y0 : R, 0 < x -> y0 < 1 -> x * y0 < x
H1:0 < sin u
u':=PI / 2:R
vlt3:u < / 4
H2:1 < PI / 2 - u

1 < PI / 2 - u
assumption. Qed.

forall x : R, x < PI / 2 -> - PI / 2 < - x

forall x : R, x < PI / 2 -> - PI / 2 < - x
intros x h; rewrite Ropp_div; apply Ropp_lt_contravar; assumption. Qed.

forall x : R, 0 < x -> - x < x

forall x : R, 0 < x -> - x < x
intros; lra. Qed.

forall x y : R, - tan x < y -> tan (- x) < y

forall x y : R, - tan x < y -> tan (- x) < y
intros; rewrite tan_neg; assumption. Qed.
y:R

{x : R | - PI / 2 < x < PI / 2 /\ tan x = y}
y:R

{x : R | - PI / 2 < x < PI / 2 /\ tan x = y}
y, ub:R
ub0:0 < ub
ubpi2:ub < PI / 2
Ptan_ub:Rabs y < tan ub

{x : R | - PI / 2 < x < PI / 2 /\ tan x = y}
y, ub:R
ub0:0 < ub
ubpi2:ub < PI / 2
Ptan_ub:Rabs y < tan ub
pr:=conj (tech_opp_tan ub y (proj2 (Rabs_def2 y (tan ub) Ptan_ub))) (proj1 (Rabs_def2 y (tan ub) Ptan_ub)):tan (- ub) < y < tan ub

{x : R | - PI / 2 < x < PI / 2 /\ tan x = y}
y, ub:R
ub0:0 < ub
ubpi2:ub < PI / 2
Ptan_ub:Rabs y < tan ub
pr:=conj (tech_opp_tan ub y (proj2 (Rabs_def2 y (tan ub) Ptan_ub))) (proj1 (Rabs_def2 y (tan ub) Ptan_ub)):tan (- ub) < y < tan ub
v:R
vl:- ub < v
vu:v < ub
vq:tan v = y

{x : R | - PI / 2 < x < PI / 2 /\ tan x = y}
y, ub:R
ub0:0 < ub
ubpi2:ub < PI / 2
Ptan_ub:Rabs y < tan ub
v:R
vl:- ub < v
vu:v < ub
vq:tan v = y

- PI / 2 < v < PI / 2 /\ tan v = y
split;[rewrite Ropp_div; split; lra | assumption]. Qed. Definition atan x := let (v, _) := pre_atan x in v.

forall x : R, - PI / 2 < atan x < PI / 2

forall x : R, - PI / 2 < atan x < PI / 2
intros x; unfold atan; destruct (pre_atan x) as [v [int _]]; exact int. Qed.

forall x : R, tan (atan x) = x

forall x : R, tan (atan x) = x
intros x; unfold atan; destruct (pre_atan x) as [v [_ q]]; exact q. Qed.

forall x : R, atan (- x) = - atan x

forall x : R, atan (- x) = - atan x
x:R
a:- (PI / 2) < atan (- x)
b:atan (- x) < PI / 2

atan (- x) = - atan x
x:R
a:- (PI / 2) < atan (- x)
b:atan (- x) < PI / 2
c:- (PI / 2) < atan x
d:atan x < PI / 2

atan (- x) = - atan x
x:R
a:- (PI / 2) < atan (- x)
b:atan (- x) < PI / 2
c:- (PI / 2) < atan x
d:atan x < PI / 2

tan (atan (- x)) = tan (- atan x)
rewrite tan_neg, !atan_right_inv; reflexivity. Qed.

forall x : R, derivable_pt atan x

forall x : R, derivable_pt atan x
x:R

derivable_pt atan x
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
P:Rabs x < tan ub

derivable_pt atan x
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
P:Rabs x < tan ub
lb_lt_ub:- ub < ub

derivable_pt atan x
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
P:Rabs x < tan ub
lb_lt_ub:- ub < ub

tan (- ub) < x < tan ub
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
P:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
derivable_pt atan x
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
P:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint':x < tan ub /\ - tan ub < x

tan (- ub) < x < tan ub
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
P:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
derivable_pt atan x
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
P:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub

derivable_pt atan x
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
P:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub

forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
P:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
derivable_pt atan x
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
P:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0

derivable_pt atan x
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
P:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0

forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
P:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
derivable_pt atan x
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub

- ub <= atan y
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
atan y <= ub
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
P:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
derivable_pt atan x
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
abs:atan y < - ub

- ub <= atan y
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
atan y <= ub
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
P:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
derivable_pt atan x
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
abs:atan y < - ub

y < tan (- ub)
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
abs:atan y < - ub
H:y < tan (- ub)
- ub <= atan y
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
atan y <= ub
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
P:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
derivable_pt atan x
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
abs:atan y < - ub

- PI / 2 < atan y
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
abs:atan y < - ub
atan y < - ub
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
abs:atan y < - ub
- ub < PI / 2
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
abs:atan y < - ub
H:y < tan (- ub)
- ub <= atan y
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
atan y <= ub
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
P:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
derivable_pt atan x
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
abs:atan y < - ub

atan y < - ub
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
abs:atan y < - ub
- ub < PI / 2
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
abs:atan y < - ub
H:y < tan (- ub)
- ub <= atan y
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
atan y <= ub
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
P:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
derivable_pt atan x
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
abs:atan y < - ub

- ub < PI / 2
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
abs:atan y < - ub
H:y < tan (- ub)
- ub <= atan y
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
atan y <= ub
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
P:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
derivable_pt atan x
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
abs:atan y < - ub
H:y < tan (- ub)

- ub <= atan y
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
atan y <= ub
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
P:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
derivable_pt atan x
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub

atan y <= ub
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
P:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
derivable_pt atan x
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
abs:ub < atan y

atan y <= ub
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
P:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
derivable_pt atan x
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
abs:ub < atan y

tan ub < y
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
abs:ub < atan y
H:tan ub < y
atan y <= ub
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
P:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
derivable_pt atan x
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
abs:ub < atan y

- PI / 2 < ub
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
abs:ub < atan y
ub < atan y
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
abs:ub < atan y
atan y < PI / 2
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
abs:ub < atan y
H:tan ub < y
atan y <= ub
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
P:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
derivable_pt atan x
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
abs:ub < atan y

ub < atan y
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
abs:ub < atan y
atan y < PI / 2
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
abs:ub < atan y
H:tan ub < y
atan y <= ub
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
P:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
derivable_pt atan x
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
abs:ub < atan y

atan y < PI / 2
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
abs:ub < atan y
H:tan ub < y
atan y <= ub
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
P:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
derivable_pt atan x
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
abs:ub < atan y
H:tan ub < y

atan y <= ub
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
P:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
derivable_pt atan x
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
P:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub

derivable_pt atan x
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
P:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub

forall x0 y : R, - ub <= x0 -> x0 < y -> y <= ub -> tan x0 < tan y
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
P:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
incr:forall x0 y : R, - ub <= x0 -> x0 < y -> y <= ub -> tan x0 < tan y
derivable_pt atan x
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
P:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y0 : R, tan (- ub) <= y0 -> y0 <= tan ub -> - ub <= atan y0 <= ub
y, z:R
l:- ub <= y
yz:y < z
u:z <= ub

- PI / 2 < y
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
P:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y0 : R, tan (- ub) <= y0 -> y0 <= tan ub -> - ub <= atan y0 <= ub
y, z:R
l:- ub <= y
yz:y < z
u:z <= ub
y < z
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
P:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y0 : R, tan (- ub) <= y0 -> y0 <= tan ub -> - ub <= atan y0 <= ub
y, z:R
l:- ub <= y
yz:y < z
u:z <= ub
z < PI / 2
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
P:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
incr:forall x0 y : R, - ub <= x0 -> x0 < y -> y <= ub -> tan x0 < tan y
derivable_pt atan x
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
P:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y0 : R, tan (- ub) <= y0 -> y0 <= tan ub -> - ub <= atan y0 <= ub
y, z:R
l:- ub <= y
yz:y < z
u:z <= ub

y < z
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
P:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y0 : R, tan (- ub) <= y0 -> y0 <= tan ub -> - ub <= atan y0 <= ub
y, z:R
l:- ub <= y
yz:y < z
u:z <= ub
z < PI / 2
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
P:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
incr:forall x0 y : R, - ub <= x0 -> x0 < y -> y <= ub -> tan x0 < tan y
derivable_pt atan x
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
P:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y0 : R, tan (- ub) <= y0 -> y0 <= tan ub -> - ub <= atan y0 <= ub
y, z:R
l:- ub <= y
yz:y < z
u:z <= ub

z < PI / 2
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
P:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
incr:forall x0 y : R, - ub <= x0 -> x0 < y -> y <= ub -> tan x0 < tan y
derivable_pt atan x
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
P:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
incr:forall x0 y : R, - ub <= x0 -> x0 < y -> y <= ub -> tan x0 < tan y

derivable_pt atan x
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
P:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
incr:forall x0 y : R, - ub <= x0 -> x0 < y -> y <= ub -> tan x0 < tan y

forall a : R, - ub <= a <= ub -> derivable_pt tan a
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
P:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
incr:forall x0 y : R, - ub <= x0 -> x0 < y -> y <= ub -> tan x0 < tan y
der:forall a : R, - ub <= a <= ub -> derivable_pt tan a
derivable_pt atan x
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
P:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
incr:forall x0 y : R, - ub <= x0 -> x0 < y -> y <= ub -> tan x0 < tan y
a:R
la:- ub <= a
ua:a <= ub

- PI / 2 < a < PI / 2
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
P:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
incr:forall x0 y : R, - ub <= x0 -> x0 < y -> y <= ub -> tan x0 < tan y
der:forall a : R, - ub <= a <= ub -> derivable_pt tan a
derivable_pt atan x
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
P:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
incr:forall x0 y : R, - ub <= x0 -> x0 < y -> y <= ub -> tan x0 < tan y
der:forall a : R, - ub <= a <= ub -> derivable_pt tan a

derivable_pt atan x
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
P:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
incr:forall x0 y : R, - ub <= x0 -> x0 < y -> y <= ub -> tan x0 < tan y
der:forall a : R, - ub <= a <= ub -> derivable_pt tan a

derive_pt tan (atan x) (derivable_pt_recip_interv_prelim1 tan atan (- ub) ub x lb_lt_ub xint inv_p int_tan incr der) <> 0
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
P:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
incr:forall x0 y : R, - ub <= x0 -> x0 < y -> y <= ub -> tan x0 < tan y
der:forall a : R, - ub <= a <= ub -> derivable_pt tan a
df_neq:derive_pt tan (atan x) (derivable_pt_recip_interv_prelim1 tan atan (- ub) ub x lb_lt_ub xint inv_p int_tan incr der) <> 0
derivable_pt atan x
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
P:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
incr:forall x0 y : R, - ub <= x0 -> x0 < y -> y <= ub -> tan x0 < tan y
der:forall a : R, - ub <= a <= ub -> derivable_pt tan a

derive_pt tan (atan x) (derivable_pt_tan (atan x) (atan_bound x)) <> 0
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
P:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
incr:forall x0 y : R, - ub <= x0 -> x0 < y -> y <= ub -> tan x0 < tan y
der:forall a : R, - ub <= a <= ub -> derivable_pt tan a
df_neq:derive_pt tan (atan x) (derivable_pt_recip_interv_prelim1 tan atan (- ub) ub x lb_lt_ub xint inv_p int_tan incr der) <> 0
derivable_pt atan x
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
P:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
incr:forall x0 y : R, - ub <= x0 -> x0 < y -> y <= ub -> tan x0 < tan y
der:forall a : R, - ub <= a <= ub -> derivable_pt tan a

1 + tan (atan x) ^ 2 <> 0
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
P:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
incr:forall x0 y : R, - ub <= x0 -> x0 < y -> y <= ub -> tan x0 < tan y
der:forall a : R, - ub <= a <= ub -> derivable_pt tan a
df_neq:derive_pt tan (atan x) (derivable_pt_recip_interv_prelim1 tan atan (- ub) ub x lb_lt_ub xint inv_p int_tan incr der) <> 0
derivable_pt atan x
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
P:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
incr:forall x0 y : R, - ub <= x0 -> x0 < y -> y <= ub -> tan x0 < tan y
der:forall a : R, - ub <= a <= ub -> derivable_pt tan a
df_neq:derive_pt tan (atan x) (derivable_pt_recip_interv_prelim1 tan atan (- ub) ub x lb_lt_ub xint inv_p int_tan incr der) <> 0

derivable_pt atan x
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
P:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
incr:forall x0 y : R, - ub <= x0 -> x0 < y -> y <= ub -> tan x0 < tan y
der:forall a : R, - ub <= a <= ub -> derivable_pt tan a
df_neq:derive_pt tan (atan x) (derivable_pt_recip_interv_prelim1 tan atan (- ub) ub x lb_lt_ub xint inv_p int_tan incr der) <> 0

derive_pt tan (atan x) (derivable_pt_recip_interv_prelim1 tan atan (- ub) ub x lb_lt_ub xint inv_p int_tan incr der) <> 0
exact df_neq. Qed.

forall x y : R, x < y -> atan x < atan y

forall x y : R, x < y -> atan x < atan y
x, y:R
d:x < y

atan x < atan y
x, y:R
d:x < y
t1:- PI / 2 < atan x < PI / 2

atan x < atan y
x, y:R
d:x < y
t1:- PI / 2 < atan x < PI / 2
t2:- PI / 2 < atan y < PI / 2

atan x < atan y
x, y:R
d:x < y
t1:- PI / 2 < atan x < PI / 2
t2:- PI / 2 < atan y < PI / 2
lt:atan x < atan y

atan x < atan y
x, y:R
d:x < y
t1:- PI / 2 < atan x < PI / 2
t2:- PI / 2 < atan y < PI / 2
bad:atan y <= atan x
atan x < atan y
x, y:R
d:x < y
t1:- PI / 2 < atan x < PI / 2
t2:- PI / 2 < atan y < PI / 2
bad:atan y <= atan x

atan x < atan y
x, y:R
d:~ y <= x
t1:- PI / 2 < atan x < PI / 2
t2:- PI / 2 < atan y < PI / 2
bad:atan y <= atan x

atan x < atan y
x, y:R
d:~ y <= x
t1:- PI / 2 < atan x < PI / 2
t2:- PI / 2 < atan y < PI / 2
bad:atan y <= atan x

y <= x
x, y:R
d:~ y <= x
t1:- PI / 2 < atan x < PI / 2
t2:- PI / 2 < atan y < PI / 2
bad:atan y <= atan x

tan (atan y) <= tan (atan x)
x, y:R
d:~ y <= x
t1:- PI / 2 < atan x < PI / 2
t2:- PI / 2 < atan y < PI / 2
ylt:atan y < atan x

tan (atan y) <= tan (atan x)
x, y:R
d:~ y <= x
t1:- PI / 2 < atan x < PI / 2
t2:- PI / 2 < atan y < PI / 2
yx:atan y = atan x
tan (atan y) <= tan (atan x)
x, y:R
d:~ y <= x
t1:- PI / 2 < atan x < PI / 2
t2:- PI / 2 < atan y < PI / 2
yx:atan y = atan x

tan (atan y) <= tan (atan x)
solve[rewrite yx; apply Rle_refl]. Qed.

atan 0 = 0

atan 0 = 0

- PI / 2 < 0 < PI / 2

tan (atan 0) = tan 0

tan (atan 0) = tan 0

0 = 0
reflexivity. Qed.

atan 1 = PI / 4

atan 1 = PI / 4
ut:PI > 0

atan 1 = PI / 4
ut:PI > 0
H:- PI / 2 < PI / 4 < PI / 2

atan 1 = PI / 4
ut:PI > 0
H:- PI / 2 < PI / 4 < PI / 2
t:- PI / 2 < atan 1 < PI / 2

atan 1 = PI / 4
ut:PI > 0
H:- PI / 2 < PI / 4 < PI / 2
t:- PI / 2 < atan 1 < PI / 2

tan (atan 1) = tan (PI / 4)
rewrite tan_PI4, atan_right_inv; reflexivity. Qed.
atan's derivative value is the function 1 / (1+x²)

forall x : R, derive_pt atan x (derivable_pt_atan x) = 1 / (1 + x²)

forall x : R, derive_pt atan x (derivable_pt_atan x) = 1 / (1 + x²)
x:R

derive_pt atan x (derivable_pt_atan x) = 1 / (1 + x²)
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub

derive_pt atan x (derivable_pt_atan x) = 1 / (1 + x²)
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub

derive_pt atan x (derivable_pt_atan x) = 1 / (1 + x²)
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub

tan (- ub) < x < tan ub
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
derive_pt atan x (derivable_pt_atan x) = 1 / (1 + x²)
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint':x < tan ub /\ - tan ub < x

tan (- ub) < x < tan ub
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
derive_pt atan x (derivable_pt_atan x) = 1 / (1 + x²)
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub

derive_pt atan x (derivable_pt_atan x) = 1 / (1 + x²)
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub

forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
derive_pt atan x (derivable_pt_atan x) = 1 / (1 + x²)
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0

derive_pt atan x (derivable_pt_atan x) = 1 / (1 + x²)
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0

forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
derive_pt atan x (derivable_pt_atan x) = 1 / (1 + x²)
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub

- ub <= atan y
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
atan y <= ub
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
derive_pt atan x (derivable_pt_atan x) = 1 / (1 + x²)
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
abs:atan y < - ub

- ub <= atan y
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
atan y <= ub
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
derive_pt atan x (derivable_pt_atan x) = 1 / (1 + x²)
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
abs:atan y < - ub

y < tan (- ub)
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
abs:atan y < - ub
H:y < tan (- ub)
- ub <= atan y
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
atan y <= ub
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
derive_pt atan x (derivable_pt_atan x) = 1 / (1 + x²)
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
abs:atan y < - ub

- PI / 2 < atan y
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
abs:atan y < - ub
atan y < - ub
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
abs:atan y < - ub
- ub < PI / 2
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
abs:atan y < - ub
H:y < tan (- ub)
- ub <= atan y
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
atan y <= ub
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
derive_pt atan x (derivable_pt_atan x) = 1 / (1 + x²)
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
abs:atan y < - ub

atan y < - ub
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
abs:atan y < - ub
- ub < PI / 2
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
abs:atan y < - ub
H:y < tan (- ub)
- ub <= atan y
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
atan y <= ub
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
derive_pt atan x (derivable_pt_atan x) = 1 / (1 + x²)
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
abs:atan y < - ub

- ub < PI / 2
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
abs:atan y < - ub
H:y < tan (- ub)
- ub <= atan y
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
atan y <= ub
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
derive_pt atan x (derivable_pt_atan x) = 1 / (1 + x²)
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
abs:atan y < - ub
H:y < tan (- ub)

- ub <= atan y
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
atan y <= ub
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
derive_pt atan x (derivable_pt_atan x) = 1 / (1 + x²)
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub

atan y <= ub
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
derive_pt atan x (derivable_pt_atan x) = 1 / (1 + x²)
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
abs:ub < atan y

atan y <= ub
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
derive_pt atan x (derivable_pt_atan x) = 1 / (1 + x²)
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
abs:ub < atan y

tan ub < y
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
abs:ub < atan y
H:tan ub < y
atan y <= ub
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
derive_pt atan x (derivable_pt_atan x) = 1 / (1 + x²)
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
abs:ub < atan y

- PI / 2 < ub
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
abs:ub < atan y
ub < atan y
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
abs:ub < atan y
atan y < PI / 2
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
abs:ub < atan y
H:tan ub < y
atan y <= ub
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
derive_pt atan x (derivable_pt_atan x) = 1 / (1 + x²)
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
abs:ub < atan y

ub < atan y
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
abs:ub < atan y
atan y < PI / 2
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
abs:ub < atan y
H:tan ub < y
atan y <= ub
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
derive_pt atan x (derivable_pt_atan x) = 1 / (1 + x²)
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
abs:ub < atan y

atan y < PI / 2
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
abs:ub < atan y
H:tan ub < y
atan y <= ub
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
derive_pt atan x (derivable_pt_atan x) = 1 / (1 + x²)
ub:R
ub0:0 < ub
ubpi:ub < PI / 2
y:R
lo:tan (- ub) <= y
up:y <= tan ub
abs:ub < atan y
H:tan ub < y

atan y <= ub
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
derive_pt atan x (derivable_pt_atan x) = 1 / (1 + x²)
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub

derive_pt atan x (derivable_pt_atan x) = 1 / (1 + x²)
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub

forall x0 y : R, - ub <= x0 -> x0 < y -> y <= ub -> tan x0 < tan y
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
incr:forall x0 y : R, - ub <= x0 -> x0 < y -> y <= ub -> tan x0 < tan y
derive_pt atan x (derivable_pt_atan x) = 1 / (1 + x²)
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y0 : R, tan (- ub) <= y0 -> y0 <= tan ub -> - ub <= atan y0 <= ub
y, z:R
l:- ub <= y
yz:y < z
u:z <= ub

- PI / 2 < y
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y0 : R, tan (- ub) <= y0 -> y0 <= tan ub -> - ub <= atan y0 <= ub
y, z:R
l:- ub <= y
yz:y < z
u:z <= ub
y < z
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y0 : R, tan (- ub) <= y0 -> y0 <= tan ub -> - ub <= atan y0 <= ub
y, z:R
l:- ub <= y
yz:y < z
u:z <= ub
z < PI / 2
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
incr:forall x0 y : R, - ub <= x0 -> x0 < y -> y <= ub -> tan x0 < tan y
derive_pt atan x (derivable_pt_atan x) = 1 / (1 + x²)
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y0 : R, tan (- ub) <= y0 -> y0 <= tan ub -> - ub <= atan y0 <= ub
y, z:R
l:- ub <= y
yz:y < z
u:z <= ub

y < z
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y0 : R, tan (- ub) <= y0 -> y0 <= tan ub -> - ub <= atan y0 <= ub
y, z:R
l:- ub <= y
yz:y < z
u:z <= ub
z < PI / 2
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
incr:forall x0 y : R, - ub <= x0 -> x0 < y -> y <= ub -> tan x0 < tan y
derive_pt atan x (derivable_pt_atan x) = 1 / (1 + x²)
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y0 : R, tan (- ub) <= y0 -> y0 <= tan ub -> - ub <= atan y0 <= ub
y, z:R
l:- ub <= y
yz:y < z
u:z <= ub

z < PI / 2
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
incr:forall x0 y : R, - ub <= x0 -> x0 < y -> y <= ub -> tan x0 < tan y
derive_pt atan x (derivable_pt_atan x) = 1 / (1 + x²)
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
incr:forall x0 y : R, - ub <= x0 -> x0 < y -> y <= ub -> tan x0 < tan y

derive_pt atan x (derivable_pt_atan x) = 1 / (1 + x²)
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
incr:forall x0 y : R, - ub <= x0 -> x0 < y -> y <= ub -> tan x0 < tan y

forall a : R, - ub <= a <= ub -> derivable_pt tan a
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
incr:forall x0 y : R, - ub <= x0 -> x0 < y -> y <= ub -> tan x0 < tan y
der:forall a : R, - ub <= a <= ub -> derivable_pt tan a
derive_pt atan x (derivable_pt_atan x) = 1 / (1 + x²)
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
incr:forall x0 y : R, - ub <= x0 -> x0 < y -> y <= ub -> tan x0 < tan y
a:R
la:- ub <= a
ua:a <= ub

- PI / 2 < a < PI / 2
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
incr:forall x0 y : R, - ub <= x0 -> x0 < y -> y <= ub -> tan x0 < tan y
der:forall a : R, - ub <= a <= ub -> derivable_pt tan a
derive_pt atan x (derivable_pt_atan x) = 1 / (1 + x²)
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
incr:forall x0 y : R, - ub <= x0 -> x0 < y -> y <= ub -> tan x0 < tan y
der:forall a : R, - ub <= a <= ub -> derivable_pt tan a

derive_pt atan x (derivable_pt_atan x) = 1 / (1 + x²)
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
incr:forall x0 y : R, - ub <= x0 -> x0 < y -> y <= ub -> tan x0 < tan y
der:forall a : R, - ub <= a <= ub -> derivable_pt tan a

derive_pt tan (atan x) (derivable_pt_recip_interv_prelim1 tan atan (- ub) ub x lb_lt_ub xint inv_p int_tan incr der) <> 0
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
incr:forall x0 y : R, - ub <= x0 -> x0 < y -> y <= ub -> tan x0 < tan y
der:forall a : R, - ub <= a <= ub -> derivable_pt tan a
df_neq:derive_pt tan (atan x) (derivable_pt_recip_interv_prelim1 tan atan (- ub) ub x lb_lt_ub xint inv_p int_tan incr der) <> 0
derive_pt atan x (derivable_pt_atan x) = 1 / (1 + x²)
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
incr:forall x0 y : R, - ub <= x0 -> x0 < y -> y <= ub -> tan x0 < tan y
der:forall a : R, - ub <= a <= ub -> derivable_pt tan a

derive_pt tan (atan x) (derivable_pt_tan (atan x) (atan_bound x)) <> 0
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
incr:forall x0 y : R, - ub <= x0 -> x0 < y -> y <= ub -> tan x0 < tan y
der:forall a : R, - ub <= a <= ub -> derivable_pt tan a
df_neq:derive_pt tan (atan x) (derivable_pt_recip_interv_prelim1 tan atan (- ub) ub x lb_lt_ub xint inv_p int_tan incr der) <> 0
derive_pt atan x (derivable_pt_atan x) = 1 / (1 + x²)
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
incr:forall x0 y : R, - ub <= x0 -> x0 < y -> y <= ub -> tan x0 < tan y
der:forall a : R, - ub <= a <= ub -> derivable_pt tan a

1 + tan (atan x) ^ 2 <> 0
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
incr:forall x0 y : R, - ub <= x0 -> x0 < y -> y <= ub -> tan x0 < tan y
der:forall a : R, - ub <= a <= ub -> derivable_pt tan a
df_neq:derive_pt tan (atan x) (derivable_pt_recip_interv_prelim1 tan atan (- ub) ub x lb_lt_ub xint inv_p int_tan incr der) <> 0
derive_pt atan x (derivable_pt_atan x) = 1 / (1 + x²)
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
incr:forall x0 y : R, - ub <= x0 -> x0 < y -> y <= ub -> tan x0 < tan y
der:forall a : R, - ub <= a <= ub -> derivable_pt tan a
df_neq:derive_pt tan (atan x) (derivable_pt_recip_interv_prelim1 tan atan (- ub) ub x lb_lt_ub xint inv_p int_tan incr der) <> 0

derive_pt atan x (derivable_pt_atan x) = 1 / (1 + x²)
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
incr:forall x0 y : R, - ub <= x0 -> x0 < y -> y <= ub -> tan x0 < tan y
der:forall a : R, - ub <= a <= ub -> derivable_pt tan a
df_neq:derive_pt tan (atan x) (derivable_pt_recip_interv_prelim1 tan atan (- ub) ub x lb_lt_ub xint inv_p int_tan incr der) <> 0
t:derive_pt atan x (derivable_pt_recip_interv tan atan (- ub) ub x lb_lt_ub xint inv_p int_tan incr der df_neq) = 1 / derive_pt tan (atan x) (der (atan x) (derive_pt_recip_interv_prelim1_1 tan atan (- ub) ub x lb_lt_ub xint incr int_tan inv_p))

derive_pt atan x (derivable_pt_atan x) = 1 / (1 + x²)
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
incr:forall x0 y : R, - ub <= x0 -> x0 < y -> y <= ub -> tan x0 < tan y
der:forall a : R, - ub <= a <= ub -> derivable_pt tan a
df_neq:derive_pt tan (atan x) (derivable_pt_recip_interv_prelim1 tan atan (- ub) ub x lb_lt_ub xint inv_p int_tan incr der) <> 0
t:derive_pt atan x (derivable_pt_recip_interv tan atan (- ub) ub x lb_lt_ub xint inv_p int_tan incr der df_neq) = 1 / derive_pt tan (atan x) (der (atan x) (derive_pt_recip_interv_prelim1_1 tan atan (- ub) ub x lb_lt_ub xint incr int_tan inv_p))

derive_pt atan x (derivable_pt_recip_interv tan atan (- ub) ub x lb_lt_ub xint inv_p int_tan incr der df_neq) = 1 / (1 + x²)
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
incr:forall x0 y : R, - ub <= x0 -> x0 < y -> y <= ub -> tan x0 < tan y
der:forall a : R, - ub <= a <= ub -> derivable_pt tan a
df_neq:derive_pt tan (atan x) (derivable_pt_recip_interv_prelim1 tan atan (- ub) ub x lb_lt_ub xint inv_p int_tan incr der) <> 0
t:derive_pt atan x (derivable_pt_recip_interv tan atan (- ub) ub x lb_lt_ub xint inv_p int_tan incr der df_neq) = 1 / derive_pt tan (atan x) (der (atan x) (derive_pt_recip_interv_prelim1_1 tan atan (- ub) ub x lb_lt_ub xint incr int_tan inv_p))

1 / derive_pt tan (atan x) (der (atan x) (derive_pt_recip_interv_prelim1_1 tan atan (- ub) ub x lb_lt_ub xint incr int_tan inv_p)) = 1 / (1 + x²)
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
incr:forall x0 y : R, - ub <= x0 -> x0 < y -> y <= ub -> tan x0 < tan y
der:forall a : R, - ub <= a <= ub -> derivable_pt tan a
df_neq:derive_pt tan (atan x) (derivable_pt_recip_interv_prelim1 tan atan (- ub) ub x lb_lt_ub xint inv_p int_tan incr der) <> 0
t:derive_pt atan x (derivable_pt_recip_interv tan atan (- ub) ub x lb_lt_ub xint inv_p int_tan incr der df_neq) = 1 / derive_pt tan (atan x) (der (atan x) (derive_pt_recip_interv_prelim1_1 tan atan (- ub) ub x lb_lt_ub xint incr int_tan inv_p))
t':- PI / 2 < atan x < PI / 2

1 / derive_pt tan (atan x) (der (atan x) (derive_pt_recip_interv_prelim1_1 tan atan (- ub) ub x lb_lt_ub xint incr int_tan inv_p)) = 1 / (1 + x²)
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
incr:forall x0 y : R, - ub <= x0 -> x0 < y -> y <= ub -> tan x0 < tan y
der:forall a : R, - ub <= a <= ub -> derivable_pt tan a
df_neq:derive_pt tan (atan x) (derivable_pt_recip_interv_prelim1 tan atan (- ub) ub x lb_lt_ub xint inv_p int_tan incr der) <> 0
t:derive_pt atan x (derivable_pt_recip_interv tan atan (- ub) ub x lb_lt_ub xint inv_p int_tan incr der df_neq) = 1 / derive_pt tan (atan x) (der (atan x) (derive_pt_recip_interv_prelim1_1 tan atan (- ub) ub x lb_lt_ub xint incr int_tan inv_p))
t':- PI / 2 < atan x < PI / 2

1 / derive_pt tan (atan x) (derivable_pt_tan (atan x) t') = 1 / (1 + x²)
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
incr:forall x0 y : R, - ub <= x0 -> x0 < y -> y <= ub -> tan x0 < tan y
der:forall a : R, - ub <= a <= ub -> derivable_pt tan a
df_neq:derive_pt tan (atan x) (derivable_pt_recip_interv_prelim1 tan atan (- ub) ub x lb_lt_ub xint inv_p int_tan incr der) <> 0
t:derive_pt atan x (derivable_pt_recip_interv tan atan (- ub) ub x lb_lt_ub xint inv_p int_tan incr der df_neq) = 1 / derive_pt tan (atan x) (der (atan x) (derive_pt_recip_interv_prelim1_1 tan atan (- ub) ub x lb_lt_ub xint incr int_tan inv_p))
t':- PI / 2 < atan x < PI / 2

1 / (1 + x ^ 2) = 1 / (1 + x²)
x, ub:R
ub0:0 < ub
ubpi:ub < PI / 2
Pub:Rabs x < tan ub
lb_lt_ub:- ub < ub
xint:tan (- ub) < x < tan ub
inv_p:forall x0 : R, tan (- ub) <= x0 -> x0 <= tan ub -> comp tan atan x0 = id x0
int_tan:forall y : R, tan (- ub) <= y -> y <= tan ub -> - ub <= atan y <= ub
incr:forall x0 y : R, - ub <= x0 -> x0 < y -> y <= ub -> tan x0 < tan y
der:forall a : R, - ub <= a <= ub -> derivable_pt tan a
df_neq:derive_pt tan (atan x) (derivable_pt_recip_interv_prelim1 tan atan (- ub) ub x lb_lt_ub xint inv_p int_tan incr der) <> 0
t:derive_pt atan x (derivable_pt_recip_interv tan atan (- ub) ub x lb_lt_ub xint inv_p int_tan incr der df_neq) = 1 / derive_pt tan (atan x) (der (atan x) (derive_pt_recip_interv_prelim1_1 tan atan (- ub) ub x lb_lt_ub xint incr int_tan inv_p))
t':- PI / 2 < atan x < PI / 2

1 / (1 + x ^ 2) = 1 / (1 + x ^ 2)
reflexivity. Qed.

forall x : R, derivable_pt_lim atan x (/ (1 + x ^ 2))

forall x : R, derivable_pt_lim atan x (/ (1 + x ^ 2))
x:R

derivable_pt_lim atan x (/ (1 + x ^ 2))
x:R

derive_pt atan x (derivable_pt_atan x) = / (1 + x ^ 2)
x:R

derive_pt atan x (derivable_pt_atan x) = / (1 + x * x)
x:R

derive_pt atan x (derivable_pt_atan x) = 1 * / (1 + x * x)
apply derive_pt_atan. Qed.

Definition of the arctangent function as the sum of the arctan power series

(* Proof taken from Guillaume Melquiond's interval package for Coq *)

Definition Ratan_seq x :=  fun n => (x ^ (2 * n + 1) / INR (2 * n + 1))%R.


forall x : R, 0 <= x <= 1 -> Un_decreasing (Ratan_seq x)

forall x : R, 0 <= x <= 1 -> Un_decreasing (Ratan_seq x)
x:R
Hx:0 <= x <= 1
n:nat

Ratan_seq x (S n) <= Ratan_seq x n
x:R
Hx:0 <= x <= 1
n:nat

x ^ (2 * S n + 1) * / INR (2 * S n + 1) <= x ^ (2 * n + 1) * / INR (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat

0 <= x ^ (2 * S n + 1)
x:R
Hx:0 <= x <= 1
n:nat
0 <= / INR (2 * S n + 1)
x:R
Hx:0 <= x <= 1
n:nat
x ^ (2 * S n + 1) <= x ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
/ INR (2 * S n + 1) <= / INR (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat

0 <= x
x:R
Hx:0 <= x <= 1
n:nat
0 <= / INR (2 * S n + 1)
x:R
Hx:0 <= x <= 1
n:nat
x ^ (2 * S n + 1) <= x ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
/ INR (2 * S n + 1) <= / INR (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat

0 <= / INR (2 * S n + 1)
x:R
Hx:0 <= x <= 1
n:nat
x ^ (2 * S n + 1) <= x ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
/ INR (2 * S n + 1) <= / INR (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat

0 < / INR (2 * S n + 1)
x:R
Hx:0 <= x <= 1
n:nat
x ^ (2 * S n + 1) <= x ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
/ INR (2 * S n + 1) <= / INR (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat

0 < INR (2 * S n + 1)
x:R
Hx:0 <= x <= 1
n:nat
x ^ (2 * S n + 1) <= x ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
/ INR (2 * S n + 1) <= / INR (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat

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

x ^ (2 * S n + 1) <= x ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
/ INR (2 * S n + 1) <= / INR (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 < x

x ^ (2 * S n + 1) <= x ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 = x
x ^ (2 * S n + 1) <= x ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
/ INR (2 * S n + 1) <= / INR (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 < x
Hx2:x < 1

x ^ (2 * S n + 1) <= x ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 < x
Hx2:x = 1
x ^ (2 * S n + 1) <= x ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 = x
x ^ (2 * S n + 1) <= x ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
/ INR (2 * S n + 1) <= / INR (2 * n + 1)
(* . 0 < x < 1 *)
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 < x
Hx2:x < 1

(/ / x) ^ (2 * S n + 1) <= (/ / x) ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 < x
Hx2:x < 1
x <> 0
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 < x
Hx2:x = 1
x ^ (2 * S n + 1) <= x ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 = x
x ^ (2 * S n + 1) <= x ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
/ INR (2 * S n + 1) <= / INR (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 < x
Hx2:x < 1
H:/ x <> 0

(/ / x) ^ (2 * S n + 1) <= (/ / x) ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 < x
Hx2:x < 1
x <> 0
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 < x
Hx2:x = 1
x ^ (2 * S n + 1) <= x ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 = x
x ^ (2 * S n + 1) <= x ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
/ INR (2 * S n + 1) <= / INR (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 < x
Hx2:x < 1
H:/ x <> 0

/ (/ x) ^ (2 * S n + 1) <= / (/ x) ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 < x
Hx2:x < 1
x <> 0
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 < x
Hx2:x = 1
x ^ (2 * S n + 1) <= x ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 = x
x ^ (2 * S n + 1) <= x ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
/ INR (2 * S n + 1) <= / INR (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 < x
Hx2:x < 1
H:/ x <> 0

/ (/ x) ^ (2 * S n + 1) < / (/ x) ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 < x
Hx2:x < 1
x <> 0
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 < x
Hx2:x = 1
x ^ (2 * S n + 1) <= x ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 = x
x ^ (2 * S n + 1) <= x ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
/ INR (2 * S n + 1) <= / INR (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 < x
Hx2:x < 1
H:/ x <> 0

0 < (/ x) ^ (2 * n + 1) * (/ x) ^ (2 * S n + 1)
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 < x
Hx2:x < 1
H:/ x <> 0
(/ x) ^ (2 * n + 1) < (/ x) ^ (2 * S n + 1)
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 < x
Hx2:x < 1
x <> 0
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 < x
Hx2:x = 1
x ^ (2 * S n + 1) <= x ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 = x
x ^ (2 * S n + 1) <= x ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
/ INR (2 * S n + 1) <= / INR (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 < x
Hx2:x < 1
H:/ x <> 0

(/ x) ^ (2 * n + 1) < (/ x) ^ (2 * S n + 1)
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 < x
Hx2:x < 1
x <> 0
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 < x
Hx2:x = 1
x ^ (2 * S n + 1) <= x ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 = x
x ^ (2 * S n + 1) <= x ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
/ INR (2 * S n + 1) <= / INR (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 < x
Hx2:x < 1
H:/ x <> 0

1 < / x
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 < x
Hx2:x < 1
H:/ x <> 0
(2 * n + 1 < 2 * S n + 1)%nat
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 < x
Hx2:x < 1
x <> 0
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 < x
Hx2:x = 1
x ^ (2 * S n + 1) <= x ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 = x
x ^ (2 * S n + 1) <= x ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
/ INR (2 * S n + 1) <= / INR (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 < x
Hx2:x < 1
H:/ x <> 0

/ 1 < / x
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 < x
Hx2:x < 1
H:/ x <> 0
(2 * n + 1 < 2 * S n + 1)%nat
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 < x
Hx2:x < 1
x <> 0
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 < x
Hx2:x = 1
x ^ (2 * S n + 1) <= x ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 = x
x ^ (2 * S n + 1) <= x ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
/ INR (2 * S n + 1) <= / INR (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 < x
Hx2:x < 1
H:/ x <> 0

0 < x * 1
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 < x
Hx2:x < 1
H:/ x <> 0
x < 1
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 < x
Hx2:x < 1
H:/ x <> 0
(2 * n + 1 < 2 * S n + 1)%nat
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 < x
Hx2:x < 1
x <> 0
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 < x
Hx2:x = 1
x ^ (2 * S n + 1) <= x ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 = x
x ^ (2 * S n + 1) <= x ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
/ INR (2 * S n + 1) <= / INR (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 < x
Hx2:x < 1
H:/ x <> 0

0 < x
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 < x
Hx2:x < 1
H:/ x <> 0
x < 1
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 < x
Hx2:x < 1
H:/ x <> 0
(2 * n + 1 < 2 * S n + 1)%nat
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 < x
Hx2:x < 1
x <> 0
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 < x
Hx2:x = 1
x ^ (2 * S n + 1) <= x ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 = x
x ^ (2 * S n + 1) <= x ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
/ INR (2 * S n + 1) <= / INR (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 < x
Hx2:x < 1
H:/ x <> 0

x < 1
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 < x
Hx2:x < 1
H:/ x <> 0
(2 * n + 1 < 2 * S n + 1)%nat
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 < x
Hx2:x < 1
x <> 0
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 < x
Hx2:x = 1
x ^ (2 * S n + 1) <= x ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 = x
x ^ (2 * S n + 1) <= x ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
/ INR (2 * S n + 1) <= / INR (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 < x
Hx2:x < 1
H:/ x <> 0

(2 * n + 1 < 2 * S n + 1)%nat
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 < x
Hx2:x < 1
x <> 0
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 < x
Hx2:x = 1
x ^ (2 * S n + 1) <= x ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 = x
x ^ (2 * S n + 1) <= x ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
/ INR (2 * S n + 1) <= / INR (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 < x
Hx2:x < 1

x <> 0
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 < x
Hx2:x = 1
x ^ (2 * S n + 1) <= x ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 = x
x ^ (2 * S n + 1) <= x ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
/ INR (2 * S n + 1) <= / INR (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 < x
Hx2:x < 1

x > 0
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 < x
Hx2:x = 1
x ^ (2 * S n + 1) <= x ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 = x
x ^ (2 * S n + 1) <= x ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
/ INR (2 * S n + 1) <= / INR (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 < x
Hx2:x = 1

x ^ (2 * S n + 1) <= x ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 = x
x ^ (2 * S n + 1) <= x ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
/ INR (2 * S n + 1) <= / INR (2 * n + 1)
(* . x = 1 *)
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 < x
Hx2:x = 1

1 ^ (2 * S n + 1) <= 1 ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 = x
x ^ (2 * S n + 1) <= x ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
/ INR (2 * S n + 1) <= / INR (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 < x
Hx2:x = 1

1 <= 1
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 = x
x ^ (2 * S n + 1) <= x ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
/ INR (2 * S n + 1) <= / INR (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 = x

x ^ (2 * S n + 1) <= x ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
/ INR (2 * S n + 1) <= / INR (2 * n + 1)
(* . x = 0 *)
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 = x

0 ^ (2 * S n + 1) <= 0 ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
/ INR (2 * S n + 1) <= / INR (2 * n + 1)
x:R
Hx:0 <= x <= 1
n:nat
Hx1:0 = x

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

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

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

0 < INR (2 * n + 1) * INR (2 * S n + 1)
x:R
Hx:0 <= x <= 1
n:nat
INR (2 * n + 1) < INR (2 * S n + 1)
x:R
Hx:0 <= x <= 1
n:nat

INR (2 * n + 1) < INR (2 * S n + 1)
x:R
Hx:0 <= x <= 1
n:nat

(2 * n + 1 < 2 * S n + 1)%nat
omega. Qed.

forall x : R, 0 <= x <= 1 -> Un_cv (Ratan_seq x) 0

forall x : R, 0 <= x <= 1 -> Un_cv (Ratan_seq x) 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0

exists N : nat, forall n : nat, (n >= N)%nat -> R_dist (Ratan_seq x n) 0 < eps
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
HN:IZR (up (/ eps)) > / eps

exists N : nat, forall n : nat, (n >= N)%nat -> R_dist (Ratan_seq x n) 0 < eps
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
HN:IZR (up (/ eps)) > / eps

(0 < up (/ eps))%Z
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
HN:IZR (up (/ eps)) > / eps
H:(0 < up (/ eps))%Z
exists N : nat, forall n : nat, (n >= N)%nat -> R_dist (Ratan_seq x n) 0 < eps
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
HN:IZR (up (/ eps)) > / eps

0 < IZR (up (/ eps))
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
HN:IZR (up (/ eps)) > / eps
H:(0 < up (/ eps))%Z
exists N : nat, forall n : nat, (n >= N)%nat -> R_dist (Ratan_seq x n) 0 < eps
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
HN:IZR (up (/ eps)) > / eps

0 < / eps
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
HN:IZR (up (/ eps)) > / eps
H:(0 < up (/ eps))%Z
exists N : nat, forall n : nat, (n >= N)%nat -> R_dist (Ratan_seq x n) 0 < eps
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
HN:IZR (up (/ eps)) > / eps

0 < eps
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
HN:IZR (up (/ eps)) > / eps
H:(0 < up (/ eps))%Z
exists N : nat, forall n : nat, (n >= N)%nat -> R_dist (Ratan_seq x n) 0 < eps
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
HN:IZR (up (/ eps)) > / eps
H:(0 < up (/ eps))%Z

exists N : nat, forall n : nat, (n >= N)%nat -> R_dist (Ratan_seq x n) 0 < eps
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
HN:IZR (up (/ eps)) > / eps
p:positive
H:(0 < Z.pos p)%Z
H0:up (/ eps) = Z.pos p

exists N : nat, forall n : nat, (n >= N)%nat -> R_dist (Ratan_seq x n) 0 < eps
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
H:(0 < Z.pos p)%Z
H0:up (/ eps) = Z.pos p

exists N : nat, forall n : nat, (n >= N)%nat -> R_dist (Ratan_seq x n) 0 < eps
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
H:(0 < Z.pos p)%Z
H0:up (/ eps) = Z.pos p

exists N : nat, forall n : nat, (n >= N)%nat -> R_dist (Ratan_seq x n) 0 < eps
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
H:(0 < Z.pos p)%Z
H0:up (/ eps) = Z.pos p
N:=Pos.to_nat p:nat

exists N0 : nat, forall n : nat, (n >= N0)%nat -> R_dist (Ratan_seq x n) 0 < eps
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
H:(0 < Z.pos p)%Z
H0:up (/ eps) = Z.pos p
N:=Pos.to_nat p:nat

exists N0 : nat, forall n : nat, (n >= N0)%nat -> R_dist (Ratan_seq x n) 0 < eps
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat

exists N0 : nat, forall n : nat, (n >= N0)%nat -> R_dist (Ratan_seq x n) 0 < eps
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat

forall n : nat, (n >= N)%nat -> R_dist (Ratan_seq x n) 0 < eps
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat

R_dist (Ratan_seq x n) 0 < eps
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat

Rabs (Ratan_seq x n - 0) < eps
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat

Rabs (Ratan_seq x n) < eps
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat

Rabs (x ^ (2 * n + 1) / INR (2 * n + 1)) < eps
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat

x ^ (2 * n + 1) / INR (2 * n + 1) < eps
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
x ^ (2 * n + 1) / INR (2 * n + 1) >= 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat

x ^ (2 * n + 1) / INR (2 * n + 1) <= 1 ^ (2 * n + 1) / INR (2 * n + 1)
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
1 ^ (2 * n + 1) / INR (2 * n + 1) < eps
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
x ^ (2 * n + 1) / INR (2 * n + 1) >= 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat

x ^ (2 * n + 1) * / INR (2 * n + 1) <= 1 ^ (2 * n + 1) * / INR (2 * n + 1)
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
1 ^ (2 * n + 1) / INR (2 * n + 1) < eps
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
x ^ (2 * n + 1) / INR (2 * n + 1) >= 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat

0 <= / INR (2 * n + 1)
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
x ^ (2 * n + 1) <= 1 ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
1 ^ (2 * n + 1) / INR (2 * n + 1) < eps
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
x ^ (2 * n + 1) / INR (2 * n + 1) >= 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat

0 < / INR (2 * n + 1)
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
x ^ (2 * n + 1) <= 1 ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
1 ^ (2 * n + 1) / INR (2 * n + 1) < eps
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
x ^ (2 * n + 1) / INR (2 * n + 1) >= 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat

0 < INR (2 * n + 1)
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
x ^ (2 * n + 1) <= 1 ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
1 ^ (2 * n + 1) / INR (2 * n + 1) < eps
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
x ^ (2 * n + 1) / INR (2 * n + 1) >= 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat

(0 < 2 * n + 1)%nat
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
x ^ (2 * n + 1) <= 1 ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
1 ^ (2 * n + 1) / INR (2 * n + 1) < eps
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
x ^ (2 * n + 1) / INR (2 * n + 1) >= 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat

x ^ (2 * n + 1) <= 1 ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
1 ^ (2 * n + 1) / INR (2 * n + 1) < eps
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
x ^ (2 * n + 1) / INR (2 * n + 1) >= 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat

0 <= x <= 1
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
1 ^ (2 * n + 1) / INR (2 * n + 1) < eps
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
x ^ (2 * n + 1) / INR (2 * n + 1) >= 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat

1 ^ (2 * n + 1) / INR (2 * n + 1) < eps
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
x ^ (2 * n + 1) / INR (2 * n + 1) >= 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat

1 / INR (2 * n + 1) < eps
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
x ^ (2 * n + 1) / INR (2 * n + 1) >= 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat

1 / INR (2 * n + 1) <= / INR (2 * N + 1)
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
/ INR (2 * N + 1) < eps
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
x ^ (2 * n + 1) / INR (2 * n + 1) >= 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat

1 * / INR (2 * n + 1) <= / INR (2 * N + 1)
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
/ INR (2 * N + 1) < eps
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
x ^ (2 * n + 1) / INR (2 * n + 1) >= 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat

/ INR (2 * n + 1) <= / INR (2 * N + 1)
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
/ INR (2 * N + 1) < eps
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
x ^ (2 * n + 1) / INR (2 * n + 1) >= 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat

0 < INR (2 * N + 1)
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
INR (2 * N + 1) <= INR (2 * n + 1)
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
/ INR (2 * N + 1) < eps
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
x ^ (2 * n + 1) / INR (2 * n + 1) >= 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat

(0 < 2 * N + 1)%nat
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
INR (2 * N + 1) <= INR (2 * n + 1)
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
/ INR (2 * N + 1) < eps
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
x ^ (2 * n + 1) / INR (2 * n + 1) >= 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat

INR (2 * N + 1) <= INR (2 * n + 1)
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
/ INR (2 * N + 1) < eps
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
x ^ (2 * n + 1) / INR (2 * n + 1) >= 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat

(2 * N + 1 <= 2 * n + 1)%nat
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
/ INR (2 * N + 1) < eps
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
x ^ (2 * n + 1) / INR (2 * n + 1) >= 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat

/ INR (2 * N + 1) < eps
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
x ^ (2 * n + 1) / INR (2 * n + 1) >= 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat

/ INR (2 * N + 1) < / / eps
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
eps <> 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
x ^ (2 * n + 1) / INR (2 * n + 1) >= 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat

0 < / eps * INR (2 * N + 1)
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
/ eps < INR (2 * N + 1)
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
eps <> 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
x ^ (2 * n + 1) / INR (2 * n + 1) >= 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat

0 < / eps
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
0 < INR (2 * N + 1)
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
/ eps < INR (2 * N + 1)
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
eps <> 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
x ^ (2 * n + 1) / INR (2 * n + 1) >= 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat

0 < INR (2 * N + 1)
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
/ eps < INR (2 * N + 1)
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
eps <> 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
x ^ (2 * n + 1) / INR (2 * n + 1) >= 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat

(0 < 2 * N + 1)%nat
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
/ eps < INR (2 * N + 1)
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
eps <> 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
x ^ (2 * n + 1) / INR (2 * n + 1) >= 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat

/ eps < INR (2 * N + 1)
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
eps <> 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
x ^ (2 * n + 1) / INR (2 * n + 1) >= 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat

/ eps < INR N
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
INR N < INR (2 * N + 1)
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
eps <> 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
x ^ (2 * n + 1) / INR (2 * n + 1) >= 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
H:IZR (up (/ eps)) > / eps

/ eps < INR N
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
INR N < INR (2 * N + 1)
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
eps <> 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
x ^ (2 * n + 1) / INR (2 * n + 1) >= 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
H:IZR (up (/ eps)) > / eps

(0 < up (/ eps))%Z
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
H:IZR (up (/ eps)) > / eps
H0:(0 < up (/ eps))%Z
/ eps < INR N
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
INR N < INR (2 * N + 1)
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
eps <> 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
x ^ (2 * n + 1) / INR (2 * n + 1) >= 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
H:IZR (up (/ eps)) > / eps

0 < IZR (up (/ eps))
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
H:IZR (up (/ eps)) > / eps
H0:(0 < up (/ eps))%Z
/ eps < INR N
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
INR N < INR (2 * N + 1)
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
eps <> 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
x ^ (2 * n + 1) / INR (2 * n + 1) >= 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
H:IZR (up (/ eps)) > / eps

0 < / eps
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
H:IZR (up (/ eps)) > / eps
H0:(0 < up (/ eps))%Z
/ eps < INR N
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
INR N < INR (2 * N + 1)
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
eps <> 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
x ^ (2 * n + 1) / INR (2 * n + 1) >= 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
H:IZR (up (/ eps)) > / eps

0 < eps
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
H:IZR (up (/ eps)) > / eps
H0:(0 < up (/ eps))%Z
/ eps < INR N
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
INR N < INR (2 * N + 1)
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
eps <> 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
x ^ (2 * n + 1) / INR (2 * n + 1) >= 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
H:IZR (up (/ eps)) > / eps
H0:(0 < up (/ eps))%Z

/ eps < INR N
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
INR N < INR (2 * N + 1)
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
eps <> 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
x ^ (2 * n + 1) / INR (2 * n + 1) >= 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
H:IZR (up (/ eps)) > / eps
H0:(0 < up (/ eps))%Z

/ eps < INR (Pos.to_nat p)
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
INR N < INR (2 * N + 1)
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
eps <> 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
x ^ (2 * n + 1) / INR (2 * n + 1) >= 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
H:IZR (up (/ eps)) > / eps
H0:(0 < up (/ eps))%Z

/ eps < IZR (Z.pos p)
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
INR N < INR (2 * N + 1)
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
eps <> 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
x ^ (2 * n + 1) / INR (2 * n + 1) >= 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat

INR N < INR (2 * N + 1)
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
eps <> 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
x ^ (2 * n + 1) / INR (2 * n + 1) >= 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat

(N < 2 * N + 1)%nat
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
eps <> 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
x ^ (2 * n + 1) / INR (2 * n + 1) >= 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat

eps <> 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
x ^ (2 * n + 1) / INR (2 * n + 1) >= 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat

eps > 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
x ^ (2 * n + 1) / INR (2 * n + 1) >= 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat

x ^ (2 * n + 1) / INR (2 * n + 1) >= 0
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat

0 <= x ^ (2 * n + 1) / INR (2 * n + 1)
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat

0 <= x ^ (2 * n + 1) * / INR (2 * n + 1)
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat

0 <= x ^ (2 * n + 1)
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
0 <= / INR (2 * n + 1)
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat

0 <= x
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat
0 <= / INR (2 * n + 1)
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat

0 <= / INR (2 * n + 1)
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat

0 < / INR (2 * n + 1)
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat

0 < INR (2 * n + 1)
x:R
Hx:0 <= x <= 1
eps:R
Heps:eps > 0
p:positive
HN:IZR (Z.pos p) > / eps
N:=Pos.to_nat p:nat
n:nat
Hn:(n >= N)%nat

(0 < 2 * n + 1)%nat
omega. Qed.
x:R
Hx:0 <= x <= 1

{l : R | Un_cv (fun N : nat => sum_f_R0 (tg_alt (Ratan_seq x)) N) l}
x:R
Hx:0 <= x <= 1

{l : R | Un_cv (fun N : nat => sum_f_R0 (tg_alt (Ratan_seq x)) N) l}
exact (alternated_series (Ratan_seq x) (Ratan_seq_decreasing _ Hx) (Ratan_seq_converging _ Hx)). Defined.

forall (x : R) (n : nat), Ratan_seq (- x) n = - Ratan_seq x n

forall (x : R) (n : nat), Ratan_seq (- x) n = - Ratan_seq x n
x:R
n:nat

(- x) ^ (2 * n + 1) / INR (2 * n + 1) = - (x ^ (2 * n + 1) / INR (2 * n + 1))
x:R
n:nat

((- x) ^ 2) ^ n * - x / INR (2 * n + 1) = - ((x ^ 2) ^ n * x / INR (2 * n + 1))
unfold Rdiv; replace ((-x) ^ 2) with (x ^ 2) by ring; ring. Qed.

forall (x : R) (n : nat), sum_f_R0 (tg_alt (Ratan_seq (- x))) n = - sum_f_R0 (tg_alt (Ratan_seq x)) n

forall (x : R) (n : nat), sum_f_R0 (tg_alt (Ratan_seq (- x))) n = - sum_f_R0 (tg_alt (Ratan_seq x)) n
x:R
n:nat

sum_f_R0 (tg_alt (Ratan_seq (- x))) n = -1 * sum_f_R0 (tg_alt (Ratan_seq x)) n
x:R
n, i:nat

(-1) ^ i * Ratan_seq (- x) i = (-1) ^ i * Ratan_seq x i * -1
rewrite Ratan_seq_opp; ring. Qed.
x:R
Hx:-1 <= x <= 1

{l : R | Un_cv (fun N : nat => sum_f_R0 (tg_alt (Ratan_seq x)) N) l}
x:R
Hx:-1 <= x <= 1

{l : R | Un_cv (fun N : nat => sum_f_R0 (tg_alt (Ratan_seq x)) N) l}
x:R
Hx:-1 <= x <= 1
r:0 <= x

{l : R | Un_cv (fun N : nat => sum_f_R0 (tg_alt (Ratan_seq x)) N) l}
x:R
Hx:-1 <= x <= 1
r:x < 0
{l : R | Un_cv (fun N : nat => sum_f_R0 (tg_alt (Ratan_seq x)) N) l}
x:R
Hx:-1 <= x <= 1
r:0 <= x
pr:0 <= x <= 1

{l : R | Un_cv (fun N : nat => sum_f_R0 (tg_alt (Ratan_seq x)) N) l}
x:R
Hx:-1 <= x <= 1
r:x < 0
{l : R | Un_cv (fun N : nat => sum_f_R0 (tg_alt (Ratan_seq x)) N) l}
x:R
Hx:-1 <= x <= 1
r:x < 0

{l : R | Un_cv (fun N : nat => sum_f_R0 (tg_alt (Ratan_seq x)) N) l}
x:R
Hx:-1 <= x <= 1
r:x < 0
pr:0 <= - x <= 1

{l : R | Un_cv (fun N : nat => sum_f_R0 (tg_alt (Ratan_seq x)) N) l}
x:R
Hx:-1 <= x <= 1
r:x < 0
pr:0 <= - x <= 1
v:R
Pv:Un_cv (fun N : nat => sum_f_R0 (tg_alt (Ratan_seq (- x))) N) v

{l : R | Un_cv (fun N : nat => sum_f_R0 (tg_alt (Ratan_seq x)) N) l}
x:R
Hx:-1 <= x <= 1
r:x < 0
pr:0 <= - x <= 1
v:R
Pv:Un_cv (fun N : nat => sum_f_R0 (tg_alt (Ratan_seq (- x))) N) v

Un_cv (fun N : nat => sum_f_R0 (tg_alt (Ratan_seq x)) N) (- v)
x:R
Hx:-1 <= x <= 1
r:x < 0
pr:0 <= - x <= 1
v:R
Pv:Un_cv (fun N : nat => sum_f_R0 (tg_alt (Ratan_seq (- x))) N) v

forall n : nat, -1 * sum_f_R0 (tg_alt (Ratan_seq (- x))) n = sum_f_R0 (tg_alt (Ratan_seq x)) n
x:R
Hx:-1 <= x <= 1
r:x < 0
pr:0 <= - x <= 1
v:R
Pv:Un_cv (fun N : nat => sum_f_R0 (tg_alt (Ratan_seq (- x))) N) v
Un_cv (fun n : nat => -1 * sum_f_R0 (tg_alt (Ratan_seq (- x))) n) (- v)
x:R
Hx:-1 <= x <= 1
r:x < 0
pr:0 <= - x <= 1
v:R
Pv:Un_cv (fun N : nat => sum_f_R0 (tg_alt (Ratan_seq (- x))) N) v

Un_cv (fun n : nat => -1 * sum_f_R0 (tg_alt (Ratan_seq (- x))) n) (- v)
x:R
Hx:-1 <= x <= 1
r:x < 0
pr:0 <= - x <= 1
v:R
Pv:Un_cv (fun N : nat => sum_f_R0 (tg_alt (Ratan_seq (- x))) N) v

Un_cv (fun n : nat => -1 * sum_f_R0 (tg_alt (Ratan_seq (- x))) n) (-1 * v)
x:R
Hx:-1 <= x <= 1
r:x < 0
pr:0 <= - x <= 1
v:R
Pv:Un_cv (fun N : nat => sum_f_R0 (tg_alt (Ratan_seq (- x))) N) v

Un_cv (fun _ : nat => -1) (-1)
solve[intros; exists 0%nat; intros; rewrite R_dist_eq; auto]. Qed.
x:R

{-1 <= x <= 1} + {~ -1 <= x <= 1}
x:R

{-1 <= x <= 1} + {~ -1 <= x <= 1}
x:R
r:x <= 1

{-1 <= x <= 1} + {~ -1 <= x <= 1}
x:R
r:1 < x
{-1 <= x <= 1} + {~ -1 <= x <= 1}
x:R
r:x <= 1
r0:-1 <= x

{-1 <= x <= 1} + {~ -1 <= x <= 1}
x:R
r:x <= 1
r0:x < -1
{-1 <= x <= 1} + {~ -1 <= x <= 1}
x:R
r:1 < x
{-1 <= x <= 1} + {~ -1 <= x <= 1}
x:R
r:x <= 1
r0:x < -1

{-1 <= x <= 1} + {~ -1 <= x <= 1}
x:R
r:1 < x
{-1 <= x <= 1} + {~ -1 <= x <= 1}
x:R
r:1 < x

{-1 <= x <= 1} + {~ -1 <= x <= 1}
right;intros [a1 a2]; lra. Qed. Definition ps_atan (x : R) : R := match in_int x with left h => let (v, _) := ps_atan_exists_1 x h in v | right h => atan x end.

Proof of the equivalence of the two definitions between -1 and 1


ps_atan 0 = 0

ps_atan 0 = 0

match in_int 0 with | left h => let (v, _) := ps_atan_exists_1 0 h in v | right _ => atan 0 end = 0
h1:-1 <= 0 <= 1

(let (v, _) := ps_atan_exists_1 0 h1 in v) = 0
h2:~ -1 <= 0 <= 1
atan 0 = 0
h1:-1 <= 0 <= 1
v:R
P:Un_cv (fun N : nat => sum_f_R0 (tg_alt (Ratan_seq 0)) N) v

v = 0
h2:~ -1 <= 0 <= 1
atan 0 = 0
h1:-1 <= 0 <= 1
v:R
P:Un_cv (fun N : nat => sum_f_R0 (tg_alt (Ratan_seq 0)) N) v

Un_cv (fun N : nat => sum_f_R0 (tg_alt (Ratan_seq 0)) N) 0
h2:~ -1 <= 0 <= 1
atan 0 = 0
h1:-1 <= 0 <= 1
v:R
P:Un_cv (fun N : nat => sum_f_R0 (tg_alt (Ratan_seq 0)) N) v

forall n : nat, 0 = sum_f_R0 (tg_alt (Ratan_seq 0)) n
h1:-1 <= 0 <= 1
v:R
P:Un_cv (fun N : nat => sum_f_R0 (tg_alt (Ratan_seq 0)) N) v
Un_cv (fun _ : nat => 0) 0
h2:~ -1 <= 0 <= 1
atan 0 = 0
h1:-1 <= 0 <= 1
v:R
P:Un_cv (fun N : nat => sum_f_R0 (tg_alt (Ratan_seq 0)) N) v
n:nat

forall n0 : nat, (n0 <= n)%nat -> tg_alt (Ratan_seq 0) n0 = 0
h1:-1 <= 0 <= 1
v:R
P:Un_cv (fun N : nat => sum_f_R0 (tg_alt (Ratan_seq 0)) N) v
Un_cv (fun _ : nat => 0) 0
h2:~ -1 <= 0 <= 1
atan 0 = 0
h1:-1 <= 0 <= 1
v:R
P:Un_cv (fun N : nat => sum_f_R0 (tg_alt (Ratan_seq 0)) N) v
n, i:nat

(-1) ^ i * (0 * 0 ^ (i + (i + 0)) / match (i + (i + 0))%nat with | 0%nat => 1 | S _ => INR (i + (i + 0)) + 1 end) = 0
h1:-1 <= 0 <= 1
v:R
P:Un_cv (fun N : nat => sum_f_R0 (tg_alt (Ratan_seq 0)) N) v
Un_cv (fun _ : nat => 0) 0
h2:~ -1 <= 0 <= 1
atan 0 = 0
h1:-1 <= 0 <= 1
v:R
P:Un_cv (fun N : nat => sum_f_R0 (tg_alt (Ratan_seq 0)) N) v

Un_cv (fun _ : nat => 0) 0
h2:~ -1 <= 0 <= 1
atan 0 = 0
h1:-1 <= 0 <= 1
v:R
P:Un_cv (fun N : nat => sum_f_R0 (tg_alt (Ratan_seq 0)) N) v
eps:R
ep:eps > 0
n:nat

Rabs (0 - 0) < eps
h2:~ -1 <= 0 <= 1
atan 0 = 0
h2:~ -1 <= 0 <= 1

atan 0 = 0
case h2; split; lra. Qed.

forall (x : R) (h : -1 <= - x <= 1) (h' : -1 <= x <= 1), proj1_sig (ps_atan_exists_1 (- x) h) = - proj1_sig (ps_atan_exists_1 x h')

forall (x : R) (h : -1 <= - x <= 1) (h' : -1 <= x <= 1), proj1_sig (ps_atan_exists_1 (- x) h) = - proj1_sig (ps_atan_exists_1 x h')
x:R
h:-1 <= - x <= 1
h':-1 <= x <= 1
v:R
Pv:Un_cv (fun N : nat => sum_f_R0 (tg_alt (Ratan_seq (- x))) N) v

proj1_sig (exist (fun l : R => Un_cv (fun N : nat => sum_f_R0 (tg_alt (Ratan_seq (- x))) N) l) v Pv) = - proj1_sig (ps_atan_exists_1 x h')
x:R
h:-1 <= - x <= 1
h':-1 <= x <= 1
v:R
Pv:Un_cv (fun N : nat => sum_f_R0 (tg_alt (Ratan_seq (- x))) N) v
u:R
Pu:Un_cv (fun N : nat => sum_f_R0 (tg_alt (Ratan_seq x)) N) u

v = - u
x:R
h:-1 <= - x <= 1
h':-1 <= x <= 1
v:R
Pv:Un_cv (fun N : nat => sum_f_R0 (tg_alt (Ratan_seq (- x))) N) v
u:R
Pu:Un_cv (fun N : nat => sum_f_R0 (tg_alt (Ratan_seq x)) N) u

Un_cv (fun N : nat => -1 * sum_f_R0 (tg_alt (Ratan_seq x)) N) (-1 * u)
x:R
h:-1 <= - x <= 1
h':-1 <= x <= 1
v:R
Pv:Un_cv (fun N : nat => sum_f_R0 (tg_alt (Ratan_seq (- x))) N) v
u:R
Pu:Un_cv (fun N : nat => sum_f_R0 (tg_alt (Ratan_seq x)) N) u
Pu':Un_cv (fun N : nat => -1 * sum_f_R0 (tg_alt (Ratan_seq x)) N) (-1 * u)
v = - u
x:R
h:-1 <= - x <= 1
h':-1 <= x <= 1
v:R
Pv:Un_cv (fun N : nat => sum_f_R0 (tg_alt (Ratan_seq (- x))) N) v
u:R
Pu:Un_cv (fun N : nat => sum_f_R0 (tg_alt (Ratan_seq x)) N) u

Un_cv (fun _ : nat => -1) (-1)
x:R
h:-1 <= - x <= 1
h':-1 <= x <= 1
v:R
Pv:Un_cv (fun N : nat => sum_f_R0 (tg_alt (Ratan_seq (- x))) N) v
u:R
Pu:Un_cv (fun N : nat => sum_f_R0 (tg_alt (Ratan_seq x)) N) u
Pu':Un_cv (fun N : nat => -1 * sum_f_R0 (tg_alt (Ratan_seq x)) N) (-1 * u)
v = - u
x:R
h:-1 <= - x <= 1
h':-1 <= x <= 1
v:R
Pv:Un_cv (fun N : nat => sum_f_R0 (tg_alt (Ratan_seq (- x))) N) v
u:R
Pu:Un_cv (fun N : nat => sum_f_R0 (tg_alt (Ratan_seq x)) N) u
Pu':Un_cv (fun N : nat => -1 * sum_f_R0 (tg_alt (Ratan_seq x)) N) (-1 * u)

v = - u
x:R
h:-1 <= - x <= 1
h':-1 <= x <= 1
v:R
Pv:Un_cv (fun N : nat => sum_f_R0 (tg_alt (Ratan_seq (- x))) N) v
u:R
Pu:Un_cv (fun N : nat => sum_f_R0 (tg_alt (Ratan_seq x)) N) u
Pu':Un_cv (fun N : nat => -1 * sum_f_R0 (tg_alt (Ratan_seq x)) N) (-1 * u)

Un_cv (fun N : nat => -1 * sum_f_R0 (tg_alt (Ratan_seq x)) N) v
x:R
h:-1 <= - x <= 1
h':-1 <= x <= 1
v:R
Pv:Un_cv (fun N : nat => sum_f_R0 (tg_alt (Ratan_seq (- x))) N) v
u:R
Pu:Un_cv (fun N : nat => sum_f_R0 (tg_alt (Ratan_seq x)) N) u
Pu':Un_cv (fun N : nat => -1 * sum_f_R0 (tg_alt (Ratan_seq x)) N) (-1 * u)
Pv':Un_cv (fun N : nat => -1 * sum_f_R0 (tg_alt (Ratan_seq x)) N) v
v = - u
x:R
h:-1 <= - x <= 1
h':-1 <= x <= 1
v:R
Pv:Un_cv (fun N : nat => sum_f_R0 (tg_alt (Ratan_seq (- x))) N) v
u:R
Pu:Un_cv (fun N : nat => sum_f_R0 (tg_alt (Ratan_seq x)) N) u
Pu':Un_cv (fun N : nat => -1 * sum_f_R0 (tg_alt (Ratan_seq x)) N) (-1 * u)
Pv':Un_cv (fun N : nat => -1 * sum_f_R0 (tg_alt (Ratan_seq x)) N) v

v = - u
x:R
h:-1 <= - x <= 1
h':-1 <= x <= 1
v:R
Pv:Un_cv (fun N : nat => sum_f_R0 (tg_alt (Ratan_seq (- x))) N) v
u:R
Pu:Un_cv (fun N : nat => sum_f_R0 (tg_alt (Ratan_seq x)) N) u
Pu':Un_cv (fun N : nat => -1 * sum_f_R0 (tg_alt (Ratan_seq x)) N) (-1 * u)
Pv':Un_cv (fun N : nat => -1 * sum_f_R0 (tg_alt (Ratan_seq x)) N) v

v = -1 * u
apply UL_sequence with (1:=Pv') (2:= Pu'). Qed.

forall x : R, ps_atan (- x) = - ps_atan x

forall x : R, ps_atan (- x) = - ps_atan x
x:R

match in_int (- x) with | left h => let (v, _) := ps_atan_exists_1 (- x) h in v | right _ => atan (- x) end = - match in_int x with | left h => let (v, _) := ps_atan_exists_1 x h in v | right _ => atan x end
x:R
inside:-1 <= - x <= 1

(let (v, _) := ps_atan_exists_1 (- x) inside in v) = - match in_int x with | left h => let (v, _) := ps_atan_exists_1 x h in v | right _ => atan x end
x:R
outside:~ -1 <= - x <= 1
atan (- x) = - match in_int x with | left h => let (v, _) := ps_atan_exists_1 x h in v | right _ => atan x end
x:R
inside:-1 <= - x <= 1
ins':-1 <= x <= 1

(let (v, _) := ps_atan_exists_1 (- x) inside in v) = - (let (v, _) := ps_atan_exists_1 x ins' in v)
x:R
inside:-1 <= - x <= 1
outs':~ -1 <= x <= 1
(let (v, _) := ps_atan_exists_1 (- x) inside in v) = - atan x
x:R
outside:~ -1 <= - x <= 1
atan (- x) = - match in_int x with | left h => let (v, _) := ps_atan_exists_1 x h in v | right _ => atan x end
x:R
inside:-1 <= - x <= 1
ins':-1 <= x <= 1

proj1_sig (ps_atan_exists_1 (- x) inside) = - proj1_sig (ps_atan_exists_1 x ins') -> (let (v, _) := ps_atan_exists_1 (- x) inside in v) = - (let (v, _) := ps_atan_exists_1 x ins' in v)
x:R
inside:-1 <= - x <= 1
outs':~ -1 <= x <= 1
(let (v, _) := ps_atan_exists_1 (- x) inside in v) = - atan x
x:R
outside:~ -1 <= - x <= 1
atan (- x) = - match in_int x with | left h => let (v, _) := ps_atan_exists_1 x h in v | right _ => atan x end
x:R
inside:-1 <= - x <= 1
outs':~ -1 <= x <= 1

(let (v, _) := ps_atan_exists_1 (- x) inside in v) = - atan x
x:R
outside:~ -1 <= - x <= 1
atan (- x) = - match in_int x with | left h => let (v, _) := ps_atan_exists_1 x h in v | right _ => atan x end
x:R
outside:~ -1 <= - x <= 1

atan (- x) = - match in_int x with | left h => let (v, _) := ps_atan_exists_1 x h in v | right _ => atan x end
x:R
outside:~ -1 <= - x <= 1
ins':-1 <= x <= 1

atan (- x) = - (let (v, _) := ps_atan_exists_1 x ins' in v)
x:R
outside:~ -1 <= - x <= 1
outs':~ -1 <= x <= 1
atan (- x) = - atan x
x:R
outside:~ -1 <= - x <= 1
outs':~ -1 <= x <= 1

atan (- x) = - atan x
apply atan_opp. Qed.
atan = ps_atan

forall (N : nat) (x : R), 0 <= x -> x <= 1 -> continuity_pt (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x

forall (N : nat) (x : R), 0 <= x -> x <= 1 -> continuity_pt (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x

forall (x : R) (N : nat), sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
Sublemma:forall (x : R) (N : nat), sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
forall (N : nat) (x : R), 0 <= x -> x <= 1 -> continuity_pt (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x
x:R
N:nat

sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
Sublemma:forall (x : R) (N : nat), sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
forall (N : nat) (x : R), 0 <= x -> x <= 1 -> continuity_pt (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x
x:R

sum_f_R0 (tg_alt (Ratan_seq x)) 0 = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) 0) (fun x0 : R => x0 ^ 2) x
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
sum_f_R0 (tg_alt (Ratan_seq x)) (S N) = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) (S N)) (fun x0 : R => x0 ^ 2) x
Sublemma:forall (x : R) (N : nat), sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
forall (N : nat) (x : R), 0 <= x -> x <= 1 -> continuity_pt (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x

sum_f_R0 (tg_alt (Ratan_seq x)) (S N) = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) (S N)) (fun x0 : R => x0 ^ 2) x
Sublemma:forall (x : R) (N : nat), sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
forall (N : nat) (x : R), 0 <= x -> x <= 1 -> continuity_pt (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x

sum_f_R0 (tg_alt (Ratan_seq x)) N + tg_alt (Ratan_seq x) (S N) = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) (S N)) (fun x0 : R => x0 ^ 2) x
Sublemma:forall (x : R) (N : nat), sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
forall (N : nat) (x : R), 0 <= x -> x <= 1 -> continuity_pt (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x

x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x + tg_alt (Ratan_seq x) (S N) = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) (S N)) (fun x0 : R => x0 ^ 2) x
Sublemma:forall (x : R) (N : nat), sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
forall (N : nat) (x : R), 0 <= x -> x <= 1 -> continuity_pt (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x

x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x + tg_alt (Ratan_seq x) (S N) = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N + (-1) ^ S N / INR (2 * S N + 1) * x0 ^ S N) (fun x0 : R => x0 ^ 2) x
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N + (-1) ^ S N / INR (2 * S N + 1) * x0 ^ S N) (fun x0 : R => x0 ^ 2) = comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) (S N)) (fun x0 : R => x0 ^ 2)
Sublemma:forall (x : R) (N : nat), sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
forall (N : nat) (x : R), 0 <= x -> x <= 1 -> continuity_pt (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x

x * sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * (x ^ 2) ^ n) N + tg_alt (Ratan_seq x) (S N) = x * (sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * (x ^ 2) ^ n) N + (-1) ^ S N / INR (2 * S N + 1) * (x ^ 2) ^ S N)
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N + (-1) ^ S N / INR (2 * S N + 1) * x0 ^ S N) (fun x0 : R => x0 ^ 2) = comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) (S N)) (fun x0 : R => x0 ^ 2)
Sublemma:forall (x : R) (N : nat), sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
forall (N : nat) (x : R), 0 <= x -> x <= 1 -> continuity_pt (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x

x * sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * (x ^ 2) ^ n) N + tg_alt (Ratan_seq x) (S N) = x * sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * (x ^ 2) ^ n) N + x * ((-1) ^ S N / INR (2 * S N + 1) * (x ^ 2) ^ S N)
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N + (-1) ^ S N / INR (2 * S N + 1) * x0 ^ S N) (fun x0 : R => x0 ^ 2) = comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) (S N)) (fun x0 : R => x0 ^ 2)
Sublemma:forall (x : R) (N : nat), sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
forall (N : nat) (x : R), 0 <= x -> x <= 1 -> continuity_pt (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x

tg_alt (Ratan_seq x) (S N) = x * ((-1) ^ S N / INR (2 * S N + 1) * (x ^ 2) ^ S N)
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N + (-1) ^ S N / INR (2 * S N + 1) * x0 ^ S N) (fun x0 : R => x0 ^ 2) = comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) (S N)) (fun x0 : R => x0 ^ 2)
Sublemma:forall (x : R) (N : nat), sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
forall (N : nat) (x : R), 0 <= x -> x <= 1 -> continuity_pt (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x

(-1) ^ S N * (x ^ (2 * S N + 1) / INR (2 * S N + 1)) = x * ((-1) ^ S N / INR (2 * S N + 1) * (x ^ 2) ^ S N)
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N + (-1) ^ S N / INR (2 * S N + 1) * x0 ^ S N) (fun x0 : R => x0 ^ 2) = comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) (S N)) (fun x0 : R => x0 ^ 2)
Sublemma:forall (x : R) (N : nat), sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
forall (N : nat) (x : R), 0 <= x -> x <= 1 -> continuity_pt (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x

(-1) ^ S N * (x ^ (2 * S N + 1) / INR (2 * S N + 1)) = x * ((-1) ^ S N / INR (2 * S N + 1)) * (x ^ 2) ^ S N
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N + (-1) ^ S N / INR (2 * S N + 1) * x0 ^ S N) (fun x0 : R => x0 ^ 2) = comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) (S N)) (fun x0 : R => x0 ^ 2)
Sublemma:forall (x : R) (N : nat), sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
forall (N : nat) (x : R), 0 <= x -> x <= 1 -> continuity_pt (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
Hyp:x = 0

(-1) ^ S N * (x ^ (2 * S N + 1) / INR (2 * S N + 1)) = x * ((-1) ^ S N / INR (2 * S N + 1)) * (x ^ 2) ^ S N
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
Hyp:x <> 0
(-1) ^ S N * (x ^ (2 * S N + 1) / INR (2 * S N + 1)) = x * ((-1) ^ S N / INR (2 * S N + 1)) * (x ^ 2) ^ S N
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N + (-1) ^ S N / INR (2 * S N + 1) * x0 ^ S N) (fun x0 : R => x0 ^ 2) = comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) (S N)) (fun x0 : R => x0 ^ 2)
Sublemma:forall (x : R) (N : nat), sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
forall (N : nat) (x : R), 0 <= x -> x <= 1 -> continuity_pt (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
Hyp:x = 0

(-1) ^ S N * (0 / INR (2 * S N + 1)) = 0 * ((-1) ^ S N / INR (2 * S N + 1)) * (0 ^ 2) ^ S N
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
Hyp:x = 0
(0 < 2 * S N + 1)%nat
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
Hyp:x <> 0
(-1) ^ S N * (x ^ (2 * S N + 1) / INR (2 * S N + 1)) = x * ((-1) ^ S N / INR (2 * S N + 1)) * (x ^ 2) ^ S N
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N + (-1) ^ S N / INR (2 * S N + 1) * x0 ^ S N) (fun x0 : R => x0 ^ 2) = comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) (S N)) (fun x0 : R => x0 ^ 2)
Sublemma:forall (x : R) (N : nat), sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
forall (N : nat) (x : R), 0 <= x -> x <= 1 -> continuity_pt (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
Hyp:x = 0

(-1) ^ S N * (0 / INR (2 * S N + 1)) = 0
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
Hyp:x = 0
(0 < 2 * S N + 1)%nat
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
Hyp:x <> 0
(-1) ^ S N * (x ^ (2 * S N + 1) / INR (2 * S N + 1)) = x * ((-1) ^ S N / INR (2 * S N + 1)) * (x ^ 2) ^ S N
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N + (-1) ^ S N / INR (2 * S N + 1) * x0 ^ S N) (fun x0 : R => x0 ^ 2) = comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) (S N)) (fun x0 : R => x0 ^ 2)
Sublemma:forall (x : R) (N : nat), sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
forall (N : nat) (x : R), 0 <= x -> x <= 1 -> continuity_pt (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
Hyp:x = 0

(0 < 2 * S N + 1)%nat
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
Hyp:x <> 0
(-1) ^ S N * (x ^ (2 * S N + 1) / INR (2 * S N + 1)) = x * ((-1) ^ S N / INR (2 * S N + 1)) * (x ^ 2) ^ S N
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N + (-1) ^ S N / INR (2 * S N + 1) * x0 ^ S N) (fun x0 : R => x0 ^ 2) = comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) (S N)) (fun x0 : R => x0 ^ 2)
Sublemma:forall (x : R) (N : nat), sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
forall (N : nat) (x : R), 0 <= x -> x <= 1 -> continuity_pt (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
Hyp:x <> 0

(-1) ^ S N * (x ^ (2 * S N + 1) / INR (2 * S N + 1)) = x * ((-1) ^ S N / INR (2 * S N + 1)) * (x ^ 2) ^ S N
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N + (-1) ^ S N / INR (2 * S N + 1) * x0 ^ S N) (fun x0 : R => x0 ^ 2) = comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) (S N)) (fun x0 : R => x0 ^ 2)
Sublemma:forall (x : R) (N : nat), sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
forall (N : nat) (x : R), 0 <= x -> x <= 1 -> continuity_pt (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
Hyp:x <> 0

(-1) ^ S N * (x ^ (2 * S N + 1) / INR (2 * S N + 1)) = x ^ (2 * S N + 1) * ((-1) ^ S N / INR (2 * S N + 1))
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
Hyp:x <> 0
x ^ (2 * S N + 1) * ((-1) ^ S N / INR (2 * S N + 1)) = x * ((-1) ^ S N / INR (2 * S N + 1)) * (x ^ 2) ^ S N
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N + (-1) ^ S N / INR (2 * S N + 1) * x0 ^ S N) (fun x0 : R => x0 ^ 2) = comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) (S N)) (fun x0 : R => x0 ^ 2)
Sublemma:forall (x : R) (N : nat), sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
forall (N : nat) (x : R), 0 <= x -> x <= 1 -> continuity_pt (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
Hyp:x <> 0

x ^ (2 * S N + 1) * / INR (2 * S N + 1) * (-1) ^ S N = x ^ (2 * S N + 1) * ((-1) ^ S N / INR (2 * S N + 1))
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
Hyp:x <> 0
x ^ (2 * S N + 1) * ((-1) ^ S N / INR (2 * S N + 1)) = x * ((-1) ^ S N / INR (2 * S N + 1)) * (x ^ 2) ^ S N
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N + (-1) ^ S N / INR (2 * S N + 1) * x0 ^ S N) (fun x0 : R => x0 ^ 2) = comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) (S N)) (fun x0 : R => x0 ^ 2)
Sublemma:forall (x : R) (N : nat), sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
forall (N : nat) (x : R), 0 <= x -> x <= 1 -> continuity_pt (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
Hyp:x <> 0

/ INR (2 * S N + 1) * (-1) ^ S N = (-1) ^ S N / INR (2 * S N + 1)
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
Hyp:x <> 0
x ^ (2 * S N + 1) * ((-1) ^ S N / INR (2 * S N + 1)) = x * ((-1) ^ S N / INR (2 * S N + 1)) * (x ^ 2) ^ S N
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N + (-1) ^ S N / INR (2 * S N + 1) * x0 ^ S N) (fun x0 : R => x0 ^ 2) = comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) (S N)) (fun x0 : R => x0 ^ 2)
Sublemma:forall (x : R) (N : nat), sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
forall (N : nat) (x : R), 0 <= x -> x <= 1 -> continuity_pt (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
Hyp:x <> 0

INR (2 * S N + 1) <> 0
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
Hyp:x <> 0
x ^ (2 * S N + 1) * ((-1) ^ S N / INR (2 * S N + 1)) = x * ((-1) ^ S N / INR (2 * S N + 1)) * (x ^ 2) ^ S N
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N + (-1) ^ S N / INR (2 * S N + 1) * x0 ^ S N) (fun x0 : R => x0 ^ 2) = comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) (S N)) (fun x0 : R => x0 ^ 2)
Sublemma:forall (x : R) (N : nat), sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
forall (N : nat) (x : R), 0 <= x -> x <= 1 -> continuity_pt (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
Hyp:x <> 0

x ^ (2 * S N + 1) * ((-1) ^ S N / INR (2 * S N + 1)) = x * ((-1) ^ S N / INR (2 * S N + 1)) * (x ^ 2) ^ S N
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N + (-1) ^ S N / INR (2 * S N + 1) * x0 ^ S N) (fun x0 : R => x0 ^ 2) = comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) (S N)) (fun x0 : R => x0 ^ 2)
Sublemma:forall (x : R) (N : nat), sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
forall (N : nat) (x : R), 0 <= x -> x <= 1 -> continuity_pt (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
Hyp:x <> 0

x ^ (2 * S N + 1) * ((-1) ^ S N / INR (2 * S N + 1)) = x * ((-1) ^ S N / INR (2 * S N + 1) * (x ^ 2) ^ S N)
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N + (-1) ^ S N / INR (2 * S N + 1) * x0 ^ S N) (fun x0 : R => x0 ^ 2) = comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) (S N)) (fun x0 : R => x0 ^ 2)
Sublemma:forall (x : R) (N : nat), sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
forall (N : nat) (x : R), 0 <= x -> x <= 1 -> continuity_pt (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
Hyp:x <> 0

x ^ (2 * S N + 1) * ((-1) ^ S N / INR (2 * S N + 1)) = (-1) ^ S N / INR (2 * S N + 1) * (x ^ 2) ^ S N * x
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
Hyp:x <> 0
(-1) ^ S N / INR (2 * S N + 1) * (x ^ 2) ^ S N * x = x * ((-1) ^ S N / INR (2 * S N + 1) * (x ^ 2) ^ S N)
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N + (-1) ^ S N / INR (2 * S N + 1) * x0 ^ S N) (fun x0 : R => x0 ^ 2) = comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) (S N)) (fun x0 : R => x0 ^ 2)
Sublemma:forall (x : R) (N : nat), sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
forall (N : nat) (x : R), 0 <= x -> x <= 1 -> continuity_pt (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
Hyp:x <> 0

x ^ (2 * S N + 1) * ((-1) ^ S N / INR (2 * S N + 1)) = (-1) ^ S N / INR (2 * S N + 1) * ((x ^ 2) ^ S N * x)
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
Hyp:x <> 0
(-1) ^ S N / INR (2 * S N + 1) * (x ^ 2) ^ S N * x = x * ((-1) ^ S N / INR (2 * S N + 1) * (x ^ 2) ^ S N)
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N + (-1) ^ S N / INR (2 * S N + 1) * x0 ^ S N) (fun x0 : R => x0 ^ 2) = comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) (S N)) (fun x0 : R => x0 ^ 2)
Sublemma:forall (x : R) (N : nat), sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
forall (N : nat) (x : R), 0 <= x -> x <= 1 -> continuity_pt (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
Hyp:x <> 0

x ^ (2 * S N + 1) * ((-1) ^ S N / INR (2 * S N + 1)) = (-1) ^ S N / INR (2 * S N + 1) * x ^ (2 * S N + 1)
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
Hyp:x <> 0
x ^ (2 * S N + 1) = (x ^ 2) ^ S N * x
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
Hyp:x <> 0
(-1) ^ S N / INR (2 * S N + 1) * (x ^ 2) ^ S N * x = x * ((-1) ^ S N / INR (2 * S N + 1) * (x ^ 2) ^ S N)
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N + (-1) ^ S N / INR (2 * S N + 1) * x0 ^ S N) (fun x0 : R => x0 ^ 2) = comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) (S N)) (fun x0 : R => x0 ^ 2)
Sublemma:forall (x : R) (N : nat), sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
forall (N : nat) (x : R), 0 <= x -> x <= 1 -> continuity_pt (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
Hyp:x <> 0

x ^ (2 * S N + 1) = (x ^ 2) ^ S N * x
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
Hyp:x <> 0
(-1) ^ S N / INR (2 * S N + 1) * (x ^ 2) ^ S N * x = x * ((-1) ^ S N / INR (2 * S N + 1) * (x ^ 2) ^ S N)
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N + (-1) ^ S N / INR (2 * S N + 1) * x0 ^ S N) (fun x0 : R => x0 ^ 2) = comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) (S N)) (fun x0 : R => x0 ^ 2)
Sublemma:forall (x : R) (N : nat), sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
forall (N : nat) (x : R), 0 <= x -> x <= 1 -> continuity_pt (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
Hyp:x <> 0

x ^ (2 * S N + 1) = x ^ (2 * S N) * x
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
Hyp:x <> 0
(-1) ^ S N / INR (2 * S N + 1) * (x ^ 2) ^ S N * x = x * ((-1) ^ S N / INR (2 * S N + 1) * (x ^ 2) ^ S N)
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N + (-1) ^ S N / INR (2 * S N + 1) * x0 ^ S N) (fun x0 : R => x0 ^ 2) = comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) (S N)) (fun x0 : R => x0 ^ 2)
Sublemma:forall (x : R) (N : nat), sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
forall (N : nat) (x : R), 0 <= x -> x <= 1 -> continuity_pt (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
Hyp:x <> 0

forall (x0 : R) (n : nat), x0 ^ n * x0 = x0 ^ (n + 1)
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
Hyp:x <> 0
Temp:forall (x0 : R) (n : nat), x0 ^ n * x0 = x0 ^ (n + 1)
x ^ (2 * S N + 1) = x ^ (2 * S N) * x
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
Hyp:x <> 0
(-1) ^ S N / INR (2 * S N + 1) * (x ^ 2) ^ S N * x = x * ((-1) ^ S N / INR (2 * S N + 1) * (x ^ 2) ^ S N)
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N + (-1) ^ S N / INR (2 * S N + 1) * x0 ^ S N) (fun x0 : R => x0 ^ 2) = comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) (S N)) (fun x0 : R => x0 ^ 2)
Sublemma:forall (x : R) (N : nat), sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
forall (N : nat) (x : R), 0 <= x -> x <= 1 -> continuity_pt (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
Hyp:x <> 0
a:R

a ^ 0 * a = a ^ (0 + 1)
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n0 : nat => (-1) ^ n0 / INR (2 * n0 + 1) * x0 ^ n0) N) (fun x0 : R => x0 ^ 2) x
Hyp:x <> 0
a:R
n:nat
IHn:a ^ n * a = a ^ (n + 1)
a ^ S n * a = a ^ (S n + 1)
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
Hyp:x <> 0
Temp:forall (x0 : R) (n : nat), x0 ^ n * x0 = x0 ^ (n + 1)
x ^ (2 * S N + 1) = x ^ (2 * S N) * x
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
Hyp:x <> 0
(-1) ^ S N / INR (2 * S N + 1) * (x ^ 2) ^ S N * x = x * ((-1) ^ S N / INR (2 * S N + 1) * (x ^ 2) ^ S N)
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N + (-1) ^ S N / INR (2 * S N + 1) * x0 ^ S N) (fun x0 : R => x0 ^ 2) = comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) (S N)) (fun x0 : R => x0 ^ 2)
Sublemma:forall (x : R) (N : nat), sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
forall (N : nat) (x : R), 0 <= x -> x <= 1 -> continuity_pt (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
Hyp:x <> 0
a:R

1 * a = a ^ (0 + 1)
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n0 : nat => (-1) ^ n0 / INR (2 * n0 + 1) * x0 ^ n0) N) (fun x0 : R => x0 ^ 2) x
Hyp:x <> 0
a:R
n:nat
IHn:a ^ n * a = a ^ (n + 1)
a ^ S n * a = a ^ (S n + 1)
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
Hyp:x <> 0
Temp:forall (x0 : R) (n : nat), x0 ^ n * x0 = x0 ^ (n + 1)
x ^ (2 * S N + 1) = x ^ (2 * S N) * x
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
Hyp:x <> 0
(-1) ^ S N / INR (2 * S N + 1) * (x ^ 2) ^ S N * x = x * ((-1) ^ S N / INR (2 * S N + 1) * (x ^ 2) ^ S N)
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N + (-1) ^ S N / INR (2 * S N + 1) * x0 ^ S N) (fun x0 : R => x0 ^ 2) = comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) (S N)) (fun x0 : R => x0 ^ 2)
Sublemma:forall (x : R) (N : nat), sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
forall (N : nat) (x : R), 0 <= x -> x <= 1 -> continuity_pt (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n0 : nat => (-1) ^ n0 / INR (2 * n0 + 1) * x0 ^ n0) N) (fun x0 : R => x0 ^ 2) x
Hyp:x <> 0
a:R
n:nat
IHn:a ^ n * a = a ^ (n + 1)

a ^ S n * a = a ^ (S n + 1)
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
Hyp:x <> 0
Temp:forall (x0 : R) (n : nat), x0 ^ n * x0 = x0 ^ (n + 1)
x ^ (2 * S N + 1) = x ^ (2 * S N) * x
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
Hyp:x <> 0
(-1) ^ S N / INR (2 * S N + 1) * (x ^ 2) ^ S N * x = x * ((-1) ^ S N / INR (2 * S N + 1) * (x ^ 2) ^ S N)
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N + (-1) ^ S N / INR (2 * S N + 1) * x0 ^ S N) (fun x0 : R => x0 ^ 2) = comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) (S N)) (fun x0 : R => x0 ^ 2)
Sublemma:forall (x : R) (N : nat), sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
forall (N : nat) (x : R), 0 <= x -> x <= 1 -> continuity_pt (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
Hyp:x <> 0
Temp:forall (x0 : R) (n : nat), x0 ^ n * x0 = x0 ^ (n + 1)

x ^ (2 * S N + 1) = x ^ (2 * S N) * x
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
Hyp:x <> 0
(-1) ^ S N / INR (2 * S N + 1) * (x ^ 2) ^ S N * x = x * ((-1) ^ S N / INR (2 * S N + 1) * (x ^ 2) ^ S N)
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N + (-1) ^ S N / INR (2 * S N + 1) * x0 ^ S N) (fun x0 : R => x0 ^ 2) = comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) (S N)) (fun x0 : R => x0 ^ 2)
Sublemma:forall (x : R) (N : nat), sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
forall (N : nat) (x : R), 0 <= x -> x <= 1 -> continuity_pt (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
Hyp:x <> 0

(-1) ^ S N / INR (2 * S N + 1) * (x ^ 2) ^ S N * x = x * ((-1) ^ S N / INR (2 * S N + 1) * (x ^ 2) ^ S N)
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N + (-1) ^ S N / INR (2 * S N + 1) * x0 ^ S N) (fun x0 : R => x0 ^ 2) = comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) (S N)) (fun x0 : R => x0 ^ 2)
Sublemma:forall (x : R) (N : nat), sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
forall (N : nat) (x : R), 0 <= x -> x <= 1 -> continuity_pt (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x
x:R
N:nat
IHN:sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x

comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) N + (-1) ^ S N / INR (2 * S N + 1) * x0 ^ S N) (fun x0 : R => x0 ^ 2) = comp (fun x0 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x0 ^ n) (S N)) (fun x0 : R => x0 ^ 2)
Sublemma:forall (x : R) (N : nat), sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x
forall (N : nat) (x : R), 0 <= x -> x <= 1 -> continuity_pt (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x
Sublemma:forall (x : R) (N : nat), sum_f_R0 (tg_alt (Ratan_seq x)) N = x * comp (fun x0 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x0 ^ n) N) (fun x0 : R => x0 ^ 2) x

forall (N : nat) (x : R), 0 <= x -> x <= 1 -> continuity_pt (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x
Sublemma:forall (x0 : R) (N0 : nat), sum_f_R0 (tg_alt (Ratan_seq x0)) N0 = x0 * comp (fun x1 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x1 ^ n) N0) (fun x1 : R => x1 ^ 2) x0
N:nat
x:R
x_lb:0 <= x
x_ub:x <= 1

continuity_pt (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x
Sublemma:forall (x0 : R) (N0 : nat), sum_f_R0 (tg_alt (Ratan_seq x0)) N0 = x0 * comp (fun x1 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x1 ^ n) N0) (fun x1 : R => x1 ^ 2) x0
N:nat
x:R
x_lb:0 <= x
x_ub:x <= 1
eps:R
eps_pos:eps > 0

exists alp : R, alp > 0 /\ (forall x0 : Base R_met, D_x no_cond x x0 /\ dist R_met x0 x < alp -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq x)) N) < eps)
Sublemma:forall (x0 : R) (N0 : nat), sum_f_R0 (tg_alt (Ratan_seq x0)) N0 = x0 * comp (fun x1 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x1 ^ n) N0) (fun x1 : R => x1 ^ 2) x0
N:nat
x:R
x_lb:0 <= x
x_ub:x <= 1
eps:R
eps_pos:eps > 0

continuity id
Sublemma:forall (x0 : R) (N0 : nat), sum_f_R0 (tg_alt (Ratan_seq x0)) N0 = x0 * comp (fun x1 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x1 ^ n) N0) (fun x1 : R => x1 ^ 2) x0
N:nat
x:R
x_lb:0 <= x
x_ub:x <= 1
eps:R
eps_pos:eps > 0
continuity_id:continuity id
exists alp : R, alp > 0 /\ (forall x0 : Base R_met, D_x no_cond x x0 /\ dist R_met x0 x < alp -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq x)) N) < eps)
Sublemma:forall (x0 : R) (N0 : nat), sum_f_R0 (tg_alt (Ratan_seq x0)) N0 = x0 * comp (fun x1 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x1 ^ n) N0) (fun x1 : R => x1 ^ 2) x0
N:nat
x:R
x_lb:0 <= x
x_ub:x <= 1
eps:R
eps_pos:eps > 0
continuity_id:continuity id

exists alp : R, alp > 0 /\ (forall x0 : Base R_met, D_x no_cond x x0 /\ dist R_met x0 x < alp -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq x)) N) < eps)
Sublemma:forall (x0 : R) (N0 : nat), sum_f_R0 (tg_alt (Ratan_seq x0)) N0 = x0 * comp (fun x1 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x1 ^ n) N0) (fun x1 : R => x1 ^ 2) x0
N:nat
x:R
x_lb:0 <= x
x_ub:x <= 1
eps:R
eps_pos:eps > 0
continuity_id:continuity id
Temp:continuity (comp (fun x1 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x1 ^ n) N) (fun x1 : R => x1 ^ 2)) -> continuity (id * comp (fun x1 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x1 ^ n) N) (fun x1 : R => x1 ^ 2))

exists alp : R, alp > 0 /\ (forall x0 : Base R_met, D_x no_cond x x0 /\ dist R_met x0 x < alp -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq x)) N) < eps)
Sublemma:forall (x0 : R) (N0 : nat), sum_f_R0 (tg_alt (Ratan_seq x0)) N0 = x0 * comp (fun x1 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x1 ^ n) N0) (fun x1 : R => x1 ^ 2) x0
N:nat
x:R
x_lb:0 <= x
x_ub:x <= 1
eps:R
eps_pos:eps > 0
continuity_id:continuity id
Temp:continuity (comp (fun x1 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x1 ^ n) N) (fun x1 : R => x1 ^ 2)) -> continuity (id * comp (fun x1 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x1 ^ n) N) (fun x1 : R => x1 ^ 2))

continuity (comp (fun x1 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x1 ^ n) N) (fun x1 : R => x1 ^ 2))
Sublemma:forall (x0 : R) (N0 : nat), sum_f_R0 (tg_alt (Ratan_seq x0)) N0 = x0 * comp (fun x1 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x1 ^ n) N0) (fun x1 : R => x1 ^ 2) x0
N:nat
x:R
x_lb:0 <= x
x_ub:x <= 1
eps:R
eps_pos:eps > 0
continuity_id:continuity id
Temp:continuity (comp (fun x1 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x1 ^ n) N) (fun x1 : R => x1 ^ 2)) -> continuity (id * comp (fun x1 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x1 ^ n) N) (fun x1 : R => x1 ^ 2))
Temp2:continuity (comp (fun x1 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x1 ^ n) N) (fun x1 : R => x1 ^ 2))
exists alp : R, alp > 0 /\ (forall x0 : Base R_met, D_x no_cond x x0 /\ dist R_met x0 x < alp -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq x)) N) < eps)
Sublemma:forall (x0 : R) (N0 : nat), sum_f_R0 (tg_alt (Ratan_seq x0)) N0 = x0 * comp (fun x1 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x1 ^ n) N0) (fun x1 : R => x1 ^ 2) x0
N:nat
x:R
x_lb:0 <= x
x_ub:x <= 1
eps:R
eps_pos:eps > 0
continuity_id:continuity id
Temp:continuity (comp (fun x1 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x1 ^ n) N) (fun x1 : R => x1 ^ 2)) -> continuity (id * comp (fun x1 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x1 ^ n) N) (fun x1 : R => x1 ^ 2))

continuity (fun x1 : R => x1 ^ 2)
Sublemma:forall (x0 : R) (N0 : nat), sum_f_R0 (tg_alt (Ratan_seq x0)) N0 = x0 * comp (fun x1 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x1 ^ n) N0) (fun x1 : R => x1 ^ 2) x0
N:nat
x:R
x_lb:0 <= x
x_ub:x <= 1
eps:R
eps_pos:eps > 0
continuity_id:continuity id
Temp:continuity (comp (fun x1 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x1 ^ n) N) (fun x1 : R => x1 ^ 2)) -> continuity (id * comp (fun x1 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x1 ^ n) N) (fun x1 : R => x1 ^ 2))
continuity (fun x1 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x1 ^ n) N)
Sublemma:forall (x0 : R) (N0 : nat), sum_f_R0 (tg_alt (Ratan_seq x0)) N0 = x0 * comp (fun x1 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x1 ^ n) N0) (fun x1 : R => x1 ^ 2) x0
N:nat
x:R
x_lb:0 <= x
x_ub:x <= 1
eps:R
eps_pos:eps > 0
continuity_id:continuity id
Temp:continuity (comp (fun x1 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x1 ^ n) N) (fun x1 : R => x1 ^ 2)) -> continuity (id * comp (fun x1 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x1 ^ n) N) (fun x1 : R => x1 ^ 2))
Temp2:continuity (comp (fun x1 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x1 ^ n) N) (fun x1 : R => x1 ^ 2))
exists alp : R, alp > 0 /\ (forall x0 : Base R_met, D_x no_cond x x0 /\ dist R_met x0 x < alp -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq x)) N) < eps)
Sublemma:forall (x0 : R) (N0 : nat), sum_f_R0 (tg_alt (Ratan_seq x0)) N0 = x0 * comp (fun x1 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x1 ^ n) N0) (fun x1 : R => x1 ^ 2) x0
N:nat
x:R
x_lb:0 <= x
x_ub:x <= 1
eps:R
eps_pos:eps > 0
continuity_id:continuity id
Temp:continuity (comp (fun x1 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x1 ^ n) N) (fun x1 : R => x1 ^ 2)) -> continuity (id * comp (fun x1 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x1 ^ n) N) (fun x1 : R => x1 ^ 2))

continuity (fun x1 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x1 ^ n) N)
Sublemma:forall (x0 : R) (N0 : nat), sum_f_R0 (tg_alt (Ratan_seq x0)) N0 = x0 * comp (fun x1 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x1 ^ n) N0) (fun x1 : R => x1 ^ 2) x0
N:nat
x:R
x_lb:0 <= x
x_ub:x <= 1
eps:R
eps_pos:eps > 0
continuity_id:continuity id
Temp:continuity (comp (fun x1 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x1 ^ n) N) (fun x1 : R => x1 ^ 2)) -> continuity (id * comp (fun x1 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x1 ^ n) N) (fun x1 : R => x1 ^ 2))
Temp2:continuity (comp (fun x1 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x1 ^ n) N) (fun x1 : R => x1 ^ 2))
exists alp : R, alp > 0 /\ (forall x0 : Base R_met, D_x no_cond x x0 /\ dist R_met x0 x < alp -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq x)) N) < eps)
Sublemma:forall (x0 : R) (N0 : nat), sum_f_R0 (tg_alt (Ratan_seq x0)) N0 = x0 * comp (fun x1 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x1 ^ n) N0) (fun x1 : R => x1 ^ 2) x0
N:nat
x:R
x_lb:0 <= x
x_ub:x <= 1
eps:R
eps_pos:eps > 0
continuity_id:continuity id
Temp:continuity (comp (fun x1 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x1 ^ n) N) (fun x1 : R => x1 ^ 2)) -> continuity (id * comp (fun x1 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x1 ^ n) N) (fun x1 : R => x1 ^ 2))
Temp2:continuity (comp (fun x1 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x1 ^ n) N) (fun x1 : R => x1 ^ 2))

exists alp : R, alp > 0 /\ (forall x0 : Base R_met, D_x no_cond x x0 /\ dist R_met x0 x < alp -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq x)) N) < eps)
Sublemma:forall (x0 : R) (N0 : nat), sum_f_R0 (tg_alt (Ratan_seq x0)) N0 = x0 * comp (fun x1 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x1 ^ n) N0) (fun x1 : R => x1 ^ 2) x0
N:nat
x:R
x_lb:0 <= x
x_ub:x <= 1
eps:R
eps_pos:eps > 0
continuity_id:continuity id
alpha:R
alpha_pos:alpha > 0
T:forall x0 : Base R_met, D_x no_cond x x0 /\ dist R_met x0 x < alpha -> dist R_met ((id * comp (fun x1 : R => sum_f_R0 (fun n : nat => ((-1) ^ n / INR (2 * n + 1) * x1 ^ n)%R) N) (fun x1 : R => x1 ^ 2))%F x0) ((id * comp (fun x1 : R => sum_f_R0 (fun n : nat => ((-1) ^ n / INR (2 * n + 1) * x1 ^ n)%R) N) (fun x1 : R => x1 ^ 2))%F x) < eps

exists alp : R, alp > 0 /\ (forall x0 : Base R_met, D_x no_cond x x0 /\ dist R_met x0 x < alp -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq x)) N) < eps)
Sublemma:forall (x0 : R) (N0 : nat), sum_f_R0 (tg_alt (Ratan_seq x0)) N0 = x0 * comp (fun x1 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x1 ^ n) N0) (fun x1 : R => x1 ^ 2) x0
N:nat
x:R
x_lb:0 <= x
x_ub:x <= 1
eps:R
eps_pos:eps > 0
continuity_id:continuity id
alpha:R
alpha_pos:alpha > 0
T:forall x0 : Base R_met, D_x no_cond x x0 /\ dist R_met x0 x < alpha -> dist R_met ((id * comp (fun x1 : R => sum_f_R0 (fun n : nat => ((-1) ^ n / INR (2 * n + 1) * x1 ^ n)%R) N) (fun x1 : R => x1 ^ 2))%F x0) ((id * comp (fun x1 : R => sum_f_R0 (fun n : nat => ((-1) ^ n / INR (2 * n + 1) * x1 ^ n)%R) N) (fun x1 : R => x1 ^ 2))%F x) < eps

alpha > 0
Sublemma:forall (x0 : R) (N0 : nat), sum_f_R0 (tg_alt (Ratan_seq x0)) N0 = x0 * comp (fun x1 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x1 ^ n) N0) (fun x1 : R => x1 ^ 2) x0
N:nat
x:R
x_lb:0 <= x
x_ub:x <= 1
eps:R
eps_pos:eps > 0
continuity_id:continuity id
alpha:R
alpha_pos:alpha > 0
T:forall x0 : Base R_met, D_x no_cond x x0 /\ dist R_met x0 x < alpha -> dist R_met ((id * comp (fun x1 : R => sum_f_R0 (fun n : nat => ((-1) ^ n / INR (2 * n + 1) * x1 ^ n)%R) N) (fun x1 : R => x1 ^ 2))%F x0) ((id * comp (fun x1 : R => sum_f_R0 (fun n : nat => ((-1) ^ n / INR (2 * n + 1) * x1 ^ n)%R) N) (fun x1 : R => x1 ^ 2))%F x) < eps
forall x0 : Base R_met, D_x no_cond x x0 /\ dist R_met x0 x < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq x)) N) < eps
Sublemma:forall (x0 : R) (N0 : nat), sum_f_R0 (tg_alt (Ratan_seq x0)) N0 = x0 * comp (fun x1 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x1 ^ n) N0) (fun x1 : R => x1 ^ 2) x0
N:nat
x:R
x_lb:0 <= x
x_ub:x <= 1
eps:R
eps_pos:eps > 0
continuity_id:continuity id
alpha:R
alpha_pos:alpha > 0
T:forall x0 : Base R_met, D_x no_cond x x0 /\ dist R_met x0 x < alpha -> dist R_met ((id * comp (fun x1 : R => sum_f_R0 (fun n : nat => ((-1) ^ n / INR (2 * n + 1) * x1 ^ n)%R) N) (fun x1 : R => x1 ^ 2))%F x0) ((id * comp (fun x1 : R => sum_f_R0 (fun n : nat => ((-1) ^ n / INR (2 * n + 1) * x1 ^ n)%R) N) (fun x1 : R => x1 ^ 2))%F x) < eps

forall x0 : Base R_met, D_x no_cond x x0 /\ dist R_met x0 x < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq x)) N) < eps
Sublemma:forall (x1 : R) (N0 : nat), sum_f_R0 (tg_alt (Ratan_seq x1)) N0 = x1 * comp (fun x2 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x2 ^ n) N0) (fun x2 : R => x2 ^ 2) x1
N:nat
x:R
x_lb:0 <= x
x_ub:x <= 1
eps:R
eps_pos:eps > 0
continuity_id:continuity id
alpha:R
alpha_pos:alpha > 0
T:forall x1 : Base R_met, D_x no_cond x x1 /\ dist R_met x1 x < alpha -> dist R_met ((id * comp (fun x2 : R => sum_f_R0 (fun n : nat => ((-1) ^ n / INR (2 * n + 1) * x2 ^ n)%R) N) (fun x2 : R => x2 ^ 2))%F x1) ((id * comp (fun x2 : R => sum_f_R0 (fun n : nat => ((-1) ^ n / INR (2 * n + 1) * x2 ^ n)%R) N) (fun x2 : R => x2 ^ 2))%F x) < eps
x0:Base R_met
x0_cond:D_x no_cond x x0 /\ dist R_met x0 x < alpha

dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq x)) N) < eps
Sublemma:forall (x1 : R) (N0 : nat), sum_f_R0 (tg_alt (Ratan_seq x1)) N0 = x1 * comp (fun x2 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x2 ^ n) N0) (fun x2 : R => x2 ^ 2) x1
N:nat
x:R
x_lb:0 <= x
x_ub:x <= 1
eps:R
eps_pos:eps > 0
continuity_id:continuity id
alpha:R
alpha_pos:alpha > 0
T:forall x1 : Base R_met, D_x no_cond x x1 /\ dist R_met x1 x < alpha -> dist R_met ((id * comp (fun x2 : R => sum_f_R0 (fun n : nat => ((-1) ^ n / INR (2 * n + 1) * x2 ^ n)%R) N) (fun x2 : R => x2 ^ 2))%F x1) ((id * comp (fun x2 : R => sum_f_R0 (fun n : nat => ((-1) ^ n / INR (2 * n + 1) * x2 ^ n)%R) N) (fun x2 : R => x2 ^ 2))%F x) < eps
x0:Base R_met
x0_cond:D_x no_cond x x0 /\ dist R_met x0 x < alpha

dist R_met (x0 * comp (fun x1 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x1 ^ n) N) (fun x1 : R => x1 ^ 2) x0) (x * comp (fun x1 : R => sum_f_R0 (fun n : nat => (-1) ^ n / INR (2 * n + 1) * x1 ^ n) N) (fun x1 : R => x1 ^ 2) x) < eps
Sublemma:forall (x1 : R) (N0 : nat), sum_f_R0 (tg_alt (Ratan_seq x1)) N0 = x1 * comp (fun x2 : R => sum_f_R0 (fun n : nat => (fun i : nat => (-1) ^ i / INR (2 * i + 1)) n * x2 ^ n) N0) (fun x2 : R => x2 ^ 2) x1
N:nat
x:R
x_lb:0 <= x
x_ub:x <= 1
eps:R
eps_pos:eps > 0
continuity_id:continuity id
alpha:R
alpha_pos:alpha > 0
T:forall x1 : Base R_met, D_x no_cond x x1 /\ dist R_met x1 x < alpha -> dist R_met ((id * comp (fun x2 : R => sum_f_R0 (fun n : nat => ((-1) ^ n / INR (2 * n + 1) * x2 ^ n)%R) N) (fun x2 : R => x2 ^ 2))%F x1) ((id * comp (fun x2 : R => sum_f_R0 (fun n : nat => ((-1) ^ n / INR (2 * n + 1) * x2 ^ n)%R) N) (fun x2 : R => x2 ^ 2))%F x) < eps
x0:Base R_met
x0_cond:D_x no_cond x x0 /\ dist R_met x0 x < alpha

D_x no_cond x x0 /\ dist R_met x0 x < alpha
intuition. Qed.
Definition of ps_atan's derivative
Definition Datan_seq := fun (x:R) (n:nat) => x ^ (2*n).


forall (x : R) (n : nat), 0 <= x < 1 -> (0 < n)%nat -> 0 <= x ^ n < 1

forall (x : R) (n : nat), 0 <= x < 1 -> (0 < n)%nat -> 0 <= x ^ n < 1
x:R
hx:0 <= x < 1

0 <= x * 1 < 1
x:R
hx:0 <= x < 1
m:nat
H:(1 <= m)%nat
IHle:0 <= x ^ m < 1
0 <= x * x ^ m < 1
x:R
hx:0 <= x < 1
m:nat
H:(1 <= m)%nat
IHle:0 <= x ^ m < 1

0 <= x * x ^ m < 1
x:R
hx:0 <= x < 1
m:nat
H:(1 <= m)%nat
IHle:0 <= x ^ m < 1

0 <= x * x ^ m
x:R
hx:0 <= x < 1
m:nat
H:(1 <= m)%nat
IHle:0 <= x ^ m < 1
x * x ^ m < 1
x:R
hx:0 <= x < 1
m:nat
H:(1 <= m)%nat
IHle:0 <= x ^ m < 1

x * x ^ m < 1
rewrite <- (Rmult_1_r 1); apply Rmult_le_0_lt_compat; intuition. Qed.

forall (x : R) (n : nat), Datan_seq (Rabs x) n = Datan_seq x n

forall (x : R) (n : nat), Datan_seq (Rabs x) n = Datan_seq x n
intros x n; unfold Datan_seq; rewrite !pow_mult, pow2_abs; reflexivity. Qed.

forall (x : R) (n : nat), 0 < x -> 0 < Datan_seq x n

forall (x : R) (n : nat), 0 < x -> 0 < Datan_seq x n
x:R
x_lb:0 < x

0 < x ^ (2 * 0)
x:R
n:nat
x_lb:0 < x
IHn:0 < x ^ (2 * n)
0 < x ^ (2 * S n)
x:R
n:nat
x_lb:0 < x
IHn:0 < x ^ (2 * n)

0 < x ^ (2 * S n)
x:R
n:nat
x_lb:0 < x
IHn:0 < x ^ (2 * n)

0 < x ^ 2 * x ^ (2 * n)
x:R
n:nat
x_lb:0 < x
IHn:0 < x ^ (2 * n)
x ^ 2 * x ^ (2 * n) = x ^ (2 * S n)
x:R
n:nat
x_lb:0 < x
IHn:0 < x ^ (2 * n)

x ^ 2 > 0
x:R
n:nat
x_lb:0 < x
IHn:0 < x ^ (2 * n)
x ^ (2 * n) > 0
x:R
n:nat
x_lb:0 < x
IHn:0 < x ^ (2 * n)
x ^ 2 * x ^ (2 * n) = x ^ (2 * S n)
x:R
n:nat
x_lb:0 < x
IHn:0 < x ^ (2 * n)

x ^ (2 * n) > 0
x:R
n:nat
x_lb:0 < x
IHn:0 < x ^ (2 * n)
x ^ 2 * x ^ (2 * n) = x ^ (2 * S n)
x:R
n:nat
x_lb:0 < x
IHn:0 < x ^ (2 * n)

x ^ 2 * x ^ (2 * n) = x ^ (2 * S n)
x:R
n:nat
x_lb:0 < x
IHn:0 < x ^ (2 * n)

x ^ 2 * x ^ (2 * n) = x ^ S (S (2 * n))
simpl ; field. Qed.

forall (x : R) (n : nat), sum_f_R0 (tg_alt (Datan_seq x)) n = (1 - (- x ^ 2) ^ S n) / (1 + x ^ 2)

forall (x : R) (n : nat), sum_f_R0 (tg_alt (Datan_seq x)) n = (1 - (- x ^ 2) ^ S n) / (1 + x ^ 2)
x:R
n:nat

sum_f_R0 (tg_alt (Datan_seq x)) n = (1 - (- x ^ 2) ^ S n) / (1 + x ^ 2)
x:R
n:nat

- x ^ 2 <> 1
x:R
n:nat
dif:- x ^ 2 <> 1
sum_f_R0 (tg_alt (Datan_seq x)) n = (1 - (- x ^ 2) ^ S n) / (1 + x ^ 2)
x:R
n:nat

- x ^ 2 <= 0
x:R
n:nat
dif:- x ^ 2 <> 1
sum_f_R0 (tg_alt (Datan_seq x)) n = (1 - (- x ^ 2) ^ S n) / (1 + x ^ 2)
x:R
n:nat
dif:- x ^ 2 <> 1

sum_f_R0 (tg_alt (Datan_seq x)) n = (1 - (- x ^ 2) ^ S n) / (1 + x ^ 2)
x:R
n:nat
dif:- x ^ 2 <> 1

sum_f_R0 (tg_alt (Datan_seq x)) n = sum_f_R0 (fun i : nat => (- x ^ 2) ^ i) n
x:R
n:nat
dif:- x ^ 2 <> 1
i:nat

(-1) ^ i * x ^ (2 * i) = (- x ^ 2) ^ i
x:R
n:nat
dif:- x ^ 2 <> 1
i:nat

(-1 * x ^ 2) ^ i = (- x ^ 2) ^ i
x:R
n:nat
dif:- x ^ 2 <> 1
i:nat

-1 * x ^ 2 = - x ^ 2
ring. Qed.

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

forall (x y : R) (n : nat), (n > 0)%nat -> 0 <= x < y -> Datan_seq x n < Datan_seq y n
x, y:R
n:nat
n_lb:(n > 0)%nat
x_encad:0 <= x < y
x_pos:x >= 0

Datan_seq x n < Datan_seq y n
x, y:R
n:nat
n_lb:(n > 0)%nat
x_encad:0 <= x < y
x_pos:x >= 0

y > 0
x, y:R
n:nat
n_lb:(n > 0)%nat
x_encad:0 <= x < y
x_pos:x >= 0
y_pos:y > 0
Datan_seq x n < Datan_seq y n
x, y:R
n:nat
n_lb:(n > 0)%nat
x_encad:0 <= x < y
x_pos:x >= 0
y_pos:y > 0

Datan_seq x n < Datan_seq y n
x, y:R
n_lb:(0 > 0)%nat
x_encad:0 <= x < y
x_pos:x >= 0
y_pos:y > 0

Datan_seq x 0 < Datan_seq y 0
x, y:R
n:nat
n_lb:(S n > 0)%nat
x_encad:0 <= x < y
x_pos:x >= 0
y_pos:y > 0
IHn:(n > 0)%nat -> Datan_seq x n < Datan_seq y n
Datan_seq x (S n) < Datan_seq y (S n)
x, y:R
n:nat
n_lb:(S n > 0)%nat
x_encad:0 <= x < y
x_pos:x >= 0
y_pos:y > 0
IHn:(n > 0)%nat -> Datan_seq x n < Datan_seq y n

Datan_seq x (S n) < Datan_seq y (S n)
x, y:R
x_encad:0 <= x < y
x_pos:x >= 0
y_pos:y > 0

x ^ (2 * 1) < y ^ (2 * 1)
x, y:R
n:nat
x_encad:0 <= x < y
x_pos:x >= 0
y_pos:y > 0
IHn:Datan_seq x (S n) < Datan_seq y (S n)
x ^ (2 * S (S n)) < y ^ (2 * S (S n))
x, y:R
x_encad:0 <= x < y
y_pos:y > 0
x_pos:x > 0

x ^ (2 * 1) < y ^ (2 * 1)
x, y:R
x_encad:0 <= x < y
y_pos:y > 0
x_pos:x = 0
x ^ (2 * 1) < y ^ (2 * 1)
x, y:R
n:nat
x_encad:0 <= x < y
x_pos:x >= 0
y_pos:y > 0
IHn:Datan_seq x (S n) < Datan_seq y (S n)
x ^ (2 * S (S n)) < y ^ (2 * S (S n))
x, y:R
y_pos:y > 0
x_pos:x > 0
H:0 <= x
H0:x < y

x * 1 > 0
x, y:R
x_encad:0 <= x < y
y_pos:y > 0
x_pos:x = 0
x ^ (2 * 1) < y ^ (2 * 1)
x, y:R
n:nat
x_encad:0 <= x < y
x_pos:x >= 0
y_pos:y > 0
IHn:Datan_seq x (S n) < Datan_seq y (S n)
x ^ (2 * S (S n)) < y ^ (2 * S (S n))
x, y:R
x_encad:0 <= x < y
y_pos:y > 0
x_pos:x = 0

x ^ (2 * 1) < y ^ (2 * 1)
x, y:R
n:nat
x_encad:0 <= x < y
x_pos:x >= 0
y_pos:y > 0
IHn:Datan_seq x (S n) < Datan_seq y (S n)
x ^ (2 * S (S n)) < y ^ (2 * S (S n))
x, y:R
x_encad:0 <= x < y
y_pos:y > 0
x_pos:x = 0

0 < y ^ (2 * 1)
x, y:R
x_encad:0 <= x < y
y_pos:y > 0
x_pos:x = 0
(0 < 2 * 1)%nat
x, y:R
n:nat
x_encad:0 <= x < y
x_pos:x >= 0
y_pos:y > 0
IHn:Datan_seq x (S n) < Datan_seq y (S n)
x ^ (2 * S (S n)) < y ^ (2 * S (S n))
x, y:R
x_encad:0 <= x < y
y_pos:y > 0
x_pos:x = 0

0 < y * y
x, y:R
x_encad:0 <= x < y
y_pos:y > 0
x_pos:x = 0
y * y = y ^ (2 * 1)
x, y:R
x_encad:0 <= x < y
y_pos:y > 0
x_pos:x = 0
(0 < 2 * 1)%nat
x, y:R
n:nat
x_encad:0 <= x < y
x_pos:x >= 0
y_pos:y > 0
IHn:Datan_seq x (S n) < Datan_seq y (S n)
x ^ (2 * S (S n)) < y ^ (2 * S (S n))
x, y:R
x_encad:0 <= x < y
y_pos:y > 0
x_pos:x = 0

y * y = y ^ (2 * 1)
x, y:R
x_encad:0 <= x < y
y_pos:y > 0
x_pos:x = 0
(0 < 2 * 1)%nat
x, y:R
n:nat
x_encad:0 <= x < y
x_pos:x >= 0
y_pos:y > 0
IHn:Datan_seq x (S n) < Datan_seq y (S n)
x ^ (2 * S (S n)) < y ^ (2 * S (S n))
x, y:R
x_encad:0 <= x < y
y_pos:y > 0
x_pos:x = 0

(0 < 2 * 1)%nat
x, y:R
n:nat
x_encad:0 <= x < y
x_pos:x >= 0
y_pos:y > 0
IHn:Datan_seq x (S n) < Datan_seq y (S n)
x ^ (2 * S (S n)) < y ^ (2 * S (S n))
x, y:R
n:nat
x_encad:0 <= x < y
x_pos:x >= 0
y_pos:y > 0
IHn:Datan_seq x (S n) < Datan_seq y (S n)

x ^ (2 * S (S n)) < y ^ (2 * S (S n))
x, y:R
n:nat
x_encad:0 <= x < y
x_pos:x >= 0
y_pos:y > 0
IHn:Datan_seq x (S n) < Datan_seq y (S n)

forall a : R, a ^ (2 * S (S n)) = a ^ 2 * a ^ (2 * S n)
x, y:R
n:nat
x_encad:0 <= x < y
x_pos:x >= 0
y_pos:y > 0
IHn:Datan_seq x (S n) < Datan_seq y (S n)
Hrew:forall a : R, a ^ (2 * S (S n)) = a ^ 2 * a ^ (2 * S n)
x ^ (2 * S (S n)) < y ^ (2 * S (S n))
n:nat
a:R

a ^ S (S (2 * S n)) = a ^ 2 * a ^ (2 * S n)
x, y:R
n:nat
x_encad:0 <= x < y
x_pos:x >= 0
y_pos:y > 0
IHn:Datan_seq x (S n) < Datan_seq y (S n)
Hrew:forall a : R, a ^ (2 * S (S n)) = a ^ 2 * a ^ (2 * S n)
x ^ (2 * S (S n)) < y ^ (2 * S (S n))
x, y:R
n:nat
x_encad:0 <= x < y
x_pos:x >= 0
y_pos:y > 0
IHn:Datan_seq x (S n) < Datan_seq y (S n)
Hrew:forall a : R, a ^ (2 * S (S n)) = a ^ 2 * a ^ (2 * S n)

x ^ (2 * S (S n)) < y ^ (2 * S (S n))
x, y:R
n:nat
x_encad:0 <= x < y
y_pos:y > 0
IHn:Datan_seq x (S n) < Datan_seq y (S n)
Hrew:forall a : R, a ^ (2 * S (S n)) = a ^ 2 * a ^ (2 * S n)
x_pos:x > 0

x ^ (2 * S (S n)) < y ^ (2 * S (S n))
x, y:R
n:nat
x_encad:0 <= x < y
y_pos:y > 0
IHn:Datan_seq x (S n) < Datan_seq y (S n)
Hrew:forall a : R, a ^ (2 * S (S n)) = a ^ 2 * a ^ (2 * S n)
x_pos:x = 0
x ^ (2 * S (S n)) < y ^ (2 * S (S n))
x, y:R
n:nat
x_encad:0 <= x < y
y_pos:y > 0
IHn:Datan_seq x (S n) < Datan_seq y (S n)
Hrew:forall a : R, a ^ (2 * S (S n)) = a ^ 2 * a ^ (2 * S n)
x_pos:x > 0

x ^ 2 * x ^ (2 * S n) < y ^ 2 * y ^ (2 * S n)
x, y:R
n:nat
x_encad:0 <= x < y
y_pos:y > 0
IHn:Datan_seq x (S n) < Datan_seq y (S n)
Hrew:forall a : R, a ^ (2 * S (S n)) = a ^ 2 * a ^ (2 * S n)
x_pos:x = 0
x ^ (2 * S (S n)) < y ^ (2 * S (S n))
x, y:R
n:nat
y_pos:y > 0
IHn:Datan_seq x (S n) < Datan_seq y (S n)
Hrew:forall a : R, a ^ (2 * S (S n)) = a ^ 2 * a ^ (2 * S n)
x_pos:x > 0
H:0 <= x
H0:x < y

x ^ 2 < y ^ 2
x, y:R
n:nat
x_encad:0 <= x < y
y_pos:y > 0
IHn:Datan_seq x (S n) < Datan_seq y (S n)
Hrew:forall a : R, a ^ (2 * S (S n)) = a ^ 2 * a ^ (2 * S n)
x_pos:x = 0
x ^ (2 * S (S n)) < y ^ (2 * S (S n))
x, y:R
n:nat
x_encad:0 <= x < y
y_pos:y > 0
IHn:Datan_seq x (S n) < Datan_seq y (S n)
Hrew:forall a : R, a ^ (2 * S (S n)) = a ^ 2 * a ^ (2 * S n)
x_pos:x = 0

x ^ (2 * S (S n)) < y ^ (2 * S (S n))
x, y:R
n:nat
x_encad:0 <= x < y
y_pos:y > 0
IHn:Datan_seq x (S n) < Datan_seq y (S n)
Hrew:forall a : R, a ^ (2 * S (S n)) = a ^ 2 * a ^ (2 * S n)
x_pos:x = 0

0 ^ (2 * S (S n)) < y ^ (2 * S (S n))
rewrite pow_i ; intuition. Qed.

forall x : R, -1 < x -> x < 1 -> Un_decreasing (Datan_seq x)

forall x : R, -1 < x -> x < 1 -> Un_decreasing (Datan_seq x)
x:R
x_lb:-1 < x
x_ub:x < 1
n:nat

Datan_seq x (S n) <= Datan_seq x n
x:R
x_lb:-1 < x
x_ub:x < 1
n:nat

x ^ (2 * S n) <= x ^ (2 * n)
x:R
x_lb:-1 < x
x_ub:x < 1
n:nat

x ^ (2 + 2 * n) <= x ^ (2 * n)
x:R
x_lb:-1 < x
x_ub:x < 1
n:nat

x ^ (2 + 2 * n) <= 1 * x ^ (2 * n)
x:R
x_lb:-1 < x
x_ub:x < 1
n:nat

x ^ 2 * x ^ (2 * n) <= 1 * x ^ (2 * n)
x:R
x_lb:-1 < x
x_ub:x < 1
n:nat

0 <= x ^ (2 * n)
x:R
x_lb:-1 < x
x_ub:x < 1
n:nat
x ^ 2 <= 1
x:R
x_lb:-1 < x
x_ub:x < 1
n:nat

x ^ 2 <= 1
x:R
x_lb:-1 < x
x_ub:x < 1
n:nat

Rabs x ^ 2 < 1
x:R
x_lb:-1 < x
x_ub:x < 1
n:nat

0 <= Rabs x < 1
x:R
x_lb:-1 < x
x_ub:x < 1
n:nat
intabs:0 <= Rabs x < 1
Rabs x ^ 2 < 1
x:R
x_lb:-1 < x
x_ub:x < 1
n:nat
intabs:0 <= Rabs x < 1

Rabs x ^ 2 < 1
x:R
x_lb:-1 < x
x_ub:x < 1
n:nat
intabs:0 <= Rabs x ^ 2 < 1

Rabs x ^ 2 < 1
x:R
x_lb:-1 < x
x_ub:x < 1
n:nat
intabs:0 <= Rabs x < 1
(0 < 2)%nat
x:R
x_lb:-1 < x
x_ub:x < 1
n:nat
intabs:0 <= Rabs x < 1

(0 < 2)%nat
omega. Qed.

forall x : R, -1 < x -> x < 1 -> Un_cv (Datan_seq x) 0

forall x : R, -1 < x -> x < 1 -> Un_cv (Datan_seq x) 0
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0

exists N : nat, forall n : nat, (n >= N)%nat -> R_dist (Datan_seq x n) 0 < eps
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0

Rabs (x ^ 2) < 1
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
x_ub2:Rabs (x ^ 2) < 1
exists N : nat, forall n : nat, (n >= N)%nat -> R_dist (Datan_seq x n) 0 < eps
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0

x ^ 2 < 1
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
x_ub2:Rabs (x ^ 2) < 1
exists N : nat, forall n : nat, (n >= N)%nat -> R_dist (Datan_seq x n) 0 < eps
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0

Rabs x ^ 2 < 1
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
x_ub2:Rabs (x ^ 2) < 1
exists N : nat, forall n : nat, (n >= N)%nat -> R_dist (Datan_seq x n) 0 < eps
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
H:0 <= Rabs x < 1

Rabs x ^ 2 < 1
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
x_ub2:Rabs (x ^ 2) < 1
exists N : nat, forall n : nat, (n >= N)%nat -> R_dist (Datan_seq x n) 0 < eps
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
x_ub2:Rabs (x ^ 2) < 1

exists N : nat, forall n : nat, (n >= N)%nat -> R_dist (Datan_seq x n) 0 < eps
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
x_ub2:Rabs (x ^ 2) < 1
N:nat
HN:forall n0 : nat, (n0 >= N)%nat -> Rabs ((x ^ 2) ^ n0) < eps
n:nat
Hn:(n >= N)%nat

R_dist (Datan_seq x n) 0 < eps
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
x_ub2:Rabs (x ^ 2) < 1
N:nat
HN:forall n0 : nat, (n0 >= N)%nat -> Rabs ((x ^ 2) ^ n0) < eps
n:nat
Hn:(n >= N)%nat

Rabs (x ^ (2 * n) - 0) < eps
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
x_ub2:Rabs (x ^ 2) < 1
N:nat
HN:forall n0 : nat, (n0 >= N)%nat -> Rabs ((x ^ 2) ^ n0) < eps
n:nat
Hn:(n >= N)%nat

Rabs ((x ^ 2) ^ n) < eps
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
x_ub2:Rabs (x ^ 2) < 1
N:nat
HN:forall n0 : nat, (n0 >= N)%nat -> Rabs ((x ^ 2) ^ n0) < eps
n:nat
Hn:(n >= N)%nat
(x ^ 2) ^ n = x ^ (2 * n) - 0
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
x_ub2:Rabs (x ^ 2) < 1
N:nat
HN:forall n0 : nat, (n0 >= N)%nat -> Rabs ((x ^ 2) ^ n0) < eps
n:nat
Hn:(n >= N)%nat

(x ^ 2) ^ n = x ^ (2 * n) - 0
rewrite pow_mult ; field. Qed.

forall x : R, -1 < x -> x < 1 -> Un_cv (fun N : nat => sum_f_R0 (tg_alt (Datan_seq x)) N) (/ (1 + x ^ 2))

forall x : R, -1 < x -> x < 1 -> Un_cv (fun N : nat => sum_f_R0 (tg_alt (Datan_seq x)) N) (/ (1 + x ^ 2))
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0

exists N : nat, forall n : nat, (n >= N)%nat -> R_dist (sum_f_R0 (tg_alt (Datan_seq x)) n) (/ (1 + x ^ 2)) < eps
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2

exists N : nat, forall n : nat, (n >= N)%nat -> R_dist (sum_f_R0 (tg_alt (Datan_seq x)) n) (/ (1 + x ^ 2)) < eps
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2

0 < 1 + x ^ 2
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
exists N : nat, forall n : nat, (n >= N)%nat -> R_dist (sum_f_R0 (tg_alt (Datan_seq x)) n) (/ (1 + x ^ 2)) < eps
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2

exists N : nat, forall n : nat, (n >= N)%nat -> R_dist (sum_f_R0 (tg_alt (Datan_seq x)) n) (/ (1 + x ^ 2)) < eps
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2

/ (1 + x ^ 2) > 0
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0
exists N : nat, forall n : nat, (n >= N)%nat -> R_dist (sum_f_R0 (tg_alt (Datan_seq x)) n) (/ (1 + x ^ 2)) < eps
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0

exists N : nat, forall n : nat, (n >= N)%nat -> R_dist (sum_f_R0 (tg_alt (Datan_seq x)) n) (/ (1 + x ^ 2)) < eps
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0

0 <= Rabs (x ^ 2) < 1
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0
x_ub2':0 <= Rabs (x ^ 2) < 1
exists N : nat, forall n : nat, (n >= N)%nat -> R_dist (sum_f_R0 (tg_alt (Datan_seq x)) n) (/ (1 + x ^ 2)) < eps
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0

0 <= Rabs x ^ 2 < 1
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0
x_ub2':0 <= Rabs (x ^ 2) < 1
exists N : nat, forall n : nat, (n >= N)%nat -> R_dist (sum_f_R0 (tg_alt (Datan_seq x)) n) (/ (1 + x ^ 2)) < eps
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0

Rabs x < 1
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0
x_ub2':0 <= Rabs (x ^ 2) < 1
exists N : nat, forall n : nat, (n >= N)%nat -> R_dist (sum_f_R0 (tg_alt (Datan_seq x)) n) (/ (1 + x ^ 2)) < eps
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0
x_ub2':0 <= Rabs (x ^ 2) < 1

exists N : nat, forall n : nat, (n >= N)%nat -> R_dist (sum_f_R0 (tg_alt (Datan_seq x)) n) (/ (1 + x ^ 2)) < eps
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0
x_ub2':0 <= Rabs (x ^ 2) < 1
x_ub2:Rabs (x ^ 2) < 1

exists N : nat, forall n : nat, (n >= N)%nat -> R_dist (sum_f_R0 (tg_alt (Datan_seq x)) n) (/ (1 + x ^ 2)) < eps
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0
x_ub2':0 <= Rabs (x ^ 2) < 1
x_ub2:Rabs (x ^ 2) < 1

(1 + x ^ 2) * eps > 0
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0
x_ub2':0 <= Rabs (x ^ 2) < 1
x_ub2:Rabs (x ^ 2) < 1
eps'_pos:(1 + x ^ 2) * eps > 0
exists N : nat, forall n : nat, (n >= N)%nat -> R_dist (sum_f_R0 (tg_alt (Datan_seq x)) n) (/ (1 + x ^ 2)) < eps
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0
x_ub2':0 <= Rabs (x ^ 2) < 1
x_ub2:Rabs (x ^ 2) < 1
eps'_pos:(1 + x ^ 2) * eps > 0

exists N : nat, forall n : nat, (n >= N)%nat -> R_dist (sum_f_R0 (tg_alt (Datan_seq x)) n) (/ (1 + x ^ 2)) < eps
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0
x_ub2':0 <= Rabs (x ^ 2) < 1
x_ub2:Rabs (x ^ 2) < 1
eps'_pos:(1 + x ^ 2) * eps > 0
N:nat
HN:forall n : nat, (n >= N)%nat -> Rabs ((x ^ 2) ^ n) < (1 + x ^ 2) * eps

forall n : nat, (n >= N)%nat -> R_dist (sum_f_R0 (tg_alt (Datan_seq x)) n) (/ (1 + x ^ 2)) < eps
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0
x_ub2':0 <= Rabs (x ^ 2) < 1
x_ub2:Rabs (x ^ 2) < 1
eps'_pos:(1 + x ^ 2) * eps > 0
N:nat
HN:forall n0 : nat, (n0 >= N)%nat -> Rabs ((x ^ 2) ^ n0) < (1 + x ^ 2) * eps
n:nat
Hn:(n >= N)%nat

R_dist (sum_f_R0 (tg_alt (Datan_seq x)) n) (/ (1 + x ^ 2)) < eps
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0
x_ub2':0 <= Rabs (x ^ 2) < 1
x_ub2:Rabs (x ^ 2) < 1
eps'_pos:(1 + x ^ 2) * eps > 0
N:nat
HN:forall n0 : nat, (n0 >= N)%nat -> Rabs ((x ^ 2) ^ n0) < (1 + x ^ 2) * eps
n:nat
Hn:(n >= N)%nat

- x ^ 2 <> 1
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0
x_ub2':0 <= Rabs (x ^ 2) < 1
x_ub2:Rabs (x ^ 2) < 1
eps'_pos:(1 + x ^ 2) * eps > 0
N:nat
HN:forall n0 : nat, (n0 >= N)%nat -> Rabs ((x ^ 2) ^ n0) < (1 + x ^ 2) * eps
n:nat
Hn:(n >= N)%nat
H1:- x ^ 2 <> 1
R_dist (sum_f_R0 (tg_alt (Datan_seq x)) n) (/ (1 + x ^ 2)) < eps
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0
x_ub2':0 <= Rabs (x ^ 2) < 1
x_ub2:Rabs (x ^ 2) < 1
eps'_pos:(1 + x ^ 2) * eps > 0
N:nat
HN:forall n0 : nat, (n0 >= N)%nat -> Rabs ((x ^ 2) ^ n0) < (1 + x ^ 2) * eps
n:nat
Hn:(n >= N)%nat

- x ^ 2 <= 0
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0
x_ub2':0 <= Rabs (x ^ 2) < 1
x_ub2:Rabs (x ^ 2) < 1
eps'_pos:(1 + x ^ 2) * eps > 0
N:nat
HN:forall n0 : nat, (n0 >= N)%nat -> Rabs ((x ^ 2) ^ n0) < (1 + x ^ 2) * eps
n:nat
Hn:(n >= N)%nat
H1:- x ^ 2 <> 1
R_dist (sum_f_R0 (tg_alt (Datan_seq x)) n) (/ (1 + x ^ 2)) < eps
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0
x_ub2':0 <= Rabs (x ^ 2) < 1
x_ub2:Rabs (x ^ 2) < 1
eps'_pos:(1 + x ^ 2) * eps > 0
N:nat
HN:forall n0 : nat, (n0 >= N)%nat -> Rabs ((x ^ 2) ^ n0) < (1 + x ^ 2) * eps
n:nat
Hn:(n >= N)%nat
H1:- x ^ 2 <> 1

R_dist (sum_f_R0 (tg_alt (Datan_seq x)) n) (/ (1 + x ^ 2)) < eps
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0
x_ub2':0 <= Rabs (x ^ 2) < 1
x_ub2:Rabs (x ^ 2) < 1
eps'_pos:(1 + x ^ 2) * eps > 0
N:nat
HN:forall n0 : nat, (n0 >= N)%nat -> Rabs ((x ^ 2) ^ n0) < (1 + x ^ 2) * eps
n:nat
Hn:(n >= N)%nat
H1:- x ^ 2 <> 1

R_dist ((1 - (- x ^ 2) ^ S n) / (1 + x ^ 2)) (/ (1 + x ^ 2)) < eps
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0
x_ub2':0 <= Rabs (x ^ 2) < 1
x_ub2:Rabs (x ^ 2) < 1
eps'_pos:(1 + x ^ 2) * eps > 0
N:nat
HN:forall n0 : nat, (n0 >= N)%nat -> Rabs ((x ^ 2) ^ n0) < (1 + x ^ 2) * eps
n:nat
Hn:(n >= N)%nat
H1:- x ^ 2 <> 1

Rabs ((1 - (- x ^ 2) ^ S n) / (1 + x ^ 2) - / (1 + x ^ 2)) < eps
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0
x_ub2':0 <= Rabs (x ^ 2) < 1
x_ub2:Rabs (x ^ 2) < 1
eps'_pos:(1 + x ^ 2) * eps > 0
N:nat
HN:forall n0 : nat, (n0 >= N)%nat -> Rabs ((x ^ 2) ^ n0) < (1 + x ^ 2) * eps
n:nat
Hn:(n >= N)%nat
H1:- x ^ 2 <> 1

forall a b : R, a / b - / b = (-1 + a) / b
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0
x_ub2':0 <= Rabs (x ^ 2) < 1
x_ub2:Rabs (x ^ 2) < 1
eps'_pos:(1 + x ^ 2) * eps > 0
N:nat
HN:forall n0 : nat, (n0 >= N)%nat -> Rabs ((x ^ 2) ^ n0) < (1 + x ^ 2) * eps
n:nat
Hn:(n >= N)%nat
H1:- x ^ 2 <> 1
tool:forall a b : R, a / b - / b = (-1 + a) / b
Rabs ((1 - (- x ^ 2) ^ S n) / (1 + x ^ 2) - / (1 + x ^ 2)) < eps
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0
x_ub2':0 <= Rabs (x ^ 2) < 1
x_ub2:Rabs (x ^ 2) < 1
eps'_pos:(1 + x ^ 2) * eps > 0
N:nat
HN:forall n0 : nat, (n0 >= N)%nat -> Rabs ((x ^ 2) ^ n0) < (1 + x ^ 2) * eps
n:nat
Hn:(n >= N)%nat
H1:- x ^ 2 <> 1
a, b:R

a * / b + - (1 * / b) = (-1 + a) * / b
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0
x_ub2':0 <= Rabs (x ^ 2) < 1
x_ub2:Rabs (x ^ 2) < 1
eps'_pos:(1 + x ^ 2) * eps > 0
N:nat
HN:forall n0 : nat, (n0 >= N)%nat -> Rabs ((x ^ 2) ^ n0) < (1 + x ^ 2) * eps
n:nat
Hn:(n >= N)%nat
H1:- x ^ 2 <> 1
tool:forall a b : R, a / b - / b = (-1 + a) / b
Rabs ((1 - (- x ^ 2) ^ S n) / (1 + x ^ 2) - / (1 + x ^ 2)) < eps
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0
x_ub2':0 <= Rabs (x ^ 2) < 1
x_ub2:Rabs (x ^ 2) < 1
eps'_pos:(1 + x ^ 2) * eps > 0
N:nat
HN:forall n0 : nat, (n0 >= N)%nat -> Rabs ((x ^ 2) ^ n0) < (1 + x ^ 2) * eps
n:nat
Hn:(n >= N)%nat
H1:- x ^ 2 <> 1
a, b:R

- (1) * / b + a * / b = -1 * / b + a * / b
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0
x_ub2':0 <= Rabs (x ^ 2) < 1
x_ub2:Rabs (x ^ 2) < 1
eps'_pos:(1 + x ^ 2) * eps > 0
N:nat
HN:forall n0 : nat, (n0 >= N)%nat -> Rabs ((x ^ 2) ^ n0) < (1 + x ^ 2) * eps
n:nat
Hn:(n >= N)%nat
H1:- x ^ 2 <> 1
tool:forall a b : R, a / b - / b = (-1 + a) / b
Rabs ((1 - (- x ^ 2) ^ S n) / (1 + x ^ 2) - / (1 + x ^ 2)) < eps
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0
x_ub2':0 <= Rabs (x ^ 2) < 1
x_ub2:Rabs (x ^ 2) < 1
eps'_pos:(1 + x ^ 2) * eps > 0
N:nat
HN:forall n0 : nat, (n0 >= N)%nat -> Rabs ((x ^ 2) ^ n0) < (1 + x ^ 2) * eps
n:nat
Hn:(n >= N)%nat
H1:- x ^ 2 <> 1
tool:forall a b : R, a / b - / b = (-1 + a) / b

Rabs ((1 - (- x ^ 2) ^ S n) / (1 + x ^ 2) - / (1 + x ^ 2)) < eps
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0
x_ub2':0 <= Rabs (x ^ 2) < 1
x_ub2:Rabs (x ^ 2) < 1
eps'_pos:(1 + x ^ 2) * eps > 0
N:nat
HN:forall n0 : nat, (n0 >= N)%nat -> Rabs ((x ^ 2) ^ n0) < (1 + x ^ 2) * eps
n:nat
Hn:(n >= N)%nat
H1:- x ^ 2 <> 1
tool:forall a b : R, a / b - / b = (-1 + a) / b
u:=1 + x ^ 2:R

Rabs ((-1 + 1 + - (- x ^ 2) ^ S n) / u) < eps
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0
x_ub2':0 <= Rabs (x ^ 2) < 1
x_ub2:Rabs (x ^ 2) < 1
eps'_pos:(1 + x ^ 2) * eps > 0
N:nat
HN:forall n0 : nat, (n0 >= N)%nat -> Rabs ((x ^ 2) ^ n0) < (1 + x ^ 2) * eps
n:nat
Hn:(n >= N)%nat
H1:- x ^ 2 <> 1
tool:forall a b : R, a / b - / b = (-1 + a) / b
u:=1 + x ^ 2:R

Rabs ((-1 + 1 + - (- x ^ 2) ^ S n) * / (1 + x ^ 2)) < eps
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0
x_ub2':0 <= Rabs (x ^ 2) < 1
x_ub2:Rabs (x ^ 2) < 1
eps'_pos:(1 + x ^ 2) * eps > 0
N:nat
HN:forall n0 : nat, (n0 >= N)%nat -> Rabs ((x ^ 2) ^ n0) < (1 + x ^ 2) * eps
n:nat
Hn:(n >= N)%nat
H1:- x ^ 2 <> 1
tool:forall a b : R, a / b - / b = (-1 + a) / b
u:=1 + x ^ 2:R

Rabs ((- (1) + 1 + - (- x ^ 2) ^ S n) * / (1 + x ^ 2)) < eps
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0
x_ub2':0 <= Rabs (x ^ 2) < 1
x_ub2:Rabs (x ^ 2) < 1
eps'_pos:(1 + x ^ 2) * eps > 0
N:nat
HN:forall n0 : nat, (n0 >= N)%nat -> Rabs ((x ^ 2) ^ n0) < (1 + x ^ 2) * eps
n:nat
Hn:(n >= N)%nat
H1:- x ^ 2 <> 1
tool:forall a b : R, a / b - / b = (-1 + a) / b
u:=1 + x ^ 2:R

Rabs ((- x ^ 2) ^ S n * / (1 + x ^ 2)) < eps
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0
x_ub2':0 <= Rabs (x ^ 2) < 1
x_ub2:Rabs (x ^ 2) < 1
eps'_pos:(1 + x ^ 2) * eps > 0
N:nat
HN:forall n0 : nat, (n0 >= N)%nat -> Rabs ((x ^ 2) ^ n0) < (1 + x ^ 2) * eps
n:nat
Hn:(n >= N)%nat
H1:- x ^ 2 <> 1

Rabs ((- x ^ 2) ^ S n) * Rabs (/ (1 + x ^ 2)) < eps
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0
x_ub2':0 <= Rabs (x ^ 2) < 1
x_ub2:Rabs (x ^ 2) < 1
eps'_pos:(1 + x ^ 2) * eps > 0
N:nat
HN:forall n0 : nat, (n0 >= N)%nat -> Rabs ((x ^ 2) ^ n0) < (1 + x ^ 2) * eps
n:nat
Hn:(n >= N)%nat
H1:- x ^ 2 <> 1

forall k : nat, Rabs ((- x ^ 2) ^ k) = Rabs ((x ^ 2) ^ k)
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0
x_ub2':0 <= Rabs (x ^ 2) < 1
x_ub2:Rabs (x ^ 2) < 1
eps'_pos:(1 + x ^ 2) * eps > 0
N:nat
HN:forall n0 : nat, (n0 >= N)%nat -> Rabs ((x ^ 2) ^ n0) < (1 + x ^ 2) * eps
n:nat
Hn:(n >= N)%nat
H1:- x ^ 2 <> 1
tool:forall k : nat, Rabs ((- x ^ 2) ^ k) = Rabs ((x ^ 2) ^ k)
Rabs ((- x ^ 2) ^ S n) * Rabs (/ (1 + x ^ 2)) < eps
x:R
Tool0:0 <= x ^ 2
k:nat
IHk:Rabs ((- x ^ 2) ^ k) = Rabs ((x ^ 2) ^ k)

Rabs ((- x ^ 2) ^ S k) = Rabs ((x ^ 2) ^ S k)
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0
x_ub2':0 <= Rabs (x ^ 2) < 1
x_ub2:Rabs (x ^ 2) < 1
eps'_pos:(1 + x ^ 2) * eps > 0
N:nat
HN:forall n0 : nat, (n0 >= N)%nat -> Rabs ((x ^ 2) ^ n0) < (1 + x ^ 2) * eps
n:nat
Hn:(n >= N)%nat
H1:- x ^ 2 <> 1
tool:forall k : nat, Rabs ((- x ^ 2) ^ k) = Rabs ((x ^ 2) ^ k)
Rabs ((- x ^ 2) ^ S n) * Rabs (/ (1 + x ^ 2)) < eps
x:R
Tool0:0 <= x ^ 2
k:nat
IHk:Rabs ((- x ^ 2) ^ k) = Rabs ((x ^ 2) ^ k)

x ^ 2 * Rabs ((x ^ 2) ^ k) = x ^ 2 * Rabs ((x ^ 2) ^ k)
x:R
Tool0:0 <= x ^ 2
k:nat
IHk:Rabs ((- x ^ 2) ^ k) = Rabs ((x ^ 2) ^ k)
0 <= x ^ 2
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0
x_ub2':0 <= Rabs (x ^ 2) < 1
x_ub2:Rabs (x ^ 2) < 1
eps'_pos:(1 + x ^ 2) * eps > 0
N:nat
HN:forall n0 : nat, (n0 >= N)%nat -> Rabs ((x ^ 2) ^ n0) < (1 + x ^ 2) * eps
n:nat
Hn:(n >= N)%nat
H1:- x ^ 2 <> 1
tool:forall k : nat, Rabs ((- x ^ 2) ^ k) = Rabs ((x ^ 2) ^ k)
Rabs ((- x ^ 2) ^ S n) * Rabs (/ (1 + x ^ 2)) < eps
x:R
Tool0:0 <= x ^ 2
k:nat
IHk:Rabs ((- x ^ 2) ^ k) = Rabs ((x ^ 2) ^ k)

0 <= x ^ 2
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0
x_ub2':0 <= Rabs (x ^ 2) < 1
x_ub2:Rabs (x ^ 2) < 1
eps'_pos:(1 + x ^ 2) * eps > 0
N:nat
HN:forall n0 : nat, (n0 >= N)%nat -> Rabs ((x ^ 2) ^ n0) < (1 + x ^ 2) * eps
n:nat
Hn:(n >= N)%nat
H1:- x ^ 2 <> 1
tool:forall k : nat, Rabs ((- x ^ 2) ^ k) = Rabs ((x ^ 2) ^ k)
Rabs ((- x ^ 2) ^ S n) * Rabs (/ (1 + x ^ 2)) < eps
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0
x_ub2':0 <= Rabs (x ^ 2) < 1
x_ub2:Rabs (x ^ 2) < 1
eps'_pos:(1 + x ^ 2) * eps > 0
N:nat
HN:forall n0 : nat, (n0 >= N)%nat -> Rabs ((x ^ 2) ^ n0) < (1 + x ^ 2) * eps
n:nat
Hn:(n >= N)%nat
H1:- x ^ 2 <> 1
tool:forall k : nat, Rabs ((- x ^ 2) ^ k) = Rabs ((x ^ 2) ^ k)

Rabs ((- x ^ 2) ^ S n) * Rabs (/ (1 + x ^ 2)) < eps
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0
x_ub2':0 <= Rabs (x ^ 2) < 1
x_ub2:Rabs (x ^ 2) < 1
eps'_pos:(1 + x ^ 2) * eps > 0
N:nat
HN:forall n0 : nat, (n0 >= N)%nat -> Rabs ((x ^ 2) ^ n0) < (1 + x ^ 2) * eps
n:nat
Hn:(n >= N)%nat
H1:- x ^ 2 <> 1

Rabs ((x ^ 2) ^ S n) * / (1 + x ^ 2) < eps
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0
x_ub2':0 <= Rabs (x ^ 2) < 1
x_ub2:Rabs (x ^ 2) < 1
eps'_pos:(1 + x ^ 2) * eps > 0
N:nat
HN:forall n0 : nat, (n0 >= N)%nat -> Rabs ((x ^ 2) ^ n0) < (1 + x ^ 2) * eps
n:nat
Hn:(n >= N)%nat
H1:- x ^ 2 <> 1

forall a b c : R, 0 < b -> a < b * c -> a * / b < c
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0
x_ub2':0 <= Rabs (x ^ 2) < 1
x_ub2:Rabs (x ^ 2) < 1
eps'_pos:(1 + x ^ 2) * eps > 0
N:nat
HN:forall n0 : nat, (n0 >= N)%nat -> Rabs ((x ^ 2) ^ n0) < (1 + x ^ 2) * eps
n:nat
Hn:(n >= N)%nat
H1:- x ^ 2 <> 1
tool:forall a b c : R, 0 < b -> a < b * c -> a * / b < c
Rabs ((x ^ 2) ^ S n) * / (1 + x ^ 2) < eps
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0
x_ub2':0 <= Rabs (x ^ 2) < 1
x_ub2:Rabs (x ^ 2) < 1
eps'_pos:(1 + x ^ 2) * eps > 0
N:nat
HN:forall n0 : nat, (n0 >= N)%nat -> Rabs ((x ^ 2) ^ n0) < (1 + x ^ 2) * eps
n:nat
Hn:(n >= N)%nat
H1:- x ^ 2 <> 1
a, b, c:R
bp:0 < b
h:a < b * c

a * / b < b * c * / b
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0
x_ub2':0 <= Rabs (x ^ 2) < 1
x_ub2:Rabs (x ^ 2) < 1
eps'_pos:(1 + x ^ 2) * eps > 0
N:nat
HN:forall n0 : nat, (n0 >= N)%nat -> Rabs ((x ^ 2) ^ n0) < (1 + x ^ 2) * eps
n:nat
Hn:(n >= N)%nat
H1:- x ^ 2 <> 1
a, b, c:R
bp:0 < b
h:a < b * c
b * c * / b = c
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0
x_ub2':0 <= Rabs (x ^ 2) < 1
x_ub2:Rabs (x ^ 2) < 1
eps'_pos:(1 + x ^ 2) * eps > 0
N:nat
HN:forall n0 : nat, (n0 >= N)%nat -> Rabs ((x ^ 2) ^ n0) < (1 + x ^ 2) * eps
n:nat
Hn:(n >= N)%nat
H1:- x ^ 2 <> 1
tool:forall a b c : R, 0 < b -> a < b * c -> a * / b < c
Rabs ((x ^ 2) ^ S n) * / (1 + x ^ 2) < eps
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0
x_ub2':0 <= Rabs (x ^ 2) < 1
x_ub2:Rabs (x ^ 2) < 1
eps'_pos:(1 + x ^ 2) * eps > 0
N:nat
HN:forall n0 : nat, (n0 >= N)%nat -> Rabs ((x ^ 2) ^ n0) < (1 + x ^ 2) * eps
n:nat
Hn:(n >= N)%nat
H1:- x ^ 2 <> 1
a, b, c:R
bp:0 < b
h:a < b * c

0 < / b
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0
x_ub2':0 <= Rabs (x ^ 2) < 1
x_ub2:Rabs (x ^ 2) < 1
eps'_pos:(1 + x ^ 2) * eps > 0
N:nat
HN:forall n0 : nat, (n0 >= N)%nat -> Rabs ((x ^ 2) ^ n0) < (1 + x ^ 2) * eps
n:nat
Hn:(n >= N)%nat
H1:- x ^ 2 <> 1
a, b, c:R
bp:0 < b
h:a < b * c
a < b * c
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0
x_ub2':0 <= Rabs (x ^ 2) < 1
x_ub2:Rabs (x ^ 2) < 1
eps'_pos:(1 + x ^ 2) * eps > 0
N:nat
HN:forall n0 : nat, (n0 >= N)%nat -> Rabs ((x ^ 2) ^ n0) < (1 + x ^ 2) * eps
n:nat
Hn:(n >= N)%nat
H1:- x ^ 2 <> 1
a, b, c:R
bp:0 < b
h:a < b * c
b * c * / b = c
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0
x_ub2':0 <= Rabs (x ^ 2) < 1
x_ub2:Rabs (x ^ 2) < 1
eps'_pos:(1 + x ^ 2) * eps > 0
N:nat
HN:forall n0 : nat, (n0 >= N)%nat -> Rabs ((x ^ 2) ^ n0) < (1 + x ^ 2) * eps
n:nat
Hn:(n >= N)%nat
H1:- x ^ 2 <> 1
tool:forall a b c : R, 0 < b -> a < b * c -> a * / b < c
Rabs ((x ^ 2) ^ S n) * / (1 + x ^ 2) < eps
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0
x_ub2':0 <= Rabs (x ^ 2) < 1
x_ub2:Rabs (x ^ 2) < 1
eps'_pos:(1 + x ^ 2) * eps > 0
N:nat
HN:forall n0 : nat, (n0 >= N)%nat -> Rabs ((x ^ 2) ^ n0) < (1 + x ^ 2) * eps
n:nat
Hn:(n >= N)%nat
H1:- x ^ 2 <> 1
a, b, c:R
bp:0 < b
h:a < b * c

a < b * c
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0
x_ub2':0 <= Rabs (x ^ 2) < 1
x_ub2:Rabs (x ^ 2) < 1
eps'_pos:(1 + x ^ 2) * eps > 0
N:nat
HN:forall n0 : nat, (n0 >= N)%nat -> Rabs ((x ^ 2) ^ n0) < (1 + x ^ 2) * eps
n:nat
Hn:(n >= N)%nat
H1:- x ^ 2 <> 1
a, b, c:R
bp:0 < b
h:a < b * c
b * c * / b = c
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0
x_ub2':0 <= Rabs (x ^ 2) < 1
x_ub2:Rabs (x ^ 2) < 1
eps'_pos:(1 + x ^ 2) * eps > 0
N:nat
HN:forall n0 : nat, (n0 >= N)%nat -> Rabs ((x ^ 2) ^ n0) < (1 + x ^ 2) * eps
n:nat
Hn:(n >= N)%nat
H1:- x ^ 2 <> 1
tool:forall a b c : R, 0 < b -> a < b * c -> a * / b < c
Rabs ((x ^ 2) ^ S n) * / (1 + x ^ 2) < eps
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0
x_ub2':0 <= Rabs (x ^ 2) < 1
x_ub2:Rabs (x ^ 2) < 1
eps'_pos:(1 + x ^ 2) * eps > 0
N:nat
HN:forall n0 : nat, (n0 >= N)%nat -> Rabs ((x ^ 2) ^ n0) < (1 + x ^ 2) * eps
n:nat
Hn:(n >= N)%nat
H1:- x ^ 2 <> 1
a, b, c:R
bp:0 < b
h:a < b * c

b * c * / b = c
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0
x_ub2':0 <= Rabs (x ^ 2) < 1
x_ub2:Rabs (x ^ 2) < 1
eps'_pos:(1 + x ^ 2) * eps > 0
N:nat
HN:forall n0 : nat, (n0 >= N)%nat -> Rabs ((x ^ 2) ^ n0) < (1 + x ^ 2) * eps
n:nat
Hn:(n >= N)%nat
H1:- x ^ 2 <> 1
tool:forall a b c : R, 0 < b -> a < b * c -> a * / b < c
Rabs ((x ^ 2) ^ S n) * / (1 + x ^ 2) < eps
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0
x_ub2':0 <= Rabs (x ^ 2) < 1
x_ub2:Rabs (x ^ 2) < 1
eps'_pos:(1 + x ^ 2) * eps > 0
N:nat
HN:forall n0 : nat, (n0 >= N)%nat -> Rabs ((x ^ 2) ^ n0) < (1 + x ^ 2) * eps
n:nat
Hn:(n >= N)%nat
H1:- x ^ 2 <> 1
tool:forall a b c : R, 0 < b -> a < b * c -> a * / b < c

Rabs ((x ^ 2) ^ S n) * / (1 + x ^ 2) < eps
x:R
x_lb:-1 < x
x_ub:x < 1
eps:R
eps_pos:eps > 0
Tool0:0 <= x ^ 2
Tool1:0 < 1 + x ^ 2
Tool2:/ (1 + x ^ 2) > 0
x_ub2':0 <= Rabs (x ^ 2) < 1
x_ub2:Rabs (x ^ 2) < 1
eps'_pos:(1 + x ^ 2) * eps > 0
N:nat
HN:forall n0 : nat, (n0 >= N)%nat -> Rabs ((x ^ 2) ^ n0) < (1 + x ^ 2) * eps
n:nat
Hn:(n >= N)%nat
H1:- x ^ 2 <> 1
tool:forall a b c : R, 0 < b -> a < b * c -> a * / b < c

Rabs ((x ^ 2) ^ S n) < (1 + x ^ 2) * eps
apply HN; omega. Qed.

forall (c : R) (r : posreal), Rabs c + r < 1 -> CVU (fun (N : nat) (x : R) => sum_f_R0 (tg_alt (Datan_seq x)) N) (fun y : R => / (1 + y ^ 2)) c r

forall (c : R) (r : posreal), Rabs c + r < 1 -> CVU (fun (N : nat) (x : R) => sum_f_R0 (tg_alt (Datan_seq x)) N) (fun y : R => / (1 + y ^ 2)) c r
c:R
r:posreal
ub_ub:Rabs c + r < 1
eps:R
eps_pos:0 < eps

exists N : nat, forall (n : nat) (y : R), (N <= n)%nat -> Boule c r y -> Rabs (/ (1 + y ^ 2) - sum_f_R0 (tg_alt (Datan_seq y)) n) < eps
c:R
r:posreal
ub_ub:Rabs c + r < 1
eps:R
eps_pos:0 < eps

forall x : R, Boule c r x -> Un_decreasing (fun n : nat => Datan_seq x n)
c:R
r:posreal
ub_ub:Rabs c + r < 1
eps:R
eps_pos:0 < eps
forall x : R, Boule c r x -> Un_cv (fun n : nat => Datan_seq x n) 0
c:R
r:posreal
ub_ub:Rabs c + r < 1
eps:R
eps_pos:0 < eps
forall x : R, Boule c r x -> Un_cv (sum_f_R0 (tg_alt (fun i : nat => Datan_seq x i))) (/ (1 + x ^ 2))
c:R
r:posreal
ub_ub:Rabs c + r < 1
eps:R
eps_pos:0 < eps
forall (x : R) (n : nat), Boule c r x -> Datan_seq x n <= Datan_seq (Rabs c + r) n
c:R
r:posreal
ub_ub:Rabs c + r < 1
eps:R
eps_pos:0 < eps
Un_cv (Datan_seq (Rabs c + r)) 0
c:R
r:posreal
ub_ub:Rabs c + r < 1
eps:R
eps_pos:0 < eps
0 < eps
c:R
r:posreal
ub_ub:Rabs c + r < 1
eps:R
eps_pos:0 < eps

forall x : R, Boule c r x -> Un_cv (fun n : nat => Datan_seq x n) 0
c:R
r:posreal
ub_ub:Rabs c + r < 1
eps:R
eps_pos:0 < eps
forall x : R, Boule c r x -> Un_cv (sum_f_R0 (tg_alt (fun i : nat => Datan_seq x i))) (/ (1 + x ^ 2))
c:R
r:posreal
ub_ub:Rabs c + r < 1
eps:R
eps_pos:0 < eps
forall (x : R) (n : nat), Boule c r x -> Datan_seq x n <= Datan_seq (Rabs c + r) n
c:R
r:posreal
ub_ub:Rabs c + r < 1
eps:R
eps_pos:0 < eps
Un_cv (Datan_seq (Rabs c + r)) 0
c:R
r:posreal
ub_ub:Rabs c + r < 1
eps:R
eps_pos:0 < eps
0 < eps
c:R
r:posreal
ub_ub:Rabs c + r < 1
eps:R
eps_pos:0 < eps

forall x : R, Boule c r x -> Un_cv (sum_f_R0 (tg_alt (fun i : nat => Datan_seq x i))) (/ (1 + x ^ 2))
c:R
r:posreal
ub_ub:Rabs c + r < 1
eps:R
eps_pos:0 < eps
forall (x : R) (n : nat), Boule c r x -> Datan_seq x n <= Datan_seq (Rabs c + r) n
c:R
r:posreal
ub_ub:Rabs c + r < 1
eps:R
eps_pos:0 < eps
Un_cv (Datan_seq (Rabs c + r)) 0
c:R
r:posreal
ub_ub:Rabs c + r < 1
eps:R
eps_pos:0 < eps
0 < eps
c:R
r:posreal
ub_ub:Rabs c + r < 1
eps:R
eps_pos:0 < eps

forall (x : R) (n : nat), Boule c r x -> Datan_seq x n <= Datan_seq (Rabs c + r) n
c:R
r:posreal
ub_ub:Rabs c + r < 1
eps:R
eps_pos:0 < eps
Un_cv (Datan_seq (Rabs c + r)) 0
c:R
r:posreal
ub_ub:Rabs c + r < 1
eps:R
eps_pos:0 < eps
0 < eps
c:R
r:posreal
ub_ub:Rabs c + r < 1
eps:R
eps_pos:0 < eps
x:R
inb:Boule c r x

Datan_seq x 0 <= Datan_seq (Rabs c + r) 0
c:R
r:posreal
ub_ub:Rabs c + r < 1
eps:R
eps_pos:0 < eps
x:R
n:nat
inb:Boule c r x
Datan_seq x (S n) <= Datan_seq (Rabs c + r) (S n)
c:R
r:posreal
ub_ub:Rabs c + r < 1
eps:R
eps_pos:0 < eps
Un_cv (Datan_seq (Rabs c + r)) 0
c:R
r:posreal
ub_ub:Rabs c + r < 1
eps:R
eps_pos:0 < eps
0 < eps
c:R
r:posreal
ub_ub:Rabs c + r < 1
eps:R
eps_pos:0 < eps
x:R
n:nat
inb:Boule c r x

Datan_seq x (S n) <= Datan_seq (Rabs c + r) (S n)
c:R
r:posreal
ub_ub:Rabs c + r < 1
eps:R
eps_pos:0 < eps
Un_cv (Datan_seq (Rabs c + r)) 0
c:R
r:posreal
ub_ub:Rabs c + r < 1
eps:R
eps_pos:0 < eps
0 < eps
c:R
r:posreal
ub_ub:Rabs c + r < 1
eps:R
eps_pos:0 < eps
x:R
n:nat
inb:Boule c r x

(S n > 0)%nat
c:R
r:posreal
ub_ub:Rabs c + r < 1
eps:R
eps_pos:0 < eps
x:R
n:nat
inb:Boule c r x
0 <= Rabs x < Rabs c + r
c:R
r:posreal
ub_ub:Rabs c + r < 1
eps:R
eps_pos:0 < eps
Un_cv (Datan_seq (Rabs c + r)) 0
c:R
r:posreal
ub_ub:Rabs c + r < 1
eps:R
eps_pos:0 < eps
0 < eps
c:R
r:posreal
ub_ub:Rabs c + r < 1
eps:R
eps_pos:0 < eps
x:R
n:nat
inb:Boule c r x

0 <= Rabs x < Rabs c + r
c:R
r:posreal
ub_ub:Rabs c + r < 1
eps:R
eps_pos:0 < eps
Un_cv (Datan_seq (Rabs c + r)) 0
c:R
r:posreal
ub_ub:Rabs c + r < 1
eps:R
eps_pos:0 < eps
0 < eps
c:R
r:posreal
ub_ub:Rabs c + r < 1
eps:R
eps_pos:0 < eps
x:R
n:nat
inb:Rabs x < Rabs c + r

0 <= Rabs x
c:R
r:posreal
ub_ub:Rabs c + r < 1
eps:R
eps_pos:0 < eps
Un_cv (Datan_seq (Rabs c + r)) 0
c:R
r:posreal
ub_ub:Rabs c + r < 1
eps:R
eps_pos:0 < eps
0 < eps
c:R
r:posreal
ub_ub:Rabs c + r < 1
eps:R
eps_pos:0 < eps

Un_cv (Datan_seq (Rabs c + r)) 0
c:R
r:posreal
ub_ub:Rabs c + r < 1
eps:R
eps_pos:0 < eps
0 < eps
c:R
r:posreal
ub_ub:Rabs c + r < 1
eps:R
eps_pos:0 < eps

-1 < Rabs c + r
c:R
r:posreal
ub_ub:Rabs c + r < 1
eps:R
eps_pos:0 < eps
Rabs c + r < 1
c:R
r:posreal
ub_ub:Rabs c + r < 1
eps:R
eps_pos:0 < eps
0 < eps
c:R
r:posreal
ub_ub:Rabs c + r < 1
eps:R
eps_pos:0 < eps

0 < Rabs c + r
c:R
r:posreal
ub_ub:Rabs c + r < 1
eps:R
eps_pos:0 < eps
Rabs c + r < 1
c:R
r:posreal
ub_ub:Rabs c + r < 1
eps:R
eps_pos:0 < eps
0 < eps
c:R
r:posreal
ub_ub:Rabs c + r < 1
eps:R
eps_pos:0 < eps

0 <= Rabs c
c:R
r:posreal
ub_ub:Rabs c + r < 1
eps:R
eps_pos:0 < eps
0 < r
c:R
r:posreal
ub_ub:Rabs c + r < 1
eps:R
eps_pos:0 < eps
Rabs c + r < 1
c:R
r:posreal
ub_ub:Rabs c + r < 1
eps:R
eps_pos:0 < eps
0 < eps
c:R
r:posreal
ub_ub:Rabs c + r < 1
eps:R
eps_pos:0 < eps

0 < r
c:R
r:posreal
ub_ub:Rabs c + r < 1
eps:R
eps_pos:0 < eps
Rabs c + r < 1
c:R
r:posreal
ub_ub:Rabs c + r < 1
eps:R
eps_pos:0 < eps
0 < eps
c:R
r:posreal
ub_ub:Rabs c + r < 1
eps:R
eps_pos:0 < eps

Rabs c + r < 1
c:R
r:posreal
ub_ub:Rabs c + r < 1
eps:R
eps_pos:0 < eps
0 < eps
c:R
r:posreal
ub_ub:Rabs c + r < 1
eps:R
eps_pos:0 < eps

0 < eps
assumption. Qed.

forall (N : nat) (x : R), -1 <= x -> x < 1 -> derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)

forall (N : nat) (x : R), -1 <= x -> x < 1 -> derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)

forall N : nat, (-1) ^ S (2 * N) = -1
Tool:forall N : nat, (-1) ^ S (2 * N) = -1
forall (N : nat) (x : R), -1 <= x -> x < 1 -> derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)

(-1) ^ S (2 * 0) = -1
n:nat
IHn:(-1) ^ S (2 * n) = -1
(-1) ^ S (2 * S n) = -1
Tool:forall N : nat, (-1) ^ S (2 * N) = -1
forall (N : nat) (x : R), -1 <= x -> x < 1 -> derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
n:nat
IHn:(-1) ^ S (2 * n) = -1

(-1) ^ S (2 * S n) = -1
Tool:forall N : nat, (-1) ^ S (2 * N) = -1
forall (N : nat) (x : R), -1 <= x -> x < 1 -> derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
n:nat
IHn:(-1) ^ S (2 * n) = -1

(-1) ^ 2 * (-1) ^ S (2 * n) = -1
n:nat
IHn:(-1) ^ S (2 * n) = -1
(-1) ^ 2 * (-1) ^ S (2 * n) = (-1) ^ S (2 * S n)
Tool:forall N : nat, (-1) ^ S (2 * N) = -1
forall (N : nat) (x : R), -1 <= x -> x < 1 -> derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
n:nat
IHn:(-1) ^ S (2 * n) = -1

(-1) ^ 2 * (-1) ^ S (2 * n) = (-1) ^ S (2 * S n)
Tool:forall N : nat, (-1) ^ S (2 * N) = -1
forall (N : nat) (x : R), -1 <= x -> x < 1 -> derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
n:nat
IHn:(-1) ^ S (2 * n) = -1

(-1) ^ (2 + S (2 * n)) = (-1) ^ S (2 * S n)
Tool:forall N : nat, (-1) ^ S (2 * N) = -1
forall (N : nat) (x : R), -1 <= x -> x < 1 -> derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
n:nat
IHn:(-1) ^ S (2 * n) = -1

(-1) ^ S (2 * S n) = (-1) ^ S (2 * S n)
n:nat
IHn:(-1) ^ S (2 * n) = -1
S (2 * S n) = (2 + S (2 * n))%nat
Tool:forall N : nat, (-1) ^ S (2 * N) = -1
forall (N : nat) (x : R), -1 <= x -> x < 1 -> derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
n:nat
IHn:(-1) ^ S (2 * n) = -1

S (2 * S n) = (2 + S (2 * n))%nat
Tool:forall N : nat, (-1) ^ S (2 * N) = -1
forall (N : nat) (x : R), -1 <= x -> x < 1 -> derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
Tool:forall N : nat, (-1) ^ S (2 * N) = -1

forall (N : nat) (x : R), -1 <= x -> x < 1 -> derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1

derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
Tool:forall N : nat, (-1) ^ S (2 * N) = -1
x:R
x_lb:-1 <= x
x_ub:x < 1

derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) 0) x (sum_f_R0 (tg_alt (Datan_seq x)) 0)
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
IHN:derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) (S N)) x (sum_f_R0 (tg_alt (Datan_seq x)) (S N))
Tool:forall N : nat, (-1) ^ S (2 * N) = -1
x:R
x_lb:-1 <= x
x_ub:x < 1

derivable_pt_lim (fun x0 : R => 1 * (x0 * 1 / 1)) x (1 * 1)
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
IHN:derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) (S N)) x (sum_f_R0 (tg_alt (Datan_seq x)) (S N))
Tool:forall N : nat, (-1) ^ S (2 * N) = -1
x:R
x_lb:-1 <= x
x_ub:x < 1
eps:R
eps_pos:0 < eps

exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((1 * ((x + h) * 1 / 1) - 1 * (x * 1 / 1)) / h - 1 * 1) < eps
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
IHN:derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) (S N)) x (sum_f_R0 (tg_alt (Datan_seq x)) (S N))
Tool:forall N : nat, (-1) ^ S (2 * N) = -1
x:R
x_lb:-1 <= x
x_ub:x < 1
eps:R
eps_pos:0 < eps
delta:posreal
Hdelta:forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((id (x + h) - id x) / h - 1) < eps

forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((1 * ((x + h) * 1 / 1) - 1 * (x * 1 / 1)) / h - 1 * 1) < eps
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
IHN:derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) (S N)) x (sum_f_R0 (tg_alt (Datan_seq x)) (S N))
Tool:forall N : nat, (-1) ^ S (2 * N) = -1
x:R
x_lb:-1 <= x
x_ub:x < 1
eps:R
eps_pos:0 < eps
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((id (x + h0) - id x) / h0 - 1) < eps
h:R
hneq:h <> 0
h_b:Rabs h < delta

Rabs ((1 * ((x + h) * 1 / 1) - 1 * (x * 1 / 1)) / h - 1 * 1) < eps
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
IHN:derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) (S N)) x (sum_f_R0 (tg_alt (Datan_seq x)) (S N))
Tool:forall N : nat, (-1) ^ S (2 * N) = -1
x:R
x_lb:-1 <= x
x_ub:x < 1
eps:R
eps_pos:0 < eps
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((id (x + h0) - id x) / h0 - 1) < eps
h:R
hneq:h <> 0
h_b:Rabs h < delta

Rabs ((id (x + h) - id x) / h - 1 * 1) < eps
Tool:forall N : nat, (-1) ^ S (2 * N) = -1
x:R
x_lb:-1 <= x
x_ub:x < 1
eps:R
eps_pos:0 < eps
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((id (x + h0) - id x) / h0 - 1) < eps
h:R
hneq:h <> 0
h_b:Rabs h < delta
id (x + h) - id x = 1 * ((x + h) * 1 / 1) - 1 * (x * 1 / 1)
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
IHN:derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) (S N)) x (sum_f_R0 (tg_alt (Datan_seq x)) (S N))
Tool:forall N : nat, (-1) ^ S (2 * N) = -1
x:R
x_lb:-1 <= x
x_ub:x < 1
eps:R
eps_pos:0 < eps
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((id (x + h0) - id x) / h0 - 1) < eps
h:R
hneq:h <> 0
h_b:Rabs h < delta

Rabs ((id (x + h) - id x) / h - 1) < eps
Tool:forall N : nat, (-1) ^ S (2 * N) = -1
x:R
x_lb:-1 <= x
x_ub:x < 1
eps:R
eps_pos:0 < eps
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((id (x + h0) - id x) / h0 - 1) < eps
h:R
hneq:h <> 0
h_b:Rabs h < delta
id (x + h) - id x = 1 * ((x + h) * 1 / 1) - 1 * (x * 1 / 1)
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
IHN:derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) (S N)) x (sum_f_R0 (tg_alt (Datan_seq x)) (S N))
Tool:forall N : nat, (-1) ^ S (2 * N) = -1
x:R
x_lb:-1 <= x
x_ub:x < 1
eps:R
eps_pos:0 < eps
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs ((id (x + h0) - id x) / h0 - 1) < eps
h:R
hneq:h <> 0
h_b:Rabs h < delta

id (x + h) - id x = 1 * ((x + h) * 1 / 1) - 1 * (x * 1 / 1)
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
IHN:derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) (S N)) x (sum_f_R0 (tg_alt (Datan_seq x)) (S N))
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
IHN:derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)

derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) (S N)) x (sum_f_R0 (tg_alt (Datan_seq x)) (S N))
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
IHN:derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
eps:R
eps_pos:0 < eps

exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) (S N) - sum_f_R0 (tg_alt (Ratan_seq x)) (S N)) / h - sum_f_R0 (tg_alt (Datan_seq x)) (S N)) < eps
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
IHN:derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0

exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) (S N) - sum_f_R0 (tg_alt (Ratan_seq x)) (S N)) / h - sum_f_R0 (tg_alt (Datan_seq x)) (S N)) < eps
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
IHN:derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:posreal
Hdelta1:forall h : R, h <> 0 -> Rabs h < delta1 -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3

exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) (S N) - sum_f_R0 (tg_alt (Ratan_seq x)) (S N)) / h - sum_f_R0 (tg_alt (Datan_seq x)) (S N)) < eps
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
IHN:derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:posreal
Hdelta1:forall h : R, h <> 0 -> Rabs h < delta1 -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3

derivable_pt_lim (fun x0 : R => tg_alt (Ratan_seq x0) (S N)) x (tg_alt (Datan_seq x) (S N))
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
IHN:derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:posreal
Hdelta1:forall h : R, h <> 0 -> Rabs h < delta1 -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
Main:derivable_pt_lim (fun x0 : R => tg_alt (Ratan_seq x0) (S N)) x (tg_alt (Datan_seq x) (S N))
exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) (S N) - sum_f_R0 (tg_alt (Ratan_seq x)) (S N)) / h - sum_f_R0 (tg_alt (Datan_seq x)) (S N)) < eps
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'

exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((tg_alt (Ratan_seq (x + h)) (S N) - tg_alt (Ratan_seq x) (S N)) / h - tg_alt (Datan_seq x) (S N)) < eps'
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
IHN:derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:posreal
Hdelta1:forall h : R, h <> 0 -> Rabs h < delta1 -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
Main:derivable_pt_lim (fun x0 : R => tg_alt (Ratan_seq x0) (S N)) x (tg_alt (Datan_seq x) (S N))
exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) (S N) - sum_f_R0 (tg_alt (Ratan_seq x)) (S N)) / h - sum_f_R0 (tg_alt (Datan_seq x)) (S N)) < eps
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h : R, h <> 0 -> Rabs h < delta -> Rabs (((x + h) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'

forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((tg_alt (Ratan_seq (x + h)) (S N) - tg_alt (Ratan_seq x) (S N)) / h - tg_alt (Datan_seq x) (S N)) < eps'
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
IHN:derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:posreal
Hdelta1:forall h : R, h <> 0 -> Rabs h < delta1 -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
Main:derivable_pt_lim (fun x0 : R => tg_alt (Ratan_seq x0) (S N)) x (tg_alt (Datan_seq x) (S N))
exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) (S N) - sum_f_R0 (tg_alt (Ratan_seq x)) (S N)) / h - sum_f_R0 (tg_alt (Datan_seq x)) (S N)) < eps
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta

Rabs (((-1) ^ S N * ((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1)) - (-1) ^ S N * (x ^ (2 * S N + 1) / INR (2 * S N + 1))) / h - (-1) ^ S N * x ^ (2 * S N)) < eps'
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
IHN:derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:posreal
Hdelta1:forall h : R, h <> 0 -> Rabs h < delta1 -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
Main:derivable_pt_lim (fun x0 : R => tg_alt (Ratan_seq x0) (S N)) x (tg_alt (Datan_seq x) (S N))
exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) (S N) - sum_f_R0 (tg_alt (Ratan_seq x)) (S N)) / h - sum_f_R0 (tg_alt (Datan_seq x)) (S N)) < eps
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta

Rabs ((-1) ^ S N * (((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1) - x ^ (2 * S N + 1) / INR (2 * S N + 1)) / h - x ^ (2 * S N))) < eps'
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
(-1) ^ S N * (((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1) - x ^ (2 * S N + 1) / INR (2 * S N + 1)) / h - x ^ (2 * S N)) = ((-1) ^ S N * ((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1)) - (-1) ^ S N * (x ^ (2 * S N + 1) / INR (2 * S N + 1))) / h - (-1) ^ S N * x ^ (2 * S N)
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
IHN:derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:posreal
Hdelta1:forall h : R, h <> 0 -> Rabs h < delta1 -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
Main:derivable_pt_lim (fun x0 : R => tg_alt (Ratan_seq x0) (S N)) x (tg_alt (Datan_seq x) (S N))
exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) (S N) - sum_f_R0 (tg_alt (Ratan_seq x)) (S N)) / h - sum_f_R0 (tg_alt (Datan_seq x)) (S N)) < eps
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta

Rabs (((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1) - x ^ (2 * S N + 1) / INR (2 * S N + 1)) / h - x ^ (2 * S N)) < eps'
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
(-1) ^ S N * (((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1) - x ^ (2 * S N + 1) / INR (2 * S N + 1)) / h - x ^ (2 * S N)) = ((-1) ^ S N * ((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1)) - (-1) ^ S N * (x ^ (2 * S N + 1) / INR (2 * S N + 1))) / h - (-1) ^ S N * x ^ (2 * S N)
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
IHN:derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:posreal
Hdelta1:forall h : R, h <> 0 -> Rabs h < delta1 -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
Main:derivable_pt_lim (fun x0 : R => tg_alt (Ratan_seq x0) (S N)) x (tg_alt (Datan_seq x) (S N))
exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) (S N) - sum_f_R0 (tg_alt (Ratan_seq x)) (S N)) / h - sum_f_R0 (tg_alt (Datan_seq x)) (S N)) < eps
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta

Rabs (/ INR (2 * S N + 1) * (((x + h) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1))) < eps'
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
/ INR (2 * S N + 1) * (((x + h) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) = ((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1) - x ^ (2 * S N + 1) / INR (2 * S N + 1)) / h - x ^ (2 * S N)
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
(-1) ^ S N * (((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1) - x ^ (2 * S N + 1) / INR (2 * S N + 1)) / h - x ^ (2 * S N)) = ((-1) ^ S N * ((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1)) - (-1) ^ S N * (x ^ (2 * S N + 1) / INR (2 * S N + 1))) / h - (-1) ^ S N * x ^ (2 * S N)
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
IHN:derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:posreal
Hdelta1:forall h : R, h <> 0 -> Rabs h < delta1 -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
Main:derivable_pt_lim (fun x0 : R => tg_alt (Ratan_seq x0) (S N)) x (tg_alt (Datan_seq x) (S N))
exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) (S N) - sum_f_R0 (tg_alt (Ratan_seq x)) (S N)) / h - sum_f_R0 (tg_alt (Datan_seq x)) (S N)) < eps
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta

Rabs (/ INR (2 * S N + 1)) * Rabs (((x + h) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
/ INR (2 * S N + 1) * (((x + h) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) = ((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1) - x ^ (2 * S N + 1) / INR (2 * S N + 1)) / h - x ^ (2 * S N)
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
(-1) ^ S N * (((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1) - x ^ (2 * S N + 1) / INR (2 * S N + 1)) / h - x ^ (2 * S N)) = ((-1) ^ S N * ((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1)) - (-1) ^ S N * (x ^ (2 * S N + 1) / INR (2 * S N + 1))) / h - (-1) ^ S N * x ^ (2 * S N)
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
IHN:derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:posreal
Hdelta1:forall h : R, h <> 0 -> Rabs h < delta1 -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
Main:derivable_pt_lim (fun x0 : R => tg_alt (Ratan_seq x0) (S N)) x (tg_alt (Datan_seq x) (S N))
exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) (S N) - sum_f_R0 (tg_alt (Ratan_seq x)) (S N)) / h - sum_f_R0 (tg_alt (Datan_seq x)) (S N)) < eps
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
Heq:((x + h) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1) = 0

Rabs (/ INR (2 * S N + 1)) * Rabs (((x + h) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
Heq:((x + h) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1) <> 0
Rabs (/ INR (2 * S N + 1)) * Rabs (((x + h) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
/ INR (2 * S N + 1) * (((x + h) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) = ((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1) - x ^ (2 * S N + 1) / INR (2 * S N + 1)) / h - x ^ (2 * S N)
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
(-1) ^ S N * (((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1) - x ^ (2 * S N + 1) / INR (2 * S N + 1)) / h - x ^ (2 * S N)) = ((-1) ^ S N * ((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1)) - (-1) ^ S N * (x ^ (2 * S N + 1) / INR (2 * S N + 1))) / h - (-1) ^ S N * x ^ (2 * S N)
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
IHN:derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:posreal
Hdelta1:forall h : R, h <> 0 -> Rabs h < delta1 -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
Main:derivable_pt_lim (fun x0 : R => tg_alt (Ratan_seq x0) (S N)) x (tg_alt (Datan_seq x) (S N))
exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) (S N) - sum_f_R0 (tg_alt (Ratan_seq x)) (S N)) / h - sum_f_R0 (tg_alt (Datan_seq x)) (S N)) < eps
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
Heq:((x + h) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1) <> 0

Rabs (/ INR (2 * S N + 1)) * Rabs (((x + h) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
/ INR (2 * S N + 1) * (((x + h) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) = ((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1) - x ^ (2 * S N + 1) / INR (2 * S N + 1)) / h - x ^ (2 * S N)
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
(-1) ^ S N * (((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1) - x ^ (2 * S N + 1) / INR (2 * S N + 1)) / h - x ^ (2 * S N)) = ((-1) ^ S N * ((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1)) - (-1) ^ S N * (x ^ (2 * S N + 1) / INR (2 * S N + 1))) / h - (-1) ^ S N * x ^ (2 * S N)
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
IHN:derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:posreal
Hdelta1:forall h : R, h <> 0 -> Rabs h < delta1 -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
Main:derivable_pt_lim (fun x0 : R => tg_alt (Ratan_seq x0) (S N)) x (tg_alt (Datan_seq x) (S N))
exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) (S N) - sum_f_R0 (tg_alt (Ratan_seq x)) (S N)) / h - sum_f_R0 (tg_alt (Datan_seq x)) (S N)) < eps
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
Heq:((x + h) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1) <> 0

Rabs (/ INR (2 * S N + 1)) * Rabs (((x + h) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < Rabs (((x + h) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1))
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
Heq:((x + h) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1) <> 0
Rabs (((x + h) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
/ INR (2 * S N + 1) * (((x + h) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) = ((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1) - x ^ (2 * S N + 1) / INR (2 * S N + 1)) / h - x ^ (2 * S N)
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
(-1) ^ S N * (((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1) - x ^ (2 * S N + 1) / INR (2 * S N + 1)) / h - x ^ (2 * S N)) = ((-1) ^ S N * ((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1)) - (-1) ^ S N * (x ^ (2 * S N + 1) / INR (2 * S N + 1))) / h - (-1) ^ S N * x ^ (2 * S N)
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
IHN:derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:posreal
Hdelta1:forall h : R, h <> 0 -> Rabs h < delta1 -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
Main:derivable_pt_lim (fun x0 : R => tg_alt (Ratan_seq x0) (S N)) x (tg_alt (Datan_seq x) (S N))
exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) (S N) - sum_f_R0 (tg_alt (Ratan_seq x)) (S N)) / h - sum_f_R0 (tg_alt (Datan_seq x)) (S N)) < eps
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
Heq:((x + h) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1) <> 0

0 < Rabs (((x + h) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1))
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
Heq:((x + h) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1) <> 0
Rabs (/ INR (2 * S N + 1)) < 1
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
Heq:((x + h) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1) <> 0
Rabs (((x + h) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
/ INR (2 * S N + 1) * (((x + h) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) = ((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1) - x ^ (2 * S N + 1) / INR (2 * S N + 1)) / h - x ^ (2 * S N)
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
(-1) ^ S N * (((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1) - x ^ (2 * S N + 1) / INR (2 * S N + 1)) / h - x ^ (2 * S N)) = ((-1) ^ S N * ((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1)) - (-1) ^ S N * (x ^ (2 * S N + 1) / INR (2 * S N + 1))) / h - (-1) ^ S N * x ^ (2 * S N)
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
IHN:derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:posreal
Hdelta1:forall h : R, h <> 0 -> Rabs h < delta1 -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
Main:derivable_pt_lim (fun x0 : R => tg_alt (Ratan_seq x0) (S N)) x (tg_alt (Datan_seq x) (S N))
exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) (S N) - sum_f_R0 (tg_alt (Ratan_seq x)) (S N)) / h - sum_f_R0 (tg_alt (Datan_seq x)) (S N)) < eps
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
Heq:((x + h) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1) <> 0

Rabs (/ INR (2 * S N + 1)) < 1
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
Heq:((x + h) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1) <> 0
Rabs (((x + h) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
/ INR (2 * S N + 1) * (((x + h) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) = ((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1) - x ^ (2 * S N + 1) / INR (2 * S N + 1)) / h - x ^ (2 * S N)
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
(-1) ^ S N * (((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1) - x ^ (2 * S N + 1) / INR (2 * S N + 1)) / h - x ^ (2 * S N)) = ((-1) ^ S N * ((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1)) - (-1) ^ S N * (x ^ (2 * S N + 1) / INR (2 * S N + 1))) / h - (-1) ^ S N * x ^ (2 * S N)
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
IHN:derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:posreal
Hdelta1:forall h : R, h <> 0 -> Rabs h < delta1 -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
Main:derivable_pt_lim (fun x0 : R => tg_alt (Ratan_seq x0) (S N)) x (tg_alt (Datan_seq x) (S N))
exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) (S N) - sum_f_R0 (tg_alt (Ratan_seq x)) (S N)) / h - sum_f_R0 (tg_alt (Datan_seq x)) (S N)) < eps
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
Heq:((x + h) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1) <> 0

/ INR (2 * S N + 1) < 1
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
Heq:((x + h) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1) <> 0
/ INR (2 * S N + 1) >= 0
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
Heq:((x + h) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1) <> 0
Rabs (((x + h) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
/ INR (2 * S N + 1) * (((x + h) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) = ((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1) - x ^ (2 * S N + 1) / INR (2 * S N + 1)) / h - x ^ (2 * S N)
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
(-1) ^ S N * (((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1) - x ^ (2 * S N + 1) / INR (2 * S N + 1)) / h - x ^ (2 * S N)) = ((-1) ^ S N * ((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1)) - (-1) ^ S N * (x ^ (2 * S N + 1) / INR (2 * S N + 1))) / h - (-1) ^ S N * x ^ (2 * S N)
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
IHN:derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:posreal
Hdelta1:forall h : R, h <> 0 -> Rabs h < delta1 -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
Main:derivable_pt_lim (fun x0 : R => tg_alt (Ratan_seq x0) (S N)) x (tg_alt (Datan_seq x) (S N))
exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) (S N) - sum_f_R0 (tg_alt (Ratan_seq x)) (S N)) / h - sum_f_R0 (tg_alt (Datan_seq x)) (S N)) < eps
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
Heq:((x + h) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1) <> 0

/ INR (2 * S N + 1) < / 1
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
Heq:((x + h) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1) <> 0
/ INR (2 * S N + 1) >= 0
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
Heq:((x + h) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1) <> 0
Rabs (((x + h) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
/ INR (2 * S N + 1) * (((x + h) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) = ((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1) - x ^ (2 * S N + 1) / INR (2 * S N + 1)) / h - x ^ (2 * S N)
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
(-1) ^ S N * (((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1) - x ^ (2 * S N + 1) / INR (2 * S N + 1)) / h - x ^ (2 * S N)) = ((-1) ^ S N * ((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1)) - (-1) ^ S N * (x ^ (2 * S N + 1) / INR (2 * S N + 1))) / h - (-1) ^ S N * x ^ (2 * S N)
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
IHN:derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:posreal
Hdelta1:forall h : R, h <> 0 -> Rabs h < delta1 -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
Main:derivable_pt_lim (fun x0 : R => tg_alt (Ratan_seq x0) (S N)) x (tg_alt (Datan_seq x) (S N))
exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) (S N) - sum_f_R0 (tg_alt (Ratan_seq x)) (S N)) / h - sum_f_R0 (tg_alt (Datan_seq x)) (S N)) < eps
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
Heq:((x + h) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1) <> 0

/ INR (2 * S N + 1) >= 0
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
Heq:((x + h) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1) <> 0
Rabs (((x + h) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
/ INR (2 * S N + 1) * (((x + h) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) = ((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1) - x ^ (2 * S N + 1) / INR (2 * S N + 1)) / h - x ^ (2 * S N)
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
(-1) ^ S N * (((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1) - x ^ (2 * S N + 1) / INR (2 * S N + 1)) / h - x ^ (2 * S N)) = ((-1) ^ S N * ((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1)) - (-1) ^ S N * (x ^ (2 * S N + 1) / INR (2 * S N + 1))) / h - (-1) ^ S N * x ^ (2 * S N)
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
IHN:derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:posreal
Hdelta1:forall h : R, h <> 0 -> Rabs h < delta1 -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
Main:derivable_pt_lim (fun x0 : R => tg_alt (Ratan_seq x0) (S N)) x (tg_alt (Datan_seq x) (S N))
exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) (S N) - sum_f_R0 (tg_alt (Ratan_seq x)) (S N)) / h - sum_f_R0 (tg_alt (Datan_seq x)) (S N)) < eps
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
Heq:((x + h) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1) <> 0

INR (2 * S N) + 1 = INR (2 * S N + 1)
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
Heq:((x + h) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1) <> 0
Rabs (((x + h) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
/ INR (2 * S N + 1) * (((x + h) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) = ((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1) - x ^ (2 * S N + 1) / INR (2 * S N + 1)) / h - x ^ (2 * S N)
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
(-1) ^ S N * (((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1) - x ^ (2 * S N + 1) / INR (2 * S N + 1)) / h - x ^ (2 * S N)) = ((-1) ^ S N * ((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1)) - (-1) ^ S N * (x ^ (2 * S N + 1) / INR (2 * S N + 1))) / h - (-1) ^ S N * x ^ (2 * S N)
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
IHN:derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:posreal
Hdelta1:forall h : R, h <> 0 -> Rabs h < delta1 -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
Main:derivable_pt_lim (fun x0 : R => tg_alt (Ratan_seq x0) (S N)) x (tg_alt (Datan_seq x) (S N))
exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) (S N) - sum_f_R0 (tg_alt (Ratan_seq x)) (S N)) / h - sum_f_R0 (tg_alt (Datan_seq x)) (S N)) < eps
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
Heq:((x + h) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1) <> 0

Rabs (((x + h) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
/ INR (2 * S N + 1) * (((x + h) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) = ((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1) - x ^ (2 * S N + 1) / INR (2 * S N + 1)) / h - x ^ (2 * S N)
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
(-1) ^ S N * (((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1) - x ^ (2 * S N + 1) / INR (2 * S N + 1)) / h - x ^ (2 * S N)) = ((-1) ^ S N * ((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1)) - (-1) ^ S N * (x ^ (2 * S N + 1) / INR (2 * S N + 1))) / h - (-1) ^ S N * x ^ (2 * S N)
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
IHN:derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:posreal
Hdelta1:forall h : R, h <> 0 -> Rabs h < delta1 -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
Main:derivable_pt_lim (fun x0 : R => tg_alt (Ratan_seq x0) (S N)) x (tg_alt (Datan_seq x) (S N))
exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) (S N) - sum_f_R0 (tg_alt (Ratan_seq x)) (S N)) / h - sum_f_R0 (tg_alt (Datan_seq x)) (S N)) < eps
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta

/ INR (2 * S N + 1) * (((x + h) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) = ((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1) - x ^ (2 * S N + 1) / INR (2 * S N + 1)) / h - x ^ (2 * S N)
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
(-1) ^ S N * (((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1) - x ^ (2 * S N + 1) / INR (2 * S N + 1)) / h - x ^ (2 * S N)) = ((-1) ^ S N * ((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1)) - (-1) ^ S N * (x ^ (2 * S N + 1) / INR (2 * S N + 1))) / h - (-1) ^ S N * x ^ (2 * S N)
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
IHN:derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:posreal
Hdelta1:forall h : R, h <> 0 -> Rabs h < delta1 -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
Main:derivable_pt_lim (fun x0 : R => tg_alt (Ratan_seq x0) (S N)) x (tg_alt (Datan_seq x) (S N))
exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) (S N) - sum_f_R0 (tg_alt (Ratan_seq x)) (S N)) / h - sum_f_R0 (tg_alt (Datan_seq x)) (S N)) < eps
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta

/ INR (2 * S N + 1) * (((x + h) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h) - / INR (2 * S N + 1) * (INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) = ((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1) - x ^ (2 * S N + 1) / INR (2 * S N + 1)) / h - x ^ (2 * S N)
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
(-1) ^ S N * (((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1) - x ^ (2 * S N + 1) / INR (2 * S N + 1)) / h - x ^ (2 * S N)) = ((-1) ^ S N * ((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1)) - (-1) ^ S N * (x ^ (2 * S N + 1) / INR (2 * S N + 1))) / h - (-1) ^ S N * x ^ (2 * S N)
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
IHN:derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:posreal
Hdelta1:forall h : R, h <> 0 -> Rabs h < delta1 -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
Main:derivable_pt_lim (fun x0 : R => tg_alt (Ratan_seq x0) (S N)) x (tg_alt (Datan_seq x) (S N))
exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) (S N) - sum_f_R0 (tg_alt (Ratan_seq x)) (S N)) / h - sum_f_R0 (tg_alt (Datan_seq x)) (S N)) < eps
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta

/ INR (2 * S N + 1) * (((x + h) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h) - x ^ (2 * S N) = ((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1) - x ^ (2 * S N + 1) / INR (2 * S N + 1)) / h - x ^ (2 * S N)
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
x ^ (2 * S N) = / INR (2 * S N + 1) * (INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1))
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
(-1) ^ S N * (((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1) - x ^ (2 * S N + 1) / INR (2 * S N + 1)) / h - x ^ (2 * S N)) = ((-1) ^ S N * ((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1)) - (-1) ^ S N * (x ^ (2 * S N + 1) / INR (2 * S N + 1))) / h - (-1) ^ S N * x ^ (2 * S N)
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
IHN:derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:posreal
Hdelta1:forall h : R, h <> 0 -> Rabs h < delta1 -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
Main:derivable_pt_lim (fun x0 : R => tg_alt (Ratan_seq x0) (S N)) x (tg_alt (Datan_seq x) (S N))
exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) (S N) - sum_f_R0 (tg_alt (Ratan_seq x)) (S N)) / h - sum_f_R0 (tg_alt (Datan_seq x)) (S N)) < eps
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta

- x ^ (2 * S N) + / INR (2 * S N + 1) * (((x + h) ^ (2 * S N + 1) + - x ^ (2 * S N + 1)) / h) = ((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1) + - (x ^ (2 * S N + 1) / INR (2 * S N + 1))) / h + - x ^ (2 * S N)
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
x ^ (2 * S N) = / INR (2 * S N + 1) * (INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1))
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
(-1) ^ S N * (((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1) - x ^ (2 * S N + 1) / INR (2 * S N + 1)) / h - x ^ (2 * S N)) = ((-1) ^ S N * ((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1)) - (-1) ^ S N * (x ^ (2 * S N + 1) / INR (2 * S N + 1))) / h - (-1) ^ S N * x ^ (2 * S N)
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
IHN:derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:posreal
Hdelta1:forall h : R, h <> 0 -> Rabs h < delta1 -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
Main:derivable_pt_lim (fun x0 : R => tg_alt (Ratan_seq x0) (S N)) x (tg_alt (Datan_seq x) (S N))
exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) (S N) - sum_f_R0 (tg_alt (Ratan_seq x)) (S N)) / h - sum_f_R0 (tg_alt (Datan_seq x)) (S N)) < eps
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta

- x ^ (2 * S N) + / INR (2 * S N + 1) * (((x + h) ^ (2 * S N + 1) + - x ^ (2 * S N + 1)) / h) = - x ^ (2 * S N) + ((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1) + - (x ^ (2 * S N + 1) / INR (2 * S N + 1))) / h
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
x ^ (2 * S N) = / INR (2 * S N + 1) * (INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1))
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
(-1) ^ S N * (((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1) - x ^ (2 * S N + 1) / INR (2 * S N + 1)) / h - x ^ (2 * S N)) = ((-1) ^ S N * ((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1)) - (-1) ^ S N * (x ^ (2 * S N + 1) / INR (2 * S N + 1))) / h - (-1) ^ S N * x ^ (2 * S N)
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
IHN:derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:posreal
Hdelta1:forall h : R, h <> 0 -> Rabs h < delta1 -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
Main:derivable_pt_lim (fun x0 : R => tg_alt (Ratan_seq x0) (S N)) x (tg_alt (Datan_seq x) (S N))
exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) (S N) - sum_f_R0 (tg_alt (Ratan_seq x)) (S N)) / h - sum_f_R0 (tg_alt (Datan_seq x)) (S N)) < eps
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta

/ INR (2 * S N + 1) * (((x + h) ^ (2 * S N + 1) + - x ^ (2 * S N + 1)) / h) = ((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1) + - (x ^ (2 * S N + 1) / INR (2 * S N + 1))) / h
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
x ^ (2 * S N) = / INR (2 * S N + 1) * (INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1))
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
(-1) ^ S N * (((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1) - x ^ (2 * S N + 1) / INR (2 * S N + 1)) / h - x ^ (2 * S N)) = ((-1) ^ S N * ((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1)) - (-1) ^ S N * (x ^ (2 * S N + 1) / INR (2 * S N + 1))) / h - (-1) ^ S N * x ^ (2 * S N)
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
IHN:derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:posreal
Hdelta1:forall h : R, h <> 0 -> Rabs h < delta1 -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
Main:derivable_pt_lim (fun x0 : R => tg_alt (Ratan_seq x0) (S N)) x (tg_alt (Datan_seq x) (S N))
exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) (S N) - sum_f_R0 (tg_alt (Ratan_seq x)) (S N)) / h - sum_f_R0 (tg_alt (Datan_seq x)) (S N)) < eps
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta

INR (2 * S N + 1) <> 0 /\ h <> 0
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
x ^ (2 * S N) = / INR (2 * S N + 1) * (INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1))
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
(-1) ^ S N * (((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1) - x ^ (2 * S N + 1) / INR (2 * S N + 1)) / h - x ^ (2 * S N)) = ((-1) ^ S N * ((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1)) - (-1) ^ S N * (x ^ (2 * S N + 1) / INR (2 * S N + 1))) / h - (-1) ^ S N * x ^ (2 * S N)
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
IHN:derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:posreal
Hdelta1:forall h : R, h <> 0 -> Rabs h < delta1 -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
Main:derivable_pt_lim (fun x0 : R => tg_alt (Ratan_seq x0) (S N)) x (tg_alt (Datan_seq x) (S N))
exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) (S N) - sum_f_R0 (tg_alt (Ratan_seq x)) (S N)) / h - sum_f_R0 (tg_alt (Datan_seq x)) (S N)) < eps
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta

x ^ (2 * S N) = / INR (2 * S N + 1) * (INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1))
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
(-1) ^ S N * (((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1) - x ^ (2 * S N + 1) / INR (2 * S N + 1)) / h - x ^ (2 * S N)) = ((-1) ^ S N * ((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1)) - (-1) ^ S N * (x ^ (2 * S N + 1) / INR (2 * S N + 1))) / h - (-1) ^ S N * x ^ (2 * S N)
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
IHN:derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:posreal
Hdelta1:forall h : R, h <> 0 -> Rabs h < delta1 -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
Main:derivable_pt_lim (fun x0 : R => tg_alt (Ratan_seq x0) (S N)) x (tg_alt (Datan_seq x) (S N))
exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) (S N) - sum_f_R0 (tg_alt (Ratan_seq x)) (S N)) / h - sum_f_R0 (tg_alt (Datan_seq x)) (S N)) < eps
N:nat
x:R

x ^ (2 * S N) = / INR (2 * S N + 1) * (INR (2 * S N + 1) * x ^ (2 * S N))
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta
(-1) ^ S N * (((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1) - x ^ (2 * S N + 1) / INR (2 * S N + 1)) / h - x ^ (2 * S N)) = ((-1) ^ S N * ((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1)) - (-1) ^ S N * (x ^ (2 * S N + 1) / INR (2 * S N + 1))) / h - (-1) ^ S N * x ^ (2 * S N)
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
IHN:derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:posreal
Hdelta1:forall h : R, h <> 0 -> Rabs h < delta1 -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
Main:derivable_pt_lim (fun x0 : R => tg_alt (Ratan_seq x0) (S N)) x (tg_alt (Datan_seq x) (S N))
exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) (S N) - sum_f_R0 (tg_alt (Ratan_seq x)) (S N)) / h - sum_f_R0 (tg_alt (Datan_seq x)) (S N)) < eps
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x, eps':R
eps'_pos:0 < eps'
delta:posreal
Hdelta:forall h0 : R, h0 <> 0 -> Rabs h0 < delta -> Rabs (((x + h0) ^ (2 * S N + 1) - x ^ (2 * S N + 1)) / h0 - INR (2 * S N + 1) * x ^ Init.Nat.pred (2 * S N + 1)) < eps'
h:R
h_neq:h <> 0
h_b:Rabs h < delta

(-1) ^ S N * (((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1) - x ^ (2 * S N + 1) / INR (2 * S N + 1)) / h - x ^ (2 * S N)) = ((-1) ^ S N * ((x + h) ^ (2 * S N + 1) / INR (2 * S N + 1)) - (-1) ^ S N * (x ^ (2 * S N + 1) / INR (2 * S N + 1))) / h - (-1) ^ S N * x ^ (2 * S N)
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
IHN:derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:posreal
Hdelta1:forall h : R, h <> 0 -> Rabs h < delta1 -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
Main:derivable_pt_lim (fun x0 : R => tg_alt (Ratan_seq x0) (S N)) x (tg_alt (Datan_seq x) (S N))
exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) (S N) - sum_f_R0 (tg_alt (Ratan_seq x)) (S N)) / h - sum_f_R0 (tg_alt (Datan_seq x)) (S N)) < eps
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
IHN:derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:posreal
Hdelta1:forall h : R, h <> 0 -> Rabs h < delta1 -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
Main:derivable_pt_lim (fun x0 : R => tg_alt (Ratan_seq x0) (S N)) x (tg_alt (Datan_seq x) (S N))

exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) (S N) - sum_f_R0 (tg_alt (Ratan_seq x)) (S N)) / h - sum_f_R0 (tg_alt (Datan_seq x)) (S N)) < eps
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
IHN:derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:posreal
Hdelta1:forall h : R, h <> 0 -> Rabs h < delta1 -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
Main:derivable_pt_lim (fun x0 : R => tg_alt (Ratan_seq x0) (S N)) x (tg_alt (Datan_seq x) (S N))
delta2:posreal
Hdelta2:forall h : R, h <> 0 -> Rabs h < delta2 -> Rabs ((tg_alt (Ratan_seq (x + h)) (S N) - tg_alt (Ratan_seq x) (S N)) / h - tg_alt (Datan_seq x) (S N)) < eps / 3

exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) (S N) - sum_f_R0 (tg_alt (Ratan_seq x)) (S N)) / h - sum_f_R0 (tg_alt (Datan_seq x)) (S N)) < eps
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
IHN:derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:R
delta1_pos:0 < delta1
Hdelta1:forall h : R, h <> 0 -> Rabs h < {| pos := delta1; cond_pos := delta1_pos |} -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
Main:derivable_pt_lim (fun x0 : R => tg_alt (Ratan_seq x0) (S N)) x (tg_alt (Datan_seq x) (S N))
delta2:R
delta2_pos:0 < delta2
Hdelta2:forall h : R, h <> 0 -> Rabs h < {| pos := delta2; cond_pos := delta2_pos |} -> Rabs ((tg_alt (Ratan_seq (x + h)) (S N) - tg_alt (Ratan_seq x) (S N)) / h - tg_alt (Datan_seq x) (S N)) < eps / 3

exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) (S N) - sum_f_R0 (tg_alt (Ratan_seq x)) (S N)) / h - sum_f_R0 (tg_alt (Datan_seq x)) (S N)) < eps
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
IHN:derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:R
delta1_pos:0 < delta1
Hdelta1:forall h : R, h <> 0 -> Rabs h < {| pos := delta1; cond_pos := delta1_pos |} -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
Main:derivable_pt_lim (fun x0 : R => tg_alt (Ratan_seq x0) (S N)) x (tg_alt (Datan_seq x) (S N))
delta2:R
delta2_pos:0 < delta2
Hdelta2:forall h : R, h <> 0 -> Rabs h < {| pos := delta2; cond_pos := delta2_pos |} -> Rabs ((tg_alt (Ratan_seq (x + h)) (S N) - tg_alt (Ratan_seq x) (S N)) / h - tg_alt (Datan_seq x) (S N)) < eps / 3
mydelta:=Rmin delta1 delta2:R

exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) (S N) - sum_f_R0 (tg_alt (Ratan_seq x)) (S N)) / h - sum_f_R0 (tg_alt (Datan_seq x)) (S N)) < eps
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
IHN:derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:R
delta1_pos:0 < delta1
Hdelta1:forall h : R, h <> 0 -> Rabs h < {| pos := delta1; cond_pos := delta1_pos |} -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
Main:derivable_pt_lim (fun x0 : R => tg_alt (Ratan_seq x0) (S N)) x (tg_alt (Datan_seq x) (S N))
delta2:R
delta2_pos:0 < delta2
Hdelta2:forall h : R, h <> 0 -> Rabs h < {| pos := delta2; cond_pos := delta2_pos |} -> Rabs ((tg_alt (Ratan_seq (x + h)) (S N) - tg_alt (Ratan_seq x) (S N)) / h - tg_alt (Datan_seq x) (S N)) < eps / 3
mydelta:=Rmin delta1 delta2:R

mydelta > 0
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
IHN:derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:R
delta1_pos:0 < delta1
Hdelta1:forall h : R, h <> 0 -> Rabs h < {| pos := delta1; cond_pos := delta1_pos |} -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
Main:derivable_pt_lim (fun x0 : R => tg_alt (Ratan_seq x0) (S N)) x (tg_alt (Datan_seq x) (S N))
delta2:R
delta2_pos:0 < delta2
Hdelta2:forall h : R, h <> 0 -> Rabs h < {| pos := delta2; cond_pos := delta2_pos |} -> Rabs ((tg_alt (Ratan_seq (x + h)) (S N) - tg_alt (Ratan_seq x) (S N)) / h - tg_alt (Datan_seq x) (S N)) < eps / 3
mydelta:=Rmin delta1 delta2:R
mydelta_pos:mydelta > 0
exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) (S N) - sum_f_R0 (tg_alt (Ratan_seq x)) (S N)) / h - sum_f_R0 (tg_alt (Datan_seq x)) (S N)) < eps
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
IHN:derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:R
delta1_pos:0 < delta1
Hdelta1:forall h : R, h <> 0 -> Rabs h < {| pos := delta1; cond_pos := delta1_pos |} -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
Main:derivable_pt_lim (fun x0 : R => tg_alt (Ratan_seq x0) (S N)) x (tg_alt (Datan_seq x) (S N))
delta2:R
delta2_pos:0 < delta2
Hdelta2:forall h : R, h <> 0 -> Rabs h < {| pos := delta2; cond_pos := delta2_pos |} -> Rabs ((tg_alt (Ratan_seq (x + h)) (S N) - tg_alt (Ratan_seq x) (S N)) / h - tg_alt (Datan_seq x) (S N)) < eps / 3
mydelta:=Rmin delta1 delta2:R
mydelta_pos:mydelta > 0

exists delta : posreal, forall h : R, h <> 0 -> Rabs h < delta -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) (S N) - sum_f_R0 (tg_alt (Ratan_seq x)) (S N)) / h - sum_f_R0 (tg_alt (Datan_seq x)) (S N)) < eps
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
IHN:derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) x (sum_f_R0 (tg_alt (Datan_seq x)) N)
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:R
delta1_pos:0 < delta1
Hdelta1:forall h0 : R, h0 <> 0 -> Rabs h0 < {| pos := delta1; cond_pos := delta1_pos |} -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h0))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h0 - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
Main:derivable_pt_lim (fun x0 : R => tg_alt (Ratan_seq x0) (S N)) x (tg_alt (Datan_seq x) (S N))
delta2:R
delta2_pos:0 < delta2
Hdelta2:forall h0 : R, h0 <> 0 -> Rabs h0 < {| pos := delta2; cond_pos := delta2_pos |} -> Rabs ((tg_alt (Ratan_seq (x + h0)) (S N) - tg_alt (Ratan_seq x) (S N)) / h0 - tg_alt (Datan_seq x) (S N)) < eps / 3
mydelta:=Rmin delta1 delta2:R
mydelta_pos:mydelta > 0
delta:={| pos := mydelta; cond_pos := mydelta_pos |}:posreal
h:R
h_neq:h <> 0
h_b:Rabs h < delta

Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) (S N) - sum_f_R0 (tg_alt (Ratan_seq x)) (S N)) / h - sum_f_R0 (tg_alt (Datan_seq x)) (S N)) < eps
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:R
delta1_pos:0 < delta1
Hdelta1:forall h0 : R, h0 <> 0 -> Rabs h0 < {| pos := delta1; cond_pos := delta1_pos |} -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h0))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h0 - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
delta2:R
delta2_pos:0 < delta2
Hdelta2:forall h0 : R, h0 <> 0 -> Rabs h0 < {| pos := delta2; cond_pos := delta2_pos |} -> Rabs ((tg_alt (Ratan_seq (x + h0)) (S N) - tg_alt (Ratan_seq x) (S N)) / h0 - tg_alt (Datan_seq x) (S N)) < eps / 3
mydelta:=Rmin delta1 delta2:R
mydelta_pos:mydelta > 0
delta:={| pos := mydelta; cond_pos := mydelta_pos |}:posreal
h:R
h_neq:h <> 0
h_b:Rabs h < delta

Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) (S N) - sum_f_R0 (tg_alt (Ratan_seq x)) (S N)) / h - sum_f_R0 (tg_alt (Datan_seq x)) (S N)) < eps
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:R
delta1_pos:0 < delta1
Hdelta1:forall h0 : R, h0 <> 0 -> Rabs h0 < {| pos := delta1; cond_pos := delta1_pos |} -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h0))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h0 - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
delta2:R
delta2_pos:0 < delta2
Hdelta2:forall h0 : R, h0 <> 0 -> Rabs h0 < {| pos := delta2; cond_pos := delta2_pos |} -> Rabs ((tg_alt (Ratan_seq (x + h0)) (S N) - tg_alt (Ratan_seq x) (S N)) / h0 - tg_alt (Datan_seq x) (S N)) < eps / 3
mydelta:=Rmin delta1 delta2:R
mydelta_pos:mydelta > 0
delta:={| pos := mydelta; cond_pos := mydelta_pos |}:posreal
h:R
h_neq:h <> 0
h_b:Rabs h < delta

Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) (S N) - sum_f_R0 (tg_alt (Ratan_seq x)) (S N)) / h + - sum_f_R0 (tg_alt (Datan_seq x)) (S N)) < eps
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:R
delta1_pos:0 < delta1
Hdelta1:forall h0 : R, h0 <> 0 -> Rabs h0 < {| pos := delta1; cond_pos := delta1_pos |} -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h0))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h0 - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
delta2:R
delta2_pos:0 < delta2
Hdelta2:forall h0 : R, h0 <> 0 -> Rabs h0 < {| pos := delta2; cond_pos := delta2_pos |} -> Rabs ((tg_alt (Ratan_seq (x + h0)) (S N) - tg_alt (Ratan_seq x) (S N)) / h0 - tg_alt (Datan_seq x) (S N)) < eps / 3
mydelta:=Rmin delta1 delta2:R
mydelta_pos:mydelta > 0
delta:={| pos := mydelta; cond_pos := mydelta_pos |}:posreal
h:R
h_neq:h <> 0
h_b:Rabs h < delta

Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) (S N) - sum_f_R0 (tg_alt (Ratan_seq x)) (S N)) / h + - sum_f_R0 (tg_alt (Datan_seq x)) (S N)) <= eps / 3 + eps / 3
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:R
delta1_pos:0 < delta1
Hdelta1:forall h0 : R, h0 <> 0 -> Rabs h0 < {| pos := delta1; cond_pos := delta1_pos |} -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h0))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h0 - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
delta2:R
delta2_pos:0 < delta2
Hdelta2:forall h0 : R, h0 <> 0 -> Rabs h0 < {| pos := delta2; cond_pos := delta2_pos |} -> Rabs ((tg_alt (Ratan_seq (x + h0)) (S N) - tg_alt (Ratan_seq x) (S N)) / h0 - tg_alt (Datan_seq x) (S N)) < eps / 3
mydelta:=Rmin delta1 delta2:R
mydelta_pos:mydelta > 0
delta:={| pos := mydelta; cond_pos := mydelta_pos |}:posreal
h:R
h_neq:h <> 0
h_b:Rabs h < delta
eps / 3 + eps / 3 < eps
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:R
delta1_pos:0 < delta1
Hdelta1:forall h0 : R, h0 <> 0 -> Rabs h0 < {| pos := delta1; cond_pos := delta1_pos |} -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h0))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h0 - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
delta2:R
delta2_pos:0 < delta2
Hdelta2:forall h0 : R, h0 <> 0 -> Rabs h0 < {| pos := delta2; cond_pos := delta2_pos |} -> Rabs ((tg_alt (Ratan_seq (x + h0)) (S N) - tg_alt (Ratan_seq x) (S N)) / h0 - tg_alt (Datan_seq x) (S N)) < eps / 3
mydelta:=Rmin delta1 delta2:R
mydelta_pos:mydelta > 0
delta:={| pos := mydelta; cond_pos := mydelta_pos |}:posreal
h:R
h_neq:h <> 0
h_b:Rabs h < delta

(sum_f_R0 (tg_alt (Ratan_seq (x + h))) (S N) - sum_f_R0 (tg_alt (Ratan_seq x)) (S N)) / h + - sum_f_R0 (tg_alt (Datan_seq x)) (S N) = (sum_f_R0 (tg_alt (Ratan_seq (x + h))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h + - sum_f_R0 (tg_alt (Datan_seq x)) N + ((tg_alt (Ratan_seq (x + h)) (S N) - tg_alt (Ratan_seq x) (S N)) / h - tg_alt (Datan_seq x) (S N))
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:R
delta1_pos:0 < delta1
Hdelta1:forall h0 : R, h0 <> 0 -> Rabs h0 < {| pos := delta1; cond_pos := delta1_pos |} -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h0))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h0 - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
delta2:R
delta2_pos:0 < delta2
Hdelta2:forall h0 : R, h0 <> 0 -> Rabs h0 < {| pos := delta2; cond_pos := delta2_pos |} -> Rabs ((tg_alt (Ratan_seq (x + h0)) (S N) - tg_alt (Ratan_seq x) (S N)) / h0 - tg_alt (Datan_seq x) (S N)) < eps / 3
mydelta:=Rmin delta1 delta2:R
mydelta_pos:mydelta > 0
delta:={| pos := mydelta; cond_pos := mydelta_pos |}:posreal
h:R
h_neq:h <> 0
h_b:Rabs h < delta
Temp:(sum_f_R0 (tg_alt (Ratan_seq (x + h))) (S N) - sum_f_R0 (tg_alt (Ratan_seq x)) (S N)) / h + - sum_f_R0 (tg_alt (Datan_seq x)) (S N) = (sum_f_R0 (tg_alt (Ratan_seq (x + h))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h + - sum_f_R0 (tg_alt (Datan_seq x)) N + ((tg_alt (Ratan_seq (x + h)) (S N) - tg_alt (Ratan_seq x) (S N)) / h - tg_alt (Datan_seq x) (S N))
Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) (S N) - sum_f_R0 (tg_alt (Ratan_seq x)) (S N)) / h + - sum_f_R0 (tg_alt (Datan_seq x)) (S N)) <= eps / 3 + eps / 3
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:R
delta1_pos:0 < delta1
Hdelta1:forall h0 : R, h0 <> 0 -> Rabs h0 < {| pos := delta1; cond_pos := delta1_pos |} -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h0))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h0 - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
delta2:R
delta2_pos:0 < delta2
Hdelta2:forall h0 : R, h0 <> 0 -> Rabs h0 < {| pos := delta2; cond_pos := delta2_pos |} -> Rabs ((tg_alt (Ratan_seq (x + h0)) (S N) - tg_alt (Ratan_seq x) (S N)) / h0 - tg_alt (Datan_seq x) (S N)) < eps / 3
mydelta:=Rmin delta1 delta2:R
mydelta_pos:mydelta > 0
delta:={| pos := mydelta; cond_pos := mydelta_pos |}:posreal
h:R
h_neq:h <> 0
h_b:Rabs h < delta
eps / 3 + eps / 3 < eps
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:R
delta1_pos:0 < delta1
Hdelta1:forall h0 : R, h0 <> 0 -> Rabs h0 < {| pos := delta1; cond_pos := delta1_pos |} -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h0))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h0 - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
delta2:R
delta2_pos:0 < delta2
Hdelta2:forall h0 : R, h0 <> 0 -> Rabs h0 < {| pos := delta2; cond_pos := delta2_pos |} -> Rabs ((tg_alt (Ratan_seq (x + h0)) (S N) - tg_alt (Ratan_seq x) (S N)) / h0 - tg_alt (Datan_seq x) (S N)) < eps / 3
mydelta:=Rmin delta1 delta2:R
mydelta_pos:mydelta > 0
delta:={| pos := mydelta; cond_pos := mydelta_pos |}:posreal
h:R
h_neq:h <> 0
h_b:Rabs h < delta
Temp:(sum_f_R0 (tg_alt (Ratan_seq (x + h))) (S N) - sum_f_R0 (tg_alt (Ratan_seq x)) (S N)) / h + - sum_f_R0 (tg_alt (Datan_seq x)) (S N) = (sum_f_R0 (tg_alt (Ratan_seq (x + h))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h + - sum_f_R0 (tg_alt (Datan_seq x)) N + ((tg_alt (Ratan_seq (x + h)) (S N) - tg_alt (Ratan_seq x) (S N)) / h - tg_alt (Datan_seq x) (S N))

Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) (S N) - sum_f_R0 (tg_alt (Ratan_seq x)) (S N)) / h + - sum_f_R0 (tg_alt (Datan_seq x)) (S N)) <= eps / 3 + eps / 3
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:R
delta1_pos:0 < delta1
Hdelta1:forall h0 : R, h0 <> 0 -> Rabs h0 < {| pos := delta1; cond_pos := delta1_pos |} -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h0))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h0 - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
delta2:R
delta2_pos:0 < delta2
Hdelta2:forall h0 : R, h0 <> 0 -> Rabs h0 < {| pos := delta2; cond_pos := delta2_pos |} -> Rabs ((tg_alt (Ratan_seq (x + h0)) (S N) - tg_alt (Ratan_seq x) (S N)) / h0 - tg_alt (Datan_seq x) (S N)) < eps / 3
mydelta:=Rmin delta1 delta2:R
mydelta_pos:mydelta > 0
delta:={| pos := mydelta; cond_pos := mydelta_pos |}:posreal
h:R
h_neq:h <> 0
h_b:Rabs h < delta
eps / 3 + eps / 3 < eps
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:R
delta1_pos:0 < delta1
Hdelta1:forall h0 : R, h0 <> 0 -> Rabs h0 < {| pos := delta1; cond_pos := delta1_pos |} -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h0))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h0 - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
delta2:R
delta2_pos:0 < delta2
Hdelta2:forall h0 : R, h0 <> 0 -> Rabs h0 < {| pos := delta2; cond_pos := delta2_pos |} -> Rabs ((tg_alt (Ratan_seq (x + h0)) (S N) - tg_alt (Ratan_seq x) (S N)) / h0 - tg_alt (Datan_seq x) (S N)) < eps / 3
mydelta:=Rmin delta1 delta2:R
mydelta_pos:mydelta > 0
delta:={| pos := mydelta; cond_pos := mydelta_pos |}:posreal
h:R
h_neq:h <> 0
h_b:Rabs h < delta
Temp:(sum_f_R0 (tg_alt (Ratan_seq (x + h))) (S N) - sum_f_R0 (tg_alt (Ratan_seq x)) (S N)) / h + - sum_f_R0 (tg_alt (Datan_seq x)) (S N) = (sum_f_R0 (tg_alt (Ratan_seq (x + h))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h + - sum_f_R0 (tg_alt (Datan_seq x)) N + ((tg_alt (Ratan_seq (x + h)) (S N) - tg_alt (Ratan_seq x) (S N)) / h - tg_alt (Datan_seq x) (S N))

Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) (S N) - sum_f_R0 (tg_alt (Ratan_seq x)) (S N)) / h + - sum_f_R0 (tg_alt (Datan_seq x)) (S N)) <= Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h + - sum_f_R0 (tg_alt (Datan_seq x)) N) + Rabs ((tg_alt (Ratan_seq (x + h)) (S N) - tg_alt (Ratan_seq x) (S N)) / h - tg_alt (Datan_seq x) (S N))
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:R
delta1_pos:0 < delta1
Hdelta1:forall h0 : R, h0 <> 0 -> Rabs h0 < {| pos := delta1; cond_pos := delta1_pos |} -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h0))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h0 - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
delta2:R
delta2_pos:0 < delta2
Hdelta2:forall h0 : R, h0 <> 0 -> Rabs h0 < {| pos := delta2; cond_pos := delta2_pos |} -> Rabs ((tg_alt (Ratan_seq (x + h0)) (S N) - tg_alt (Ratan_seq x) (S N)) / h0 - tg_alt (Datan_seq x) (S N)) < eps / 3
mydelta:=Rmin delta1 delta2:R
mydelta_pos:mydelta > 0
delta:={| pos := mydelta; cond_pos := mydelta_pos |}:posreal
h:R
h_neq:h <> 0
h_b:Rabs h < delta
Temp:(sum_f_R0 (tg_alt (Ratan_seq (x + h))) (S N) - sum_f_R0 (tg_alt (Ratan_seq x)) (S N)) / h + - sum_f_R0 (tg_alt (Datan_seq x)) (S N) = (sum_f_R0 (tg_alt (Ratan_seq (x + h))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h + - sum_f_R0 (tg_alt (Datan_seq x)) N + ((tg_alt (Ratan_seq (x + h)) (S N) - tg_alt (Ratan_seq x) (S N)) / h - tg_alt (Datan_seq x) (S N))
Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h + - sum_f_R0 (tg_alt (Datan_seq x)) N) + Rabs ((tg_alt (Ratan_seq (x + h)) (S N) - tg_alt (Ratan_seq x) (S N)) / h - tg_alt (Datan_seq x) (S N)) <= eps / 3 + eps / 3
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:R
delta1_pos:0 < delta1
Hdelta1:forall h0 : R, h0 <> 0 -> Rabs h0 < {| pos := delta1; cond_pos := delta1_pos |} -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h0))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h0 - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
delta2:R
delta2_pos:0 < delta2
Hdelta2:forall h0 : R, h0 <> 0 -> Rabs h0 < {| pos := delta2; cond_pos := delta2_pos |} -> Rabs ((tg_alt (Ratan_seq (x + h0)) (S N) - tg_alt (Ratan_seq x) (S N)) / h0 - tg_alt (Datan_seq x) (S N)) < eps / 3
mydelta:=Rmin delta1 delta2:R
mydelta_pos:mydelta > 0
delta:={| pos := mydelta; cond_pos := mydelta_pos |}:posreal
h:R
h_neq:h <> 0
h_b:Rabs h < delta
eps / 3 + eps / 3 < eps
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:R
delta1_pos:0 < delta1
Hdelta1:forall h0 : R, h0 <> 0 -> Rabs h0 < {| pos := delta1; cond_pos := delta1_pos |} -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h0))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h0 - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
delta2:R
delta2_pos:0 < delta2
Hdelta2:forall h0 : R, h0 <> 0 -> Rabs h0 < {| pos := delta2; cond_pos := delta2_pos |} -> Rabs ((tg_alt (Ratan_seq (x + h0)) (S N) - tg_alt (Ratan_seq x) (S N)) / h0 - tg_alt (Datan_seq x) (S N)) < eps / 3
mydelta:=Rmin delta1 delta2:R
mydelta_pos:mydelta > 0
delta:={| pos := mydelta; cond_pos := mydelta_pos |}:posreal
h:R
h_neq:h <> 0
h_b:Rabs h < delta
Temp:(sum_f_R0 (tg_alt (Ratan_seq (x + h))) (S N) - sum_f_R0 (tg_alt (Ratan_seq x)) (S N)) / h + - sum_f_R0 (tg_alt (Datan_seq x)) (S N) = (sum_f_R0 (tg_alt (Ratan_seq (x + h))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h + - sum_f_R0 (tg_alt (Datan_seq x)) N + ((tg_alt (Ratan_seq (x + h)) (S N) - tg_alt (Ratan_seq x) (S N)) / h - tg_alt (Datan_seq x) (S N))

Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h + - sum_f_R0 (tg_alt (Datan_seq x)) N) + Rabs ((tg_alt (Ratan_seq (x + h)) (S N) - tg_alt (Ratan_seq x) (S N)) / h - tg_alt (Datan_seq x) (S N)) <= eps / 3 + eps / 3
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:R
delta1_pos:0 < delta1
Hdelta1:forall h0 : R, h0 <> 0 -> Rabs h0 < {| pos := delta1; cond_pos := delta1_pos |} -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h0))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h0 - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
delta2:R
delta2_pos:0 < delta2
Hdelta2:forall h0 : R, h0 <> 0 -> Rabs h0 < {| pos := delta2; cond_pos := delta2_pos |} -> Rabs ((tg_alt (Ratan_seq (x + h0)) (S N) - tg_alt (Ratan_seq x) (S N)) / h0 - tg_alt (Datan_seq x) (S N)) < eps / 3
mydelta:=Rmin delta1 delta2:R
mydelta_pos:mydelta > 0
delta:={| pos := mydelta; cond_pos := mydelta_pos |}:posreal
h:R
h_neq:h <> 0
h_b:Rabs h < delta
eps / 3 + eps / 3 < eps
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:R
delta1_pos:0 < delta1
Hdelta1:forall h0 : R, (h0 = 0 -> False) -> Rabs h0 < {| pos := delta1; cond_pos := delta1_pos |} -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h0))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h0 - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
delta2:R
delta2_pos:0 < delta2
Hdelta2:forall h0 : R, (h0 = 0 -> False) -> Rabs h0 < {| pos := delta2; cond_pos := delta2_pos |} -> Rabs ((tg_alt (Ratan_seq (x + h0)) (S N) - tg_alt (Ratan_seq x) (S N)) / h0 - tg_alt (Datan_seq x) (S N)) < eps / 3
mydelta:=Rmin delta1 delta2:R
mydelta_pos:mydelta > 0
delta:={| pos := mydelta; cond_pos := mydelta_pos |}:posreal
h:R
h_neq:h = 0 -> False
h_b:Rabs h < delta
Temp:(sum_f_R0 (tg_alt (Ratan_seq (x + h))) (S N) - sum_f_R0 (tg_alt (Ratan_seq x)) (S N)) / h + - sum_f_R0 (tg_alt (Datan_seq x)) (S N) = (sum_f_R0 (tg_alt (Ratan_seq (x + h))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h + - sum_f_R0 (tg_alt (Datan_seq x)) N + ((tg_alt (Ratan_seq (x + h)) (S N) - tg_alt (Ratan_seq x) (S N)) / h - tg_alt (Datan_seq x) (S N))

{| pos := Rmin delta1 delta2; cond_pos := mydelta_pos |} <= {| pos := delta1; cond_pos := delta1_pos |}
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:R
delta1_pos:0 < delta1
Hdelta1:forall h0 : R, (h0 = 0 -> False) -> Rabs h0 < {| pos := delta1; cond_pos := delta1_pos |} -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h0))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h0 - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
delta2:R
delta2_pos:0 < delta2
Hdelta2:forall h0 : R, (h0 = 0 -> False) -> Rabs h0 < {| pos := delta2; cond_pos := delta2_pos |} -> Rabs ((tg_alt (Ratan_seq (x + h0)) (S N) - tg_alt (Ratan_seq x) (S N)) / h0 - tg_alt (Datan_seq x) (S N)) < eps / 3
mydelta:=Rmin delta1 delta2:R
mydelta_pos:mydelta > 0
delta:={| pos := mydelta; cond_pos := mydelta_pos |}:posreal
h:R
h_neq:h = 0 -> False
h_b:Rabs h < delta
Temp:(sum_f_R0 (tg_alt (Ratan_seq (x + h))) (S N) - sum_f_R0 (tg_alt (Ratan_seq x)) (S N)) / h + - sum_f_R0 (tg_alt (Datan_seq x)) (S N) = (sum_f_R0 (tg_alt (Ratan_seq (x + h))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h + - sum_f_R0 (tg_alt (Datan_seq x)) N + ((tg_alt (Ratan_seq (x + h)) (S N) - tg_alt (Ratan_seq x) (S N)) / h - tg_alt (Datan_seq x) (S N))
{| pos := Rmin delta1 delta2; cond_pos := mydelta_pos |} <= {| pos := delta2; cond_pos := delta2_pos |}
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:R
delta1_pos:0 < delta1
Hdelta1:forall h0 : R, h0 <> 0 -> Rabs h0 < {| pos := delta1; cond_pos := delta1_pos |} -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h0))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h0 - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
delta2:R
delta2_pos:0 < delta2
Hdelta2:forall h0 : R, h0 <> 0 -> Rabs h0 < {| pos := delta2; cond_pos := delta2_pos |} -> Rabs ((tg_alt (Ratan_seq (x + h0)) (S N) - tg_alt (Ratan_seq x) (S N)) / h0 - tg_alt (Datan_seq x) (S N)) < eps / 3
mydelta:=Rmin delta1 delta2:R
mydelta_pos:mydelta > 0
delta:={| pos := mydelta; cond_pos := mydelta_pos |}:posreal
h:R
h_neq:h <> 0
h_b:Rabs h < delta
eps / 3 + eps / 3 < eps
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:R
delta1_pos:0 < delta1
Hdelta1:forall h0 : R, (h0 = 0 -> False) -> Rabs h0 < {| pos := delta1; cond_pos := delta1_pos |} -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h0))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h0 - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
delta2:R
delta2_pos:0 < delta2
Hdelta2:forall h0 : R, (h0 = 0 -> False) -> Rabs h0 < {| pos := delta2; cond_pos := delta2_pos |} -> Rabs ((tg_alt (Ratan_seq (x + h0)) (S N) - tg_alt (Ratan_seq x) (S N)) / h0 - tg_alt (Datan_seq x) (S N)) < eps / 3
mydelta:=Rmin delta1 delta2:R
mydelta_pos:mydelta > 0
delta:={| pos := mydelta; cond_pos := mydelta_pos |}:posreal
h:R
h_neq:h = 0 -> False
h_b:Rabs h < delta
Temp:(sum_f_R0 (tg_alt (Ratan_seq (x + h))) (S N) - sum_f_R0 (tg_alt (Ratan_seq x)) (S N)) / h + - sum_f_R0 (tg_alt (Datan_seq x)) (S N) = (sum_f_R0 (tg_alt (Ratan_seq (x + h))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h + - sum_f_R0 (tg_alt (Datan_seq x)) N + ((tg_alt (Ratan_seq (x + h)) (S N) - tg_alt (Ratan_seq x) (S N)) / h - tg_alt (Datan_seq x) (S N))

{| pos := Rmin delta1 delta2; cond_pos := mydelta_pos |} <= {| pos := delta2; cond_pos := delta2_pos |}
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:R
delta1_pos:0 < delta1
Hdelta1:forall h0 : R, h0 <> 0 -> Rabs h0 < {| pos := delta1; cond_pos := delta1_pos |} -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h0))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h0 - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
delta2:R
delta2_pos:0 < delta2
Hdelta2:forall h0 : R, h0 <> 0 -> Rabs h0 < {| pos := delta2; cond_pos := delta2_pos |} -> Rabs ((tg_alt (Ratan_seq (x + h0)) (S N) - tg_alt (Ratan_seq x) (S N)) / h0 - tg_alt (Datan_seq x) (S N)) < eps / 3
mydelta:=Rmin delta1 delta2:R
mydelta_pos:mydelta > 0
delta:={| pos := mydelta; cond_pos := mydelta_pos |}:posreal
h:R
h_neq:h <> 0
h_b:Rabs h < delta
eps / 3 + eps / 3 < eps
Tool:forall N0 : nat, (-1) ^ S (2 * N0) = -1
N:nat
x:R
x_lb:-1 <= x
x_ub:x < 1
eps:R
eps_pos:0 < eps
eps_3_pos:eps / 3 > 0
delta1:R
delta1_pos:0 < delta1
Hdelta1:forall h0 : R, h0 <> 0 -> Rabs h0 < {| pos := delta1; cond_pos := delta1_pos |} -> Rabs ((sum_f_R0 (tg_alt (Ratan_seq (x + h0))) N - sum_f_R0 (tg_alt (Ratan_seq x)) N) / h0 - sum_f_R0 (tg_alt (Datan_seq x)) N) < eps / 3
delta2:R
delta2_pos:0 < delta2
Hdelta2:forall h0 : R, h0 <> 0 -> Rabs h0 < {| pos := delta2; cond_pos := delta2_pos |} -> Rabs ((tg_alt (Ratan_seq (x + h0)) (S N) - tg_alt (Ratan_seq x) (S N)) / h0 - tg_alt (Datan_seq x) (S N)) < eps / 3
mydelta:=Rmin delta1 delta2:R
mydelta_pos:mydelta > 0
delta:={| pos := mydelta; cond_pos := mydelta_pos |}:posreal
h:R
h_neq:h <> 0
h_b:Rabs h < delta

eps / 3 + eps / 3 < eps
lra. Qed.

CVU (fun (N : nat) (x : R) => sum_f_R0 (tg_alt (Ratan_seq x)) N) ps_atan (/ 2) {| pos := / 2; cond_pos := pos_half_prf |}

CVU (fun (N : nat) (x : R) => sum_f_R0 (tg_alt (Ratan_seq x)) N) ps_atan (/ 2) {| pos := / 2; cond_pos := pos_half_prf |}

forall x : R, Boule (/ 2) pos_half x -> Un_decreasing (fun n : nat => Ratan_seq x n)

forall x : R, Boule (/ 2) pos_half x -> Un_cv (fun n : nat => Ratan_seq x n) 0

forall x : R, Boule (/ 2) pos_half x -> Un_cv (sum_f_R0 (tg_alt (fun i : nat => Ratan_seq x i))) (ps_atan x)

forall (x : R) (n : nat), Boule (/ 2) pos_half x -> Ratan_seq x n <= PI_tg n

Un_cv PI_tg 0

forall x : R, Boule (/ 2) pos_half x -> Un_cv (fun n : nat => Ratan_seq x n) 0

forall x : R, Boule (/ 2) pos_half x -> Un_cv (sum_f_R0 (tg_alt (fun i : nat => Ratan_seq x i))) (ps_atan x)

forall (x : R) (n : nat), Boule (/ 2) pos_half x -> Ratan_seq x n <= PI_tg n

Un_cv PI_tg 0

forall x : R, Boule (/ 2) pos_half x -> Un_cv (sum_f_R0 (tg_alt (fun i : nat => Ratan_seq x i))) (ps_atan x)

forall (x : R) (n : nat), Boule (/ 2) pos_half x -> Ratan_seq x n <= PI_tg n

Un_cv PI_tg 0
x:R
b:0 <= x <= 1

Un_cv (sum_f_R0 (tg_alt (fun i : nat => Ratan_seq x i))) (ps_atan x)

forall (x : R) (n : nat), Boule (/ 2) pos_half x -> Ratan_seq x n <= PI_tg n

Un_cv PI_tg 0
x:R
b:0 <= x <= 1
inside:-1 <= x <= 1

Un_cv (sum_f_R0 (tg_alt (fun i : nat => Ratan_seq x i))) (let (v, _) := ps_atan_exists_1 x inside in v)

forall (x : R) (n : nat), Boule (/ 2) pos_half x -> Ratan_seq x n <= PI_tg n

Un_cv PI_tg 0
x:R
b:0 <= x <= 1
inside:-1 <= x <= 1
v:R
Pv:Un_cv (fun N : nat => sum_f_R0 (tg_alt (Ratan_seq x)) N) v

Un_cv (sum_f_R0 (tg_alt (fun i : nat => Ratan_seq x i))) v

forall (x : R) (n : nat), Boule (/ 2) pos_half x -> Ratan_seq x n <= PI_tg n

Un_cv PI_tg 0

forall (x : R) (n : nat), Boule (/ 2) pos_half x -> Ratan_seq x n <= PI_tg n

Un_cv PI_tg 0
x:R
n:nat
b:0 <= x <= 1

Ratan_seq x n <= PI_tg n

Un_cv PI_tg 0
x:R
n:nat
b:0 <= x <= 1

x ^ (2 * n + 1) / INR (2 * n + 1) <= 1 * / INR (2 * n + 1)

Un_cv PI_tg 0
x:R
n:nat
b:0 <= x <= 1

0 <= / INR (2 * n + 1)
x:R
n:nat
b:0 <= x <= 1
x ^ (2 * n + 1) <= 1

Un_cv PI_tg 0
x:R
n:nat
b:0 <= x <= 1

x ^ (2 * n + 1) <= 1

Un_cv PI_tg 0

Un_cv PI_tg 0
exact PI_tg_cv. Qed.

CVU (fun (N : nat) (x : R) => sum_f_R0 (tg_alt (Ratan_seq x)) N) ps_atan 0 {| pos := 1; cond_pos := Rlt_0_1 |}

CVU (fun (N : nat) (x : R) => sum_f_R0 (tg_alt (Ratan_seq x)) N) ps_atan 0 {| pos := 1; cond_pos := Rlt_0_1 |}
eps:R
ep:0 < eps
N:nat
Pn:forall (n : nat) (y : R), (N <= n)%nat -> Boule (/ 2) {| pos := / 2; cond_pos := pos_half_prf |} y -> Rabs (ps_atan y - sum_f_R0 (tg_alt (Ratan_seq y)) n) < eps

exists N0 : nat, forall (n : nat) (y : R), (N0 <= n)%nat -> Boule 0 {| pos := 1; cond_pos := Rlt_0_1 |} y -> Rabs (ps_atan y - sum_f_R0 (tg_alt (Ratan_seq y)) n) < eps
eps:R
ep:0 < eps
N:nat
Pn:forall (n0 : nat) (y : R), (N <= n0)%nat -> Boule (/ 2) {| pos := / 2; cond_pos := pos_half_prf |} y -> Rabs (ps_atan y - sum_f_R0 (tg_alt (Ratan_seq y)) n0) < eps
n:nat
x:R
nN:(N <= n)%nat
b_y:Boule 0 {| pos := 1; cond_pos := Rlt_0_1 |} x

Rabs (ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) n) < eps
eps:R
ep:0 < eps
N:nat
Pn:forall (n0 : nat) (y : R), (N <= n0)%nat -> Boule (/ 2) {| pos := / 2; cond_pos := pos_half_prf |} y -> Rabs (ps_atan y - sum_f_R0 (tg_alt (Ratan_seq y)) n0) < eps
n:nat
x:R
nN:(N <= n)%nat
b_y:Boule 0 {| pos := 1; cond_pos := Rlt_0_1 |} x
xgt0:0 < x

Rabs (ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) n) < eps
eps:R
ep:0 < eps
N:nat
Pn:forall (n0 : nat) (y : R), (N <= n0)%nat -> Boule (/ 2) {| pos := / 2; cond_pos := pos_half_prf |} y -> Rabs (ps_atan y - sum_f_R0 (tg_alt (Ratan_seq y)) n0) < eps
n:nat
x:R
nN:(N <= n)%nat
b_y:Boule 0 {| pos := 1; cond_pos := Rlt_0_1 |} x
x0:0 = x
Rabs (ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) n) < eps
eps:R
ep:0 < eps
N:nat
Pn:forall (n0 : nat) (y : R), (N <= n0)%nat -> Boule (/ 2) {| pos := / 2; cond_pos := pos_half_prf |} y -> Rabs (ps_atan y - sum_f_R0 (tg_alt (Ratan_seq y)) n0) < eps
n:nat
x:R
nN:(N <= n)%nat
b_y:Boule 0 {| pos := 1; cond_pos := Rlt_0_1 |} x
x0:0 > x
Rabs (ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) n) < eps
eps:R
ep:0 < eps
N:nat
Pn:forall (n0 : nat) (y : R), (N <= n0)%nat -> Boule (/ 2) {| pos := / 2; cond_pos := pos_half_prf |} y -> Rabs (ps_atan y - sum_f_R0 (tg_alt (Ratan_seq y)) n0) < eps
n:nat
x:R
nN:(N <= n)%nat
b_y:Boule 0 {| pos := 1; cond_pos := Rlt_0_1 |} x
xgt0:0 < x

Boule (/ 2) {| pos := / 2; cond_pos := pos_half_prf |} x
eps:R
ep:0 < eps
N:nat
Pn:forall (n0 : nat) (y : R), (N <= n0)%nat -> Boule (/ 2) {| pos := / 2; cond_pos := pos_half_prf |} y -> Rabs (ps_atan y - sum_f_R0 (tg_alt (Ratan_seq y)) n0) < eps
n:nat
x:R
nN:(N <= n)%nat
b_y:Boule 0 {| pos := 1; cond_pos := Rlt_0_1 |} x
xgt0:0 < x
H:Boule (/ 2) {| pos := / 2; cond_pos := pos_half_prf |} x
Rabs (ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) n) < eps
eps:R
ep:0 < eps
N:nat
Pn:forall (n0 : nat) (y : R), (N <= n0)%nat -> Boule (/ 2) {| pos := / 2; cond_pos := pos_half_prf |} y -> Rabs (ps_atan y - sum_f_R0 (tg_alt (Ratan_seq y)) n0) < eps
n:nat
x:R
nN:(N <= n)%nat
b_y:Boule 0 {| pos := 1; cond_pos := Rlt_0_1 |} x
x0:0 = x
Rabs (ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) n) < eps
eps:R
ep:0 < eps
N:nat
Pn:forall (n0 : nat) (y : R), (N <= n0)%nat -> Boule (/ 2) {| pos := / 2; cond_pos := pos_half_prf |} y -> Rabs (ps_atan y - sum_f_R0 (tg_alt (Ratan_seq y)) n0) < eps
n:nat
x:R
nN:(N <= n)%nat
b_y:Boule 0 {| pos := 1; cond_pos := Rlt_0_1 |} x
x0:0 > x
Rabs (ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) n) < eps
eps:R
ep:0 < eps
N:nat
Pn:forall (n0 : nat) (y : R), (N <= n0)%nat -> Boule (/ 2) {| pos := / 2; cond_pos := pos_half_prf |} y -> Rabs (ps_atan y - sum_f_R0 (tg_alt (Ratan_seq y)) n0) < eps
n:nat
x:R
nN:(N <= n)%nat
xgt0:0 < x
b_y:x - 0 < 1 /\ - (1) < x - 0

Rabs (x - / 2) < / 2
eps:R
ep:0 < eps
N:nat
Pn:forall (n0 : nat) (y : R), (N <= n0)%nat -> Boule (/ 2) {| pos := / 2; cond_pos := pos_half_prf |} y -> Rabs (ps_atan y - sum_f_R0 (tg_alt (Ratan_seq y)) n0) < eps
n:nat
x:R
nN:(N <= n)%nat
b_y:Boule 0 {| pos := 1; cond_pos := Rlt_0_1 |} x
xgt0:0 < x
H:Boule (/ 2) {| pos := / 2; cond_pos := pos_half_prf |} x
Rabs (ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) n) < eps
eps:R
ep:0 < eps
N:nat
Pn:forall (n0 : nat) (y : R), (N <= n0)%nat -> Boule (/ 2) {| pos := / 2; cond_pos := pos_half_prf |} y -> Rabs (ps_atan y - sum_f_R0 (tg_alt (Ratan_seq y)) n0) < eps
n:nat
x:R
nN:(N <= n)%nat
b_y:Boule 0 {| pos := 1; cond_pos := Rlt_0_1 |} x
x0:0 = x
Rabs (ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) n) < eps
eps:R
ep:0 < eps
N:nat
Pn:forall (n0 : nat) (y : R), (N <= n0)%nat -> Boule (/ 2) {| pos := / 2; cond_pos := pos_half_prf |} y -> Rabs (ps_atan y - sum_f_R0 (tg_alt (Ratan_seq y)) n0) < eps
n:nat
x:R
nN:(N <= n)%nat
b_y:Boule 0 {| pos := 1; cond_pos := Rlt_0_1 |} x
x0:0 > x
Rabs (ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) n) < eps
eps:R
ep:0 < eps
N:nat
Pn:forall (n0 : nat) (y : R), (N <= n0)%nat -> Boule (/ 2) {| pos := / 2; cond_pos := pos_half_prf |} y -> Rabs (ps_atan y - sum_f_R0 (tg_alt (Ratan_seq y)) n0) < eps
n:nat
x:R
nN:(N <= n)%nat
b_y:Boule 0 {| pos := 1; cond_pos := Rlt_0_1 |} x
xgt0:0 < x
H:Boule (/ 2) {| pos := / 2; cond_pos := pos_half_prf |} x

Rabs (ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) n) < eps
eps:R
ep:0 < eps
N:nat
Pn:forall (n0 : nat) (y : R), (N <= n0)%nat -> Boule (/ 2) {| pos := / 2; cond_pos := pos_half_prf |} y -> Rabs (ps_atan y - sum_f_R0 (tg_alt (Ratan_seq y)) n0) < eps
n:nat
x:R
nN:(N <= n)%nat
b_y:Boule 0 {| pos := 1; cond_pos := Rlt_0_1 |} x
x0:0 = x
Rabs (ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) n) < eps
eps:R
ep:0 < eps
N:nat
Pn:forall (n0 : nat) (y : R), (N <= n0)%nat -> Boule (/ 2) {| pos := / 2; cond_pos := pos_half_prf |} y -> Rabs (ps_atan y - sum_f_R0 (tg_alt (Ratan_seq y)) n0) < eps
n:nat
x:R
nN:(N <= n)%nat
b_y:Boule 0 {| pos := 1; cond_pos := Rlt_0_1 |} x
x0:0 > x
Rabs (ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) n) < eps
eps:R
ep:0 < eps
N:nat
Pn:forall (n0 : nat) (y : R), (N <= n0)%nat -> Boule (/ 2) {| pos := / 2; cond_pos := pos_half_prf |} y -> Rabs (ps_atan y - sum_f_R0 (tg_alt (Ratan_seq y)) n0) < eps
n:nat
x:R
nN:(N <= n)%nat
b_y:Boule 0 {| pos := 1; cond_pos := Rlt_0_1 |} x
x0:0 = x

Rabs (ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) n) < eps
eps:R
ep:0 < eps
N:nat
Pn:forall (n0 : nat) (y : R), (N <= n0)%nat -> Boule (/ 2) {| pos := / 2; cond_pos := pos_half_prf |} y -> Rabs (ps_atan y - sum_f_R0 (tg_alt (Ratan_seq y)) n0) < eps
n:nat
x:R
nN:(N <= n)%nat
b_y:Boule 0 {| pos := 1; cond_pos := Rlt_0_1 |} x
x0:0 > x
Rabs (ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) n) < eps
eps:R
ep:0 < eps
N:nat
Pn:forall (n0 : nat) (y : R), (N <= n0)%nat -> Boule (/ 2) {| pos := / 2; cond_pos := pos_half_prf |} y -> Rabs (ps_atan y - sum_f_R0 (tg_alt (Ratan_seq y)) n0) < eps
n:nat
x:R
nN:(N <= n)%nat
b_y:Boule 0 {| pos := 1; cond_pos := Rlt_0_1 |} x
x0:0 = x

Rabs (0 - sum_f_R0 (tg_alt (Ratan_seq 0)) n) < eps
eps:R
ep:0 < eps
N:nat
Pn:forall (n0 : nat) (y : R), (N <= n0)%nat -> Boule (/ 2) {| pos := / 2; cond_pos := pos_half_prf |} y -> Rabs (ps_atan y - sum_f_R0 (tg_alt (Ratan_seq y)) n0) < eps
n:nat
x:R
nN:(N <= n)%nat
b_y:Boule 0 {| pos := 1; cond_pos := Rlt_0_1 |} x
x0:0 > x
Rabs (ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) n) < eps
eps:R
ep:0 < eps
N:nat
Pn:forall (n0 : nat) (y : R), (N <= n0)%nat -> Boule (/ 2) {| pos := / 2; cond_pos := pos_half_prf |} y -> Rabs (ps_atan y - sum_f_R0 (tg_alt (Ratan_seq y)) n0) < eps
n:nat
x:R
nN:(N <= n)%nat
b_y:Boule 0 {| pos := 1; cond_pos := Rlt_0_1 |} x
x0:0 = x

0 < eps
eps:R
ep:0 < eps
N:nat
Pn:forall (n0 : nat) (y : R), (N <= n0)%nat -> Boule (/ 2) {| pos := / 2; cond_pos := pos_half_prf |} y -> Rabs (ps_atan y - sum_f_R0 (tg_alt (Ratan_seq y)) n0) < eps
n:nat
x:R
nN:(N <= n)%nat
b_y:Boule 0 {| pos := 1; cond_pos := Rlt_0_1 |} x
x0:0 = x
0 <= 0
eps:R
ep:0 < eps
N:nat
Pn:forall (n0 : nat) (y : R), (N <= n0)%nat -> Boule (/ 2) {| pos := / 2; cond_pos := pos_half_prf |} y -> Rabs (ps_atan y - sum_f_R0 (tg_alt (Ratan_seq y)) n0) < eps
n:nat
x:R
nN:(N <= n)%nat
b_y:Boule 0 {| pos := 1; cond_pos := Rlt_0_1 |} x
x0:0 = x
forall i : nat, (i <= n)%nat -> 0 = tg_alt (Ratan_seq 0) i
eps:R
ep:0 < eps
N:nat
Pn:forall (n0 : nat) (y : R), (N <= n0)%nat -> Boule (/ 2) {| pos := / 2; cond_pos := pos_half_prf |} y -> Rabs (ps_atan y - sum_f_R0 (tg_alt (Ratan_seq y)) n0) < eps
n:nat
x:R
nN:(N <= n)%nat
b_y:Boule 0 {| pos := 1; cond_pos := Rlt_0_1 |} x
x0:0 > x
Rabs (ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) n) < eps
eps:R
ep:0 < eps
N:nat
Pn:forall (n0 : nat) (y : R), (N <= n0)%nat -> Boule (/ 2) {| pos := / 2; cond_pos := pos_half_prf |} y -> Rabs (ps_atan y - sum_f_R0 (tg_alt (Ratan_seq y)) n0) < eps
n:nat
x:R
nN:(N <= n)%nat
b_y:Boule 0 {| pos := 1; cond_pos := Rlt_0_1 |} x
x0:0 = x

0 <= 0
eps:R
ep:0 < eps
N:nat
Pn:forall (n0 : nat) (y : R), (N <= n0)%nat -> Boule (/ 2) {| pos := / 2; cond_pos := pos_half_prf |} y -> Rabs (ps_atan y - sum_f_R0 (tg_alt (Ratan_seq y)) n0) < eps
n:nat
x:R
nN:(N <= n)%nat
b_y:Boule 0 {| pos := 1; cond_pos := Rlt_0_1 |} x
x0:0 = x
forall i : nat, (i <= n)%nat -> 0 = tg_alt (Ratan_seq 0) i
eps:R
ep:0 < eps
N:nat
Pn:forall (n0 : nat) (y : R), (N <= n0)%nat -> Boule (/ 2) {| pos := / 2; cond_pos := pos_half_prf |} y -> Rabs (ps_atan y - sum_f_R0 (tg_alt (Ratan_seq y)) n0) < eps
n:nat
x:R
nN:(N <= n)%nat
b_y:Boule 0 {| pos := 1; cond_pos := Rlt_0_1 |} x
x0:0 > x
Rabs (ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) n) < eps
eps:R
ep:0 < eps
N:nat
Pn:forall (n0 : nat) (y : R), (N <= n0)%nat -> Boule (/ 2) {| pos := / 2; cond_pos := pos_half_prf |} y -> Rabs (ps_atan y - sum_f_R0 (tg_alt (Ratan_seq y)) n0) < eps
n:nat
x:R
nN:(N <= n)%nat
b_y:Boule 0 {| pos := 1; cond_pos := Rlt_0_1 |} x
x0:0 = x

forall i : nat, (i <= n)%nat -> 0 = tg_alt (Ratan_seq 0) i
eps:R
ep:0 < eps
N:nat
Pn:forall (n0 : nat) (y : R), (N <= n0)%nat -> Boule (/ 2) {| pos := / 2; cond_pos := pos_half_prf |} y -> Rabs (ps_atan y - sum_f_R0 (tg_alt (Ratan_seq y)) n0) < eps
n:nat
x:R
nN:(N <= n)%nat
b_y:Boule 0 {| pos := 1; cond_pos := Rlt_0_1 |} x
x0:0 > x
Rabs (ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) n) < eps
eps:R
ep:0 < eps
N:nat
Pn:forall (n0 : nat) (y : R), (N <= n0)%nat -> Boule (/ 2) {| pos := / 2; cond_pos := pos_half_prf |} y -> Rabs (ps_atan y - sum_f_R0 (tg_alt (Ratan_seq y)) n0) < eps
n:nat
x:R
nN:(N <= n)%nat
b_y:Boule 0 {| pos := 1; cond_pos := Rlt_0_1 |} x
x0:0 = x
i:nat

0 = (-1) ^ i * (0 * 0 ^ (i + (i + 0)) * / match (i + (i + 0))%nat with | 0%nat => 1 | S _ => INR (i + (i + 0)) + 1 end)
eps:R
ep:0 < eps
N:nat
Pn:forall (n0 : nat) (y : R), (N <= n0)%nat -> Boule (/ 2) {| pos := / 2; cond_pos := pos_half_prf |} y -> Rabs (ps_atan y - sum_f_R0 (tg_alt (Ratan_seq y)) n0) < eps
n:nat
x:R
nN:(N <= n)%nat
b_y:Boule 0 {| pos := 1; cond_pos := Rlt_0_1 |} x
x0:0 > x
Rabs (ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) n) < eps
eps:R
ep:0 < eps
N:nat
Pn:forall (n0 : nat) (y : R), (N <= n0)%nat -> Boule (/ 2) {| pos := / 2; cond_pos := pos_half_prf |} y -> Rabs (ps_atan y - sum_f_R0 (tg_alt (Ratan_seq y)) n0) < eps
n:nat
x:R
nN:(N <= n)%nat
b_y:Boule 0 {| pos := 1; cond_pos := Rlt_0_1 |} x
x0:0 > x

Rabs (ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) n) < eps
eps:R
ep:0 < eps
N:nat
Pn:forall (n0 : nat) (y : R), (N <= n0)%nat -> Boule (/ 2) {| pos := / 2; cond_pos := pos_half_prf |} y -> Rabs (ps_atan y - sum_f_R0 (tg_alt (Ratan_seq y)) n0) < eps
n:nat
x:R
nN:(N <= n)%nat
b_y:Boule 0 {| pos := 1; cond_pos := Rlt_0_1 |} x
x0:0 > x

Rabs (- (ps_atan (- x) - sum_f_R0 (tg_alt (Ratan_seq (- x))) n)) < eps
eps:R
ep:0 < eps
N:nat
Pn:forall (n0 : nat) (y : R), (N <= n0)%nat -> Boule (/ 2) {| pos := / 2; cond_pos := pos_half_prf |} y -> Rabs (ps_atan y - sum_f_R0 (tg_alt (Ratan_seq y)) n0) < eps
n:nat
x:R
nN:(N <= n)%nat
b_y:Boule 0 {| pos := 1; cond_pos := Rlt_0_1 |} x
x0:0 > x
- (ps_atan (- x) - sum_f_R0 (tg_alt (Ratan_seq (- x))) n) = ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) n
eps:R
ep:0 < eps
N:nat
Pn:forall (n0 : nat) (y : R), (N <= n0)%nat -> Boule (/ 2) {| pos := / 2; cond_pos := pos_half_prf |} y -> Rabs (ps_atan y - sum_f_R0 (tg_alt (Ratan_seq y)) n0) < eps
n:nat
x:R
nN:(N <= n)%nat
b_y:Boule 0 {| pos := 1; cond_pos := Rlt_0_1 |} x
x0:0 > x

Rabs (ps_atan (- x) - sum_f_R0 (tg_alt (Ratan_seq (- x))) n) < eps
eps:R
ep:0 < eps
N:nat
Pn:forall (n0 : nat) (y : R), (N <= n0)%nat -> Boule (/ 2) {| pos := / 2; cond_pos := pos_half_prf |} y -> Rabs (ps_atan y - sum_f_R0 (tg_alt (Ratan_seq y)) n0) < eps
n:nat
x:R
nN:(N <= n)%nat
b_y:Boule 0 {| pos := 1; cond_pos := Rlt_0_1 |} x
x0:0 > x
- (ps_atan (- x) - sum_f_R0 (tg_alt (Ratan_seq (- x))) n) = ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) n
eps:R
ep:0 < eps
N:nat
Pn:forall (n0 : nat) (y : R), (N <= n0)%nat -> Boule (/ 2) {| pos := / 2; cond_pos := pos_half_prf |} y -> Rabs (ps_atan y - sum_f_R0 (tg_alt (Ratan_seq y)) n0) < eps
n:nat
x:R
nN:(N <= n)%nat
b_y:Boule 0 {| pos := 1; cond_pos := Rlt_0_1 |} x
x0:0 > x

Boule (/ 2) {| pos := / 2; cond_pos := pos_half_prf |} (- x)
eps:R
ep:0 < eps
N:nat
Pn:forall (n0 : nat) (y : R), (N <= n0)%nat -> Boule (/ 2) {| pos := / 2; cond_pos := pos_half_prf |} y -> Rabs (ps_atan y - sum_f_R0 (tg_alt (Ratan_seq y)) n0) < eps
n:nat
x:R
nN:(N <= n)%nat
b_y:Boule 0 {| pos := 1; cond_pos := Rlt_0_1 |} x
x0:0 > x
H:Boule (/ 2) {| pos := / 2; cond_pos := pos_half_prf |} (- x)
Rabs (ps_atan (- x) - sum_f_R0 (tg_alt (Ratan_seq (- x))) n) < eps
eps:R
ep:0 < eps
N:nat
Pn:forall (n0 : nat) (y : R), (N <= n0)%nat -> Boule (/ 2) {| pos := / 2; cond_pos := pos_half_prf |} y -> Rabs (ps_atan y - sum_f_R0 (tg_alt (Ratan_seq y)) n0) < eps
n:nat
x:R
nN:(N <= n)%nat
b_y:Boule 0 {| pos := 1; cond_pos := Rlt_0_1 |} x
x0:0 > x
- (ps_atan (- x) - sum_f_R0 (tg_alt (Ratan_seq (- x))) n) = ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) n
eps:R
ep:0 < eps
N:nat
Pn:forall (n0 : nat) (y : R), (N <= n0)%nat -> Boule (/ 2) {| pos := / 2; cond_pos := pos_half_prf |} y -> Rabs (ps_atan y - sum_f_R0 (tg_alt (Ratan_seq y)) n0) < eps
n:nat
x:R
nN:(N <= n)%nat
x0:0 > x
b_y:x - 0 < 1 /\ - (1) < x - 0

Rabs (- x - / 2) < / 2
eps:R
ep:0 < eps
N:nat
Pn:forall (n0 : nat) (y : R), (N <= n0)%nat -> Boule (/ 2) {| pos := / 2; cond_pos := pos_half_prf |} y -> Rabs (ps_atan y - sum_f_R0 (tg_alt (Ratan_seq y)) n0) < eps
n:nat
x:R
nN:(N <= n)%nat
b_y:Boule 0 {| pos := 1; cond_pos := Rlt_0_1 |} x
x0:0 > x
H:Boule (/ 2) {| pos := / 2; cond_pos := pos_half_prf |} (- x)
Rabs (ps_atan (- x) - sum_f_R0 (tg_alt (Ratan_seq (- x))) n) < eps
eps:R
ep:0 < eps
N:nat
Pn:forall (n0 : nat) (y : R), (N <= n0)%nat -> Boule (/ 2) {| pos := / 2; cond_pos := pos_half_prf |} y -> Rabs (ps_atan y - sum_f_R0 (tg_alt (Ratan_seq y)) n0) < eps
n:nat
x:R
nN:(N <= n)%nat
b_y:Boule 0 {| pos := 1; cond_pos := Rlt_0_1 |} x
x0:0 > x
- (ps_atan (- x) - sum_f_R0 (tg_alt (Ratan_seq (- x))) n) = ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) n
eps:R
ep:0 < eps
N:nat
Pn:forall (n0 : nat) (y : R), (N <= n0)%nat -> Boule (/ 2) {| pos := / 2; cond_pos := pos_half_prf |} y -> Rabs (ps_atan y - sum_f_R0 (tg_alt (Ratan_seq y)) n0) < eps
n:nat
x:R
nN:(N <= n)%nat
b_y:Boule 0 {| pos := 1; cond_pos := Rlt_0_1 |} x
x0:0 > x
H:Boule (/ 2) {| pos := / 2; cond_pos := pos_half_prf |} (- x)

Rabs (ps_atan (- x) - sum_f_R0 (tg_alt (Ratan_seq (- x))) n) < eps
eps:R
ep:0 < eps
N:nat
Pn:forall (n0 : nat) (y : R), (N <= n0)%nat -> Boule (/ 2) {| pos := / 2; cond_pos := pos_half_prf |} y -> Rabs (ps_atan y - sum_f_R0 (tg_alt (Ratan_seq y)) n0) < eps
n:nat
x:R
nN:(N <= n)%nat
b_y:Boule 0 {| pos := 1; cond_pos := Rlt_0_1 |} x
x0:0 > x
- (ps_atan (- x) - sum_f_R0 (tg_alt (Ratan_seq (- x))) n) = ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) n
eps:R
ep:0 < eps
N:nat
Pn:forall (n0 : nat) (y : R), (N <= n0)%nat -> Boule (/ 2) {| pos := / 2; cond_pos := pos_half_prf |} y -> Rabs (ps_atan y - sum_f_R0 (tg_alt (Ratan_seq y)) n0) < eps
n:nat
x:R
nN:(N <= n)%nat
b_y:Boule 0 {| pos := 1; cond_pos := Rlt_0_1 |} x
x0:0 > x

- (ps_atan (- x) - sum_f_R0 (tg_alt (Ratan_seq (- x))) n) = ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) n
eps:R
ep:0 < eps
N:nat
Pn:forall (n0 : nat) (y : R), (N <= n0)%nat -> Boule (/ 2) {| pos := / 2; cond_pos := pos_half_prf |} y -> Rabs (ps_atan y - sum_f_R0 (tg_alt (Ratan_seq y)) n0) < eps
n:nat
x:R
nN:(N <= n)%nat
b_y:Boule 0 {| pos := 1; cond_pos := Rlt_0_1 |} x
x0:0 > x

- - ps_atan x + - - - sum_f_R0 (tg_alt (Ratan_seq x)) n = ps_atan x + - sum_f_R0 (tg_alt (Ratan_seq x)) n
rewrite !Ropp_involutive; reflexivity. Qed.

forall n : nat, PI_tg n = Ratan_seq 1 n

forall n : nat, PI_tg n = Ratan_seq 1 n
n:nat

/ INR (2 * n + 1) = / INR (2 * n + 1)
reflexivity. Qed.

forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> forall x : R, -1 < x -> x < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x)) n - ps_atan x) < eps

forall eps : R, eps > 0 -> exists N : nat, forall n : nat, (n >= N)%nat -> forall x : R, -1 < x -> x < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x)) n - ps_atan x) < eps
eps:R
ep:eps > 0

exists N : nat, forall n : nat, (n >= N)%nat -> forall x : R, -1 < x -> x < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x)) n - ps_atan x) < eps
eps:R
ep:eps > 0
N1:nat
PN1:forall (n : nat) (y : R), (N1 <= n)%nat -> Boule 0 {| pos := 1; cond_pos := Rlt_0_1 |} y -> Rabs (ps_atan y - sum_f_R0 (tg_alt (Ratan_seq y)) n) < eps

exists N : nat, forall n : nat, (n >= N)%nat -> forall x : R, -1 < x -> x < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x)) n - ps_atan x) < eps
eps:R
ep:eps > 0
N1:nat
PN1:forall (n0 : nat) (y : R), (N1 <= n0)%nat -> Boule 0 {| pos := 1; cond_pos := Rlt_0_1 |} y -> Rabs (ps_atan y - sum_f_R0 (tg_alt (Ratan_seq y)) n0) < eps
n:nat
nN:(n >= N1)%nat
x:R
xm1:-1 < x
x1:x < 1

Rabs (ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) n) < eps
eps:R
ep:eps > 0
N1:nat
PN1:forall (n0 : nat) (y : R), (N1 <= n0)%nat -> Boule 0 {| pos := 1; cond_pos := Rlt_0_1 |} y -> Rabs (ps_atan y - sum_f_R0 (tg_alt (Ratan_seq y)) n0) < eps
n:nat
nN:(n >= N1)%nat
x:R
xm1:-1 < x
x1:x < 1

Boule 0 {| pos := 1; cond_pos := Rlt_0_1 |} x
unfold Boule; simpl; rewrite Rminus_0_r; apply Rabs_def1; assumption. Qed.

continuity (fun x : R => / (1 + x ^ 2))

continuity (fun x : R => / (1 + x ^ 2))

continuity (fun x : R => 1 + x ^ 2)

forall x : R, 1 + x ^ 2 <> 0

continuity (fun _ : R => 1)

continuity (fun x : R => x ^ 2)

forall x : R, 1 + x ^ 2 <> 0

continuity (fun x : R => x ^ 2)

forall x : R, 1 + x ^ 2 <> 0

forall x : R, 1 + x ^ 2 <> 0
x:R

x ^ 2 >= 0
x:R

x² >= 0
x:R
x² = x ^ 2
x:R

x² = x ^ 2
unfold Rsqr ; field. Qed.

forall x : R, -1 < x < 1 -> derivable_pt_lim ps_atan x ((fun y : R => / (1 + y ^ 2)) x)

forall x : R, -1 < x < 1 -> derivable_pt_lim ps_atan x ((fun y : R => / (1 + y ^ 2)) x)
x:R
x_encad:-1 < x < 1

derivable_pt_lim ps_atan x ((fun y : R => / (1 + y ^ 2)) x)
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1

derivable_pt_lim ps_atan x (/ (1 + x ^ 2))
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1

derivable_pt_lim ps_atan x ((fun u : R => / (1 + u ^ 2)) x)
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1
t:forall (fn fn' : nat -> R -> R) (f g : R -> R) (x0 c0 : R) (r0 : posreal), Boule c0 r0 x0 -> (forall (y : R) (n : nat), Boule c0 r0 y -> derivable_pt_lim (fn n) y (fn' n y)) -> (forall y : R, Boule c0 r0 y -> Un_cv (fun n : nat => fn n y) (f y)) -> CVU fn' g c0 r0 -> (forall y : R, Boule c0 r0 y -> continuity_pt g y) -> derivable_pt_lim f x0 (g x0)

derivable_pt_lim ps_atan x ((fun u : R => / (1 + u ^ 2)) x)
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1
t:forall (fn fn' : nat -> R -> R) (f g : R -> R) (x0 c0 : R) (r0 : posreal), Boule c0 r0 x0 -> (forall (y : R) (n : nat), Boule c0 r0 y -> derivable_pt_lim (fn n) y (fn' n y)) -> (forall y : R, Boule c0 r0 y -> Un_cv (fun n : nat => fn n y) (f y)) -> CVU fn' g c0 r0 -> (forall y : R, Boule c0 r0 y -> continuity_pt g y) -> derivable_pt_lim f x0 (g x0)

Boule c r x
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1
t:forall (fn fn' : nat -> R -> R) (f g : R -> R) (x0 c0 : R) (r0 : posreal), Boule c0 r0 x0 -> (forall (y : R) (n : nat), Boule c0 r0 y -> derivable_pt_lim (fn n) y (fn' n y)) -> (forall y : R, Boule c0 r0 y -> Un_cv (fun n : nat => fn n y) (f y)) -> CVU fn' g c0 r0 -> (forall y : R, Boule c0 r0 y -> continuity_pt g y) -> derivable_pt_lim f x0 (g x0)
forall (y : R) (n : nat), Boule c r y -> derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) n) y (sum_f_R0 (tg_alt (Datan_seq y)) n)
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1
t:forall (fn fn' : nat -> R -> R) (f g : R -> R) (x0 c0 : R) (r0 : posreal), Boule c0 r0 x0 -> (forall (y : R) (n : nat), Boule c0 r0 y -> derivable_pt_lim (fn n) y (fn' n y)) -> (forall y : R, Boule c0 r0 y -> Un_cv (fun n : nat => fn n y) (f y)) -> CVU fn' g c0 r0 -> (forall y : R, Boule c0 r0 y -> continuity_pt g y) -> derivable_pt_lim f x0 (g x0)
forall y : R, Boule c r y -> Un_cv (fun n : nat => sum_f_R0 (tg_alt (Ratan_seq y)) n) (ps_atan y)
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1
t:forall (fn fn' : nat -> R -> R) (f g : R -> R) (x0 c0 : R) (r0 : posreal), Boule c0 r0 x0 -> (forall (y : R) (n : nat), Boule c0 r0 y -> derivable_pt_lim (fn n) y (fn' n y)) -> (forall y : R, Boule c0 r0 y -> Un_cv (fun n : nat => fn n y) (f y)) -> CVU fn' g c0 r0 -> (forall y : R, Boule c0 r0 y -> continuity_pt g y) -> derivable_pt_lim f x0 (g x0)
CVU (fun (N : nat) (x0 : R) => sum_f_R0 (tg_alt (Datan_seq x0)) N) (fun u : R => / (1 + u ^ 2)) c r
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1
t:forall (fn fn' : nat -> R -> R) (f g : R -> R) (x0 c0 : R) (r0 : posreal), Boule c0 r0 x0 -> (forall (y : R) (n : nat), Boule c0 r0 y -> derivable_pt_lim (fn n) y (fn' n y)) -> (forall y : R, Boule c0 r0 y -> Un_cv (fun n : nat => fn n y) (f y)) -> CVU fn' g c0 r0 -> (forall y : R, Boule c0 r0 y -> continuity_pt g y) -> derivable_pt_lim f x0 (g x0)
forall y : R, Boule c r y -> continuity_pt (fun u : R => / (1 + u ^ 2)) y
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1
t:forall (fn fn' : nat -> R -> R) (f g : R -> R) (x0 c0 : R) (r0 : posreal), Boule c0 r0 x0 -> (forall (y : R) (n : nat), Boule c0 r0 y -> derivable_pt_lim (fn n) y (fn' n y)) -> (forall y : R, Boule c0 r0 y -> Un_cv (fun n : nat => fn n y) (f y)) -> CVU fn' g c0 r0 -> (forall y : R, Boule c0 r0 y -> continuity_pt g y) -> derivable_pt_lim f x0 (g x0)

forall (y : R) (n : nat), Boule c r y -> derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) n) y (sum_f_R0 (tg_alt (Datan_seq y)) n)
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1
t:forall (fn fn' : nat -> R -> R) (f g : R -> R) (x0 c0 : R) (r0 : posreal), Boule c0 r0 x0 -> (forall (y : R) (n : nat), Boule c0 r0 y -> derivable_pt_lim (fn n) y (fn' n y)) -> (forall y : R, Boule c0 r0 y -> Un_cv (fun n : nat => fn n y) (f y)) -> CVU fn' g c0 r0 -> (forall y : R, Boule c0 r0 y -> continuity_pt g y) -> derivable_pt_lim f x0 (g x0)
forall y : R, Boule c r y -> Un_cv (fun n : nat => sum_f_R0 (tg_alt (Ratan_seq y)) n) (ps_atan y)
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1
t:forall (fn fn' : nat -> R -> R) (f g : R -> R) (x0 c0 : R) (r0 : posreal), Boule c0 r0 x0 -> (forall (y : R) (n : nat), Boule c0 r0 y -> derivable_pt_lim (fn n) y (fn' n y)) -> (forall y : R, Boule c0 r0 y -> Un_cv (fun n : nat => fn n y) (f y)) -> CVU fn' g c0 r0 -> (forall y : R, Boule c0 r0 y -> continuity_pt g y) -> derivable_pt_lim f x0 (g x0)
CVU (fun (N : nat) (x0 : R) => sum_f_R0 (tg_alt (Datan_seq x0)) N) (fun u : R => / (1 + u ^ 2)) c r
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1
t:forall (fn fn' : nat -> R -> R) (f g : R -> R) (x0 c0 : R) (r0 : posreal), Boule c0 r0 x0 -> (forall (y : R) (n : nat), Boule c0 r0 y -> derivable_pt_lim (fn n) y (fn' n y)) -> (forall y : R, Boule c0 r0 y -> Un_cv (fun n : nat => fn n y) (f y)) -> CVU fn' g c0 r0 -> (forall y : R, Boule c0 r0 y -> continuity_pt g y) -> derivable_pt_lim f x0 (g x0)
forall y : R, Boule c r y -> continuity_pt (fun u : R => / (1 + u ^ 2)) y
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1
t:forall (fn fn' : nat -> R -> R) (f g : R -> R) (x0 c0 : R) (r0 : posreal), Boule c0 r0 x0 -> (forall (y0 : R) (n : nat), Boule c0 r0 y0 -> derivable_pt_lim (fn n) y0 (fn' n y0)) -> (forall y0 : R, Boule c0 r0 y0 -> Un_cv (fun n : nat => fn n y0) (f y0)) -> CVU fn' g c0 r0 -> (forall y0 : R, Boule c0 r0 y0 -> continuity_pt g y0) -> derivable_pt_lim f x0 (g x0)
y:R
N:nat
H:y - c < r
H0:- r < y - c

derivable_pt_lim (fun x0 : R => sum_f_R0 (tg_alt (Ratan_seq x0)) N) y (sum_f_R0 (tg_alt (Datan_seq y)) N)
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1
t:forall (fn fn' : nat -> R -> R) (f g : R -> R) (x0 c0 : R) (r0 : posreal), Boule c0 r0 x0 -> (forall (y : R) (n : nat), Boule c0 r0 y -> derivable_pt_lim (fn n) y (fn' n y)) -> (forall y : R, Boule c0 r0 y -> Un_cv (fun n : nat => fn n y) (f y)) -> CVU fn' g c0 r0 -> (forall y : R, Boule c0 r0 y -> continuity_pt g y) -> derivable_pt_lim f x0 (g x0)
forall y : R, Boule c r y -> Un_cv (fun n : nat => sum_f_R0 (tg_alt (Ratan_seq y)) n) (ps_atan y)
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1
t:forall (fn fn' : nat -> R -> R) (f g : R -> R) (x0 c0 : R) (r0 : posreal), Boule c0 r0 x0 -> (forall (y : R) (n : nat), Boule c0 r0 y -> derivable_pt_lim (fn n) y (fn' n y)) -> (forall y : R, Boule c0 r0 y -> Un_cv (fun n : nat => fn n y) (f y)) -> CVU fn' g c0 r0 -> (forall y : R, Boule c0 r0 y -> continuity_pt g y) -> derivable_pt_lim f x0 (g x0)
CVU (fun (N : nat) (x0 : R) => sum_f_R0 (tg_alt (Datan_seq x0)) N) (fun u : R => / (1 + u ^ 2)) c r
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1
t:forall (fn fn' : nat -> R -> R) (f g : R -> R) (x0 c0 : R) (r0 : posreal), Boule c0 r0 x0 -> (forall (y : R) (n : nat), Boule c0 r0 y -> derivable_pt_lim (fn n) y (fn' n y)) -> (forall y : R, Boule c0 r0 y -> Un_cv (fun n : nat => fn n y) (f y)) -> CVU fn' g c0 r0 -> (forall y : R, Boule c0 r0 y -> continuity_pt g y) -> derivable_pt_lim f x0 (g x0)
forall y : R, Boule c r y -> continuity_pt (fun u : R => / (1 + u ^ 2)) y
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1
t:forall (fn fn' : nat -> R -> R) (f g : R -> R) (x0 c0 : R) (r0 : posreal), Boule c0 r0 x0 -> (forall (y0 : R) (n : nat), Boule c0 r0 y0 -> derivable_pt_lim (fn n) y0 (fn' n y0)) -> (forall y0 : R, Boule c0 r0 y0 -> Un_cv (fun n : nat => fn n y0) (f y0)) -> CVU fn' g c0 r0 -> (forall y0 : R, Boule c0 r0 y0 -> continuity_pt g y0) -> derivable_pt_lim f x0 (g x0)
y:R
N:nat
H:y - c < r
H0:- r < y - c

-1 <= y
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1
t:forall (fn fn' : nat -> R -> R) (f g : R -> R) (x0 c0 : R) (r0 : posreal), Boule c0 r0 x0 -> (forall (y0 : R) (n : nat), Boule c0 r0 y0 -> derivable_pt_lim (fn n) y0 (fn' n y0)) -> (forall y0 : R, Boule c0 r0 y0 -> Un_cv (fun n : nat => fn n y0) (f y0)) -> CVU fn' g c0 r0 -> (forall y0 : R, Boule c0 r0 y0 -> continuity_pt g y0) -> derivable_pt_lim f x0 (g x0)
y:R
N:nat
H:y - c < r
H0:- r < y - c
y < 1
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1
t:forall (fn fn' : nat -> R -> R) (f g : R -> R) (x0 c0 : R) (r0 : posreal), Boule c0 r0 x0 -> (forall (y : R) (n : nat), Boule c0 r0 y -> derivable_pt_lim (fn n) y (fn' n y)) -> (forall y : R, Boule c0 r0 y -> Un_cv (fun n : nat => fn n y) (f y)) -> CVU fn' g c0 r0 -> (forall y : R, Boule c0 r0 y -> continuity_pt g y) -> derivable_pt_lim f x0 (g x0)
forall y : R, Boule c r y -> Un_cv (fun n : nat => sum_f_R0 (tg_alt (Ratan_seq y)) n) (ps_atan y)
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1
t:forall (fn fn' : nat -> R -> R) (f g : R -> R) (x0 c0 : R) (r0 : posreal), Boule c0 r0 x0 -> (forall (y : R) (n : nat), Boule c0 r0 y -> derivable_pt_lim (fn n) y (fn' n y)) -> (forall y : R, Boule c0 r0 y -> Un_cv (fun n : nat => fn n y) (f y)) -> CVU fn' g c0 r0 -> (forall y : R, Boule c0 r0 y -> continuity_pt g y) -> derivable_pt_lim f x0 (g x0)
CVU (fun (N : nat) (x0 : R) => sum_f_R0 (tg_alt (Datan_seq x0)) N) (fun u : R => / (1 + u ^ 2)) c r
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1
t:forall (fn fn' : nat -> R -> R) (f g : R -> R) (x0 c0 : R) (r0 : posreal), Boule c0 r0 x0 -> (forall (y : R) (n : nat), Boule c0 r0 y -> derivable_pt_lim (fn n) y (fn' n y)) -> (forall y : R, Boule c0 r0 y -> Un_cv (fun n : nat => fn n y) (f y)) -> CVU fn' g c0 r0 -> (forall y : R, Boule c0 r0 y -> continuity_pt g y) -> derivable_pt_lim f x0 (g x0)
forall y : R, Boule c r y -> continuity_pt (fun u : R => / (1 + u ^ 2)) y
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1
t:forall (fn fn' : nat -> R -> R) (f g : R -> R) (x0 c0 : R) (r0 : posreal), Boule c0 r0 x0 -> (forall (y0 : R) (n : nat), Boule c0 r0 y0 -> derivable_pt_lim (fn n) y0 (fn' n y0)) -> (forall y0 : R, Boule c0 r0 y0 -> Un_cv (fun n : nat => fn n y0) (f y0)) -> CVU fn' g c0 r0 -> (forall y0 : R, Boule c0 r0 y0 -> continuity_pt g y0) -> derivable_pt_lim f x0 (g x0)
y:R
N:nat
H:y - c < r
H0:- r < y - c

y < 1
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1
t:forall (fn fn' : nat -> R -> R) (f g : R -> R) (x0 c0 : R) (r0 : posreal), Boule c0 r0 x0 -> (forall (y : R) (n : nat), Boule c0 r0 y -> derivable_pt_lim (fn n) y (fn' n y)) -> (forall y : R, Boule c0 r0 y -> Un_cv (fun n : nat => fn n y) (f y)) -> CVU fn' g c0 r0 -> (forall y : R, Boule c0 r0 y -> continuity_pt g y) -> derivable_pt_lim f x0 (g x0)
forall y : R, Boule c r y -> Un_cv (fun n : nat => sum_f_R0 (tg_alt (Ratan_seq y)) n) (ps_atan y)
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1
t:forall (fn fn' : nat -> R -> R) (f g : R -> R) (x0 c0 : R) (r0 : posreal), Boule c0 r0 x0 -> (forall (y : R) (n : nat), Boule c0 r0 y -> derivable_pt_lim (fn n) y (fn' n y)) -> (forall y : R, Boule c0 r0 y -> Un_cv (fun n : nat => fn n y) (f y)) -> CVU fn' g c0 r0 -> (forall y : R, Boule c0 r0 y -> continuity_pt g y) -> derivable_pt_lim f x0 (g x0)
CVU (fun (N : nat) (x0 : R) => sum_f_R0 (tg_alt (Datan_seq x0)) N) (fun u : R => / (1 + u ^ 2)) c r
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1
t:forall (fn fn' : nat -> R -> R) (f g : R -> R) (x0 c0 : R) (r0 : posreal), Boule c0 r0 x0 -> (forall (y : R) (n : nat), Boule c0 r0 y -> derivable_pt_lim (fn n) y (fn' n y)) -> (forall y : R, Boule c0 r0 y -> Un_cv (fun n : nat => fn n y) (f y)) -> CVU fn' g c0 r0 -> (forall y : R, Boule c0 r0 y -> continuity_pt g y) -> derivable_pt_lim f x0 (g x0)
forall y : R, Boule c r y -> continuity_pt (fun u : R => / (1 + u ^ 2)) y
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1
t:forall (fn fn' : nat -> R -> R) (f g : R -> R) (x0 c0 : R) (r0 : posreal), Boule c0 r0 x0 -> (forall (y : R) (n : nat), Boule c0 r0 y -> derivable_pt_lim (fn n) y (fn' n y)) -> (forall y : R, Boule c0 r0 y -> Un_cv (fun n : nat => fn n y) (f y)) -> CVU fn' g c0 r0 -> (forall y : R, Boule c0 r0 y -> continuity_pt g y) -> derivable_pt_lim f x0 (g x0)

forall y : R, Boule c r y -> Un_cv (fun n : nat => sum_f_R0 (tg_alt (Ratan_seq y)) n) (ps_atan y)
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1
t:forall (fn fn' : nat -> R -> R) (f g : R -> R) (x0 c0 : R) (r0 : posreal), Boule c0 r0 x0 -> (forall (y : R) (n : nat), Boule c0 r0 y -> derivable_pt_lim (fn n) y (fn' n y)) -> (forall y : R, Boule c0 r0 y -> Un_cv (fun n : nat => fn n y) (f y)) -> CVU fn' g c0 r0 -> (forall y : R, Boule c0 r0 y -> continuity_pt g y) -> derivable_pt_lim f x0 (g x0)
CVU (fun (N : nat) (x0 : R) => sum_f_R0 (tg_alt (Datan_seq x0)) N) (fun u : R => / (1 + u ^ 2)) c r
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1
t:forall (fn fn' : nat -> R -> R) (f g : R -> R) (x0 c0 : R) (r0 : posreal), Boule c0 r0 x0 -> (forall (y : R) (n : nat), Boule c0 r0 y -> derivable_pt_lim (fn n) y (fn' n y)) -> (forall y : R, Boule c0 r0 y -> Un_cv (fun n : nat => fn n y) (f y)) -> CVU fn' g c0 r0 -> (forall y : R, Boule c0 r0 y -> continuity_pt g y) -> derivable_pt_lim f x0 (g x0)
forall y : R, Boule c r y -> continuity_pt (fun u : R => / (1 + u ^ 2)) y
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1
t:forall (fn fn' : nat -> R -> R) (f g : R -> R) (x0 c0 : R) (r0 : posreal), Boule c0 r0 x0 -> (forall (y0 : R) (n : nat), Boule c0 r0 y0 -> derivable_pt_lim (fn n) y0 (fn' n y0)) -> (forall y0 : R, Boule c0 r0 y0 -> Un_cv (fun n : nat => fn n y0) (f y0)) -> CVU fn' g c0 r0 -> (forall y0 : R, Boule c0 r0 y0 -> continuity_pt g y0) -> derivable_pt_lim f x0 (g x0)
y:R
H:y - c < r
H0:- r < y - c

Un_cv (fun n : nat => sum_f_R0 (tg_alt (Ratan_seq y)) n) (ps_atan y)
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1
t:forall (fn fn' : nat -> R -> R) (f g : R -> R) (x0 c0 : R) (r0 : posreal), Boule c0 r0 x0 -> (forall (y : R) (n : nat), Boule c0 r0 y -> derivable_pt_lim (fn n) y (fn' n y)) -> (forall y : R, Boule c0 r0 y -> Un_cv (fun n : nat => fn n y) (f y)) -> CVU fn' g c0 r0 -> (forall y : R, Boule c0 r0 y -> continuity_pt g y) -> derivable_pt_lim f x0 (g x0)
CVU (fun (N : nat) (x0 : R) => sum_f_R0 (tg_alt (Datan_seq x0)) N) (fun u : R => / (1 + u ^ 2)) c r
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1
t:forall (fn fn' : nat -> R -> R) (f g : R -> R) (x0 c0 : R) (r0 : posreal), Boule c0 r0 x0 -> (forall (y : R) (n : nat), Boule c0 r0 y -> derivable_pt_lim (fn n) y (fn' n y)) -> (forall y : R, Boule c0 r0 y -> Un_cv (fun n : nat => fn n y) (f y)) -> CVU fn' g c0 r0 -> (forall y : R, Boule c0 r0 y -> continuity_pt g y) -> derivable_pt_lim f x0 (g x0)
forall y : R, Boule c r y -> continuity_pt (fun u : R => / (1 + u ^ 2)) y
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1
t:forall (fn fn' : nat -> R -> R) (f g : R -> R) (x0 c0 : R) (r0 : posreal), Boule c0 r0 x0 -> (forall (y0 : R) (n : nat), Boule c0 r0 y0 -> derivable_pt_lim (fn n) y0 (fn' n y0)) -> (forall y0 : R, Boule c0 r0 y0 -> Un_cv (fun n : nat => fn n y0) (f y0)) -> CVU fn' g c0 r0 -> (forall y0 : R, Boule c0 r0 y0 -> continuity_pt g y0) -> derivable_pt_lim f x0 (g x0)
y:R
H:y - c < r
H0:- r < y - c
y_gt_0:-1 < y

Un_cv (fun n : nat => sum_f_R0 (tg_alt (Ratan_seq y)) n) (ps_atan y)
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1
t:forall (fn fn' : nat -> R -> R) (f g : R -> R) (x0 c0 : R) (r0 : posreal), Boule c0 r0 x0 -> (forall (y : R) (n : nat), Boule c0 r0 y -> derivable_pt_lim (fn n) y (fn' n y)) -> (forall y : R, Boule c0 r0 y -> Un_cv (fun n : nat => fn n y) (f y)) -> CVU fn' g c0 r0 -> (forall y : R, Boule c0 r0 y -> continuity_pt g y) -> derivable_pt_lim f x0 (g x0)
CVU (fun (N : nat) (x0 : R) => sum_f_R0 (tg_alt (Datan_seq x0)) N) (fun u : R => / (1 + u ^ 2)) c r
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1
t:forall (fn fn' : nat -> R -> R) (f g : R -> R) (x0 c0 : R) (r0 : posreal), Boule c0 r0 x0 -> (forall (y : R) (n : nat), Boule c0 r0 y -> derivable_pt_lim (fn n) y (fn' n y)) -> (forall y : R, Boule c0 r0 y -> Un_cv (fun n : nat => fn n y) (f y)) -> CVU fn' g c0 r0 -> (forall y : R, Boule c0 r0 y -> continuity_pt g y) -> derivable_pt_lim f x0 (g x0)
forall y : R, Boule c r y -> continuity_pt (fun u : R => / (1 + u ^ 2)) y
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1
t:forall (fn fn' : nat -> R -> R) (f g : R -> R) (x0 c0 : R) (r0 : posreal), Boule c0 r0 x0 -> (forall (y0 : R) (n : nat), Boule c0 r0 y0 -> derivable_pt_lim (fn n) y0 (fn' n y0)) -> (forall y0 : R, Boule c0 r0 y0 -> Un_cv (fun n : nat => fn n y0) (f y0)) -> CVU fn' g c0 r0 -> (forall y0 : R, Boule c0 r0 y0 -> continuity_pt g y0) -> derivable_pt_lim f x0 (g x0)
y:R
H:y - c < r
H0:- r < y - c
y_gt_0:-1 < y
y_lt_1:y < 1

Un_cv (fun n : nat => sum_f_R0 (tg_alt (Ratan_seq y)) n) (ps_atan y)
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1
t:forall (fn fn' : nat -> R -> R) (f g : R -> R) (x0 c0 : R) (r0 : posreal), Boule c0 r0 x0 -> (forall (y : R) (n : nat), Boule c0 r0 y -> derivable_pt_lim (fn n) y (fn' n y)) -> (forall y : R, Boule c0 r0 y -> Un_cv (fun n : nat => fn n y) (f y)) -> CVU fn' g c0 r0 -> (forall y : R, Boule c0 r0 y -> continuity_pt g y) -> derivable_pt_lim f x0 (g x0)
CVU (fun (N : nat) (x0 : R) => sum_f_R0 (tg_alt (Datan_seq x0)) N) (fun u : R => / (1 + u ^ 2)) c r
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1
t:forall (fn fn' : nat -> R -> R) (f g : R -> R) (x0 c0 : R) (r0 : posreal), Boule c0 r0 x0 -> (forall (y : R) (n : nat), Boule c0 r0 y -> derivable_pt_lim (fn n) y (fn' n y)) -> (forall y : R, Boule c0 r0 y -> Un_cv (fun n : nat => fn n y) (f y)) -> CVU fn' g c0 r0 -> (forall y : R, Boule c0 r0 y -> continuity_pt g y) -> derivable_pt_lim f x0 (g x0)
forall y : R, Boule c r y -> continuity_pt (fun u : R => / (1 + u ^ 2)) y
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1
t:forall (fn fn' : nat -> R -> R) (f g : R -> R) (x0 c0 : R) (r0 : posreal), Boule c0 r0 x0 -> (forall (y0 : R) (n : nat), Boule c0 r0 y0 -> derivable_pt_lim (fn n) y0 (fn' n y0)) -> (forall y0 : R, Boule c0 r0 y0 -> Un_cv (fun n : nat => fn n y0) (f y0)) -> CVU fn' g c0 r0 -> (forall y0 : R, Boule c0 r0 y0 -> continuity_pt g y0) -> derivable_pt_lim f x0 (g x0)
y:R
H:y - c < r
H0:- r < y - c
y_gt_0:-1 < y
y_lt_1:y < 1
eps:R
eps_pos:eps > 0

forall x0 : nat, (forall n : nat, (n >= x0)%nat -> forall x1 : R, -1 < x1 -> x1 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x1)) n - ps_atan x1) < eps) -> exists N : nat, forall n : nat, (n >= N)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq y)) n) (ps_atan y) < eps
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1
t:forall (fn fn' : nat -> R -> R) (f g : R -> R) (x0 c0 : R) (r0 : posreal), Boule c0 r0 x0 -> (forall (y : R) (n : nat), Boule c0 r0 y -> derivable_pt_lim (fn n) y (fn' n y)) -> (forall y : R, Boule c0 r0 y -> Un_cv (fun n : nat => fn n y) (f y)) -> CVU fn' g c0 r0 -> (forall y : R, Boule c0 r0 y -> continuity_pt g y) -> derivable_pt_lim f x0 (g x0)
CVU (fun (N : nat) (x0 : R) => sum_f_R0 (tg_alt (Datan_seq x0)) N) (fun u : R => / (1 + u ^ 2)) c r
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1
t:forall (fn fn' : nat -> R -> R) (f g : R -> R) (x0 c0 : R) (r0 : posreal), Boule c0 r0 x0 -> (forall (y : R) (n : nat), Boule c0 r0 y -> derivable_pt_lim (fn n) y (fn' n y)) -> (forall y : R, Boule c0 r0 y -> Un_cv (fun n : nat => fn n y) (f y)) -> CVU fn' g c0 r0 -> (forall y : R, Boule c0 r0 y -> continuity_pt g y) -> derivable_pt_lim f x0 (g x0)
forall y : R, Boule c r y -> continuity_pt (fun u : R => / (1 + u ^ 2)) y
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1
t:forall (fn fn' : nat -> R -> R) (f g : R -> R) (x0 c0 : R) (r0 : posreal), Boule c0 r0 x0 -> (forall (y : R) (n : nat), Boule c0 r0 y -> derivable_pt_lim (fn n) y (fn' n y)) -> (forall y : R, Boule c0 r0 y -> Un_cv (fun n : nat => fn n y) (f y)) -> CVU fn' g c0 r0 -> (forall y : R, Boule c0 r0 y -> continuity_pt g y) -> derivable_pt_lim f x0 (g x0)

CVU (fun (N : nat) (x0 : R) => sum_f_R0 (tg_alt (Datan_seq x0)) N) (fun u : R => / (1 + u ^ 2)) c r
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1
t:forall (fn fn' : nat -> R -> R) (f g : R -> R) (x0 c0 : R) (r0 : posreal), Boule c0 r0 x0 -> (forall (y : R) (n : nat), Boule c0 r0 y -> derivable_pt_lim (fn n) y (fn' n y)) -> (forall y : R, Boule c0 r0 y -> Un_cv (fun n : nat => fn n y) (f y)) -> CVU fn' g c0 r0 -> (forall y : R, Boule c0 r0 y -> continuity_pt g y) -> derivable_pt_lim f x0 (g x0)
forall y : R, Boule c r y -> continuity_pt (fun u : R => / (1 + u ^ 2)) y
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1
t:forall (fn fn' : nat -> R -> R) (f g : R -> R) (x0 c0 : R) (r0 : posreal), Boule c0 r0 x0 -> (forall (y : R) (n : nat), Boule c0 r0 y -> derivable_pt_lim (fn n) y (fn' n y)) -> (forall y : R, Boule c0 r0 y -> Un_cv (fun n : nat => fn n y) (f y)) -> CVU fn' g c0 r0 -> (forall y : R, Boule c0 r0 y -> continuity_pt g y) -> derivable_pt_lim f x0 (g x0)

Rabs c + r < 1
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1
t:forall (fn fn' : nat -> R -> R) (f g : R -> R) (x0 c0 : R) (r0 : posreal), Boule c0 r0 x0 -> (forall (y : R) (n : nat), Boule c0 r0 y -> derivable_pt_lim (fn n) y (fn' n y)) -> (forall y : R, Boule c0 r0 y -> Un_cv (fun n : nat => fn n y) (f y)) -> CVU fn' g c0 r0 -> (forall y : R, Boule c0 r0 y -> continuity_pt g y) -> derivable_pt_lim f x0 (g x0)
forall y : R, Boule c r y -> continuity_pt (fun u : R => / (1 + u ^ 2)) y
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1
t:forall (fn fn' : nat -> R -> R) (f g : R -> R) (x0 c0 : R) (r0 : posreal), Boule c0 r0 x0 -> (forall (y : R) (n : nat), Boule c0 r0 y -> derivable_pt_lim (fn n) y (fn' n y)) -> (forall y : R, Boule c0 r0 y -> Un_cv (fun n : nat => fn n y) (f y)) -> CVU fn' g c0 r0 -> (forall y : R, Boule c0 r0 y -> continuity_pt g y) -> derivable_pt_lim f x0 (g x0)

Rabs c + r < 1
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1
t:forall (fn fn' : nat -> R -> R) (f g : R -> R) (x0 c0 : R) (r0 : posreal), Boule c0 r0 x0 -> (forall (y : R) (n : nat), Boule c0 r0 y -> derivable_pt_lim (fn n) y (fn' n y)) -> (forall y : R, Boule c0 r0 y -> Un_cv (fun n : nat => fn n y) (f y)) -> CVU fn' g c0 r0 -> (forall y : R, Boule c0 r0 y -> continuity_pt g y) -> derivable_pt_lim f x0 (g x0)
forall y : R, Boule c r y -> continuity_pt (fun u : R => / (1 + u ^ 2)) y
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1
t:forall (fn fn' : nat -> R -> R) (f g : R -> R) (x0 c0 : R) (r0 : posreal), Boule c0 r0 x0 -> (forall (y : R) (n : nat), Boule c0 r0 y -> derivable_pt_lim (fn n) y (fn' n y)) -> (forall y : R, Boule c0 r0 y -> Un_cv (fun n : nat => fn n y) (f y)) -> CVU fn' g c0 r0 -> (forall y : R, Boule c0 r0 y -> continuity_pt g y) -> derivable_pt_lim f x0 (g x0)

Rabs c + r < 1
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1
t:forall (fn fn' : nat -> R -> R) (f g : R -> R) (x0 c0 : R) (r0 : posreal), Boule c0 r0 x0 -> (forall (y : R) (n : nat), Boule c0 r0 y -> derivable_pt_lim (fn n) y (fn' n y)) -> (forall y : R, Boule c0 r0 y -> Un_cv (fun n : nat => fn n y) (f y)) -> CVU fn' g c0 r0 -> (forall y : R, Boule c0 r0 y -> continuity_pt g y) -> derivable_pt_lim f x0 (g x0)
forall y : R, Boule c r y -> continuity_pt (fun u : R => / (1 + u ^ 2)) y
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1
t:forall (fn fn' : nat -> R -> R) (f g : R -> R) (x0 c0 : R) (r0 : posreal), Boule c0 r0 x0 -> (forall (y : R) (n : nat), Boule c0 r0 y -> derivable_pt_lim (fn n) y (fn' n y)) -> (forall y : R, Boule c0 r0 y -> Un_cv (fun n : nat => fn n y) (f y)) -> CVU fn' g c0 r0 -> (forall y : R, Boule c0 r0 y -> continuity_pt g y) -> derivable_pt_lim f x0 (g x0)

Rabs c + r < 1
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1
t:forall (fn fn' : nat -> R -> R) (f g : R -> R) (x0 c0 : R) (r0 : posreal), Boule c0 r0 x0 -> (forall (y : R) (n : nat), Boule c0 r0 y -> derivable_pt_lim (fn n) y (fn' n y)) -> (forall y : R, Boule c0 r0 y -> Un_cv (fun n : nat => fn n y) (f y)) -> CVU fn' g c0 r0 -> (forall y : R, Boule c0 r0 y -> continuity_pt g y) -> derivable_pt_lim f x0 (g x0)
forall y : R, Boule c r y -> continuity_pt (fun u : R => / (1 + u ^ 2)) y
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1
t:forall (fn fn' : nat -> R -> R) (f g : R -> R) (x0 c0 : R) (r0 : posreal), Boule c0 r0 x0 -> (forall (y : R) (n : nat), Boule c0 r0 y -> derivable_pt_lim (fn n) y (fn' n y)) -> (forall y : R, Boule c0 r0 y -> Un_cv (fun n : nat => fn n y) (f y)) -> CVU fn' g c0 r0 -> (forall y : R, Boule c0 r0 y -> continuity_pt g y) -> derivable_pt_lim f x0 (g x0)

Rabs c < 1 - r
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1
t:forall (fn fn' : nat -> R -> R) (f g : R -> R) (x0 c0 : R) (r0 : posreal), Boule c0 r0 x0 -> (forall (y : R) (n : nat), Boule c0 r0 y -> derivable_pt_lim (fn n) y (fn' n y)) -> (forall y : R, Boule c0 r0 y -> Un_cv (fun n : nat => fn n y) (f y)) -> CVU fn' g c0 r0 -> (forall y : R, Boule c0 r0 y -> continuity_pt g y) -> derivable_pt_lim f x0 (g x0)
H:Rabs c < 1 - r
Rabs c + r < 1
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1
t:forall (fn fn' : nat -> R -> R) (f g : R -> R) (x0 c0 : R) (r0 : posreal), Boule c0 r0 x0 -> (forall (y : R) (n : nat), Boule c0 r0 y -> derivable_pt_lim (fn n) y (fn' n y)) -> (forall y : R, Boule c0 r0 y -> Un_cv (fun n : nat => fn n y) (f y)) -> CVU fn' g c0 r0 -> (forall y : R, Boule c0 r0 y -> continuity_pt g y) -> derivable_pt_lim f x0 (g x0)
forall y : R, Boule c r y -> continuity_pt (fun u : R => / (1 + u ^ 2)) y
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1
t:forall (fn fn' : nat -> R -> R) (f g : R -> R) (x0 c0 : R) (r0 : posreal), Boule c0 r0 x0 -> (forall (y : R) (n : nat), Boule c0 r0 y -> derivable_pt_lim (fn n) y (fn' n y)) -> (forall y : R, Boule c0 r0 y -> Un_cv (fun n : nat => fn n y) (f y)) -> CVU fn' g c0 r0 -> (forall y : R, Boule c0 r0 y -> continuity_pt g y) -> derivable_pt_lim f x0 (g x0)
H:Rabs c < 1 - r

Rabs c + r < 1
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1
t:forall (fn fn' : nat -> R -> R) (f g : R -> R) (x0 c0 : R) (r0 : posreal), Boule c0 r0 x0 -> (forall (y : R) (n : nat), Boule c0 r0 y -> derivable_pt_lim (fn n) y (fn' n y)) -> (forall y : R, Boule c0 r0 y -> Un_cv (fun n : nat => fn n y) (f y)) -> CVU fn' g c0 r0 -> (forall y : R, Boule c0 r0 y -> continuity_pt g y) -> derivable_pt_lim f x0 (g x0)
forall y : R, Boule c r y -> continuity_pt (fun u : R => / (1 + u ^ 2)) y
x:R
x_encad:-1 < x < 1
c:R
r:posreal
Pcr1:Boule c r x
P1:-1 < c - r
P2:c + r < 1
t:forall (fn fn' : nat -> R -> R) (f g : R -> R) (x0 c0 : R) (r0 : posreal), Boule c0 r0 x0 -> (forall (y : R) (n : nat), Boule c0 r0 y -> derivable_pt_lim (fn n) y (fn' n y)) -> (forall y : R, Boule c0 r0 y -> Un_cv (fun n : nat => fn n y) (f y)) -> CVU fn' g c0 r0 -> (forall y : R, Boule c0 r0 y -> continuity_pt g y) -> derivable_pt_lim f x0 (g x0)

forall y : R, Boule c r y -> continuity_pt (fun u : R => / (1 + u ^ 2)) y
intros; apply Datan_continuity. Qed.

forall x : R, -1 < x < 1 -> derivable_pt ps_atan x

forall x : R, -1 < x < 1 -> derivable_pt ps_atan x
x:R
x_encad:-1 < x < 1

derivable_pt ps_atan x
exists (/(1+x^2)) ; apply derivable_pt_lim_ps_atan; assumption. Qed.

forall eps : R, eps > 0 -> exists alp : R, alp > 0 /\ (forall x : R, x < 1 -> 0 < x -> R_dist x 1 < alp -> dist R_met (ps_atan x) (Alt_PI / 4) < eps)

forall eps : R, eps > 0 -> exists alp : R, alp > 0 /\ (forall x : R, x < 1 -> 0 < x -> R_dist x 1 < alp -> dist R_met (ps_atan x) (Alt_PI / 4) < eps)
eps:R
eps_pos:eps > 0

exists alp : R, alp > 0 /\ (forall x : R, x < 1 -> 0 < x -> R_dist x 1 < alp -> dist R_met (ps_atan x) (Alt_PI / 4) < eps)
eps:R
eps_pos:eps > 0
eps_3_pos:eps / 3 > 0

exists alp : R, alp > 0 /\ (forall x : R, x < 1 -> 0 < x -> R_dist x 1 < alp -> dist R_met (ps_atan x) (Alt_PI / 4) < eps)
eps:R
eps_pos:eps > 0
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x : R, -1 < x -> x < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x)) n - ps_atan x) < eps / 3

exists alp : R, alp > 0 /\ (forall x : R, x < 1 -> 0 < x -> R_dist x 1 < alp -> dist R_met (ps_atan x) (Alt_PI / 4) < eps)
eps:R
eps_pos:eps > 0
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x : R, -1 < x -> x < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x)) n - ps_atan x) < eps / 3

exists alp : R, alp > 0 /\ (forall x : R, x < 1 -> 0 < x -> R_dist x 1 < alp -> dist R_met (ps_atan x) (4 * (let (a, _) := exist_PI in a) / 4) < eps)
eps:R
eps_pos:eps > 0
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x : R, -1 < x -> x < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x)) n - ps_atan x) < eps / 3
v:R
Pv:Un_cv (fun N : nat => sum_f_R0 (tg_alt PI_tg) N) v

exists alp : R, alp > 0 /\ (forall x : R, x < 1 -> 0 < x -> R_dist x 1 < alp -> dist R_met (ps_atan x) v < eps)
eps:R
eps_pos:eps > 0
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x : R, -1 < x -> x < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x)) n - ps_atan x) < eps / 3
v:R
Pv:Un_cv (fun N : nat => sum_f_R0 (tg_alt PI_tg) N) v

Un_cv (sum_f_R0 (tg_alt (Ratan_seq 1))) v
eps:R
eps_pos:eps > 0
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x : R, -1 < x -> x < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x)) n - ps_atan x) < eps / 3
v:R
Pv:Un_cv (fun N : nat => sum_f_R0 (tg_alt PI_tg) N) v
Pv':Un_cv (sum_f_R0 (tg_alt (Ratan_seq 1))) v
exists alp : R, alp > 0 /\ (forall x : R, x < 1 -> 0 < x -> R_dist x 1 < alp -> dist R_met (ps_atan x) v < eps)
eps:R
eps_pos:eps > 0
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x : R, -1 < x -> x < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x)) n - ps_atan x) < eps / 3
v:R
Pv:Un_cv (fun N : nat => sum_f_R0 (tg_alt PI_tg) N) v

forall n : nat, sum_f_R0 (tg_alt PI_tg) n = sum_f_R0 (tg_alt (Ratan_seq 1)) n
eps:R
eps_pos:eps > 0
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x : R, -1 < x -> x < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x)) n - ps_atan x) < eps / 3
v:R
Pv:Un_cv (fun N : nat => sum_f_R0 (tg_alt PI_tg) N) v
Pv':Un_cv (sum_f_R0 (tg_alt (Ratan_seq 1))) v
exists alp : R, alp > 0 /\ (forall x : R, x < 1 -> 0 < x -> R_dist x 1 < alp -> dist R_met (ps_atan x) v < eps)
eps:R
eps_pos:eps > 0
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x : R, -1 < x -> x < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x)) n - ps_atan x) < eps / 3
v:R
Pv:Un_cv (fun N : nat => sum_f_R0 (tg_alt PI_tg) N) v
Pv':Un_cv (sum_f_R0 (tg_alt (Ratan_seq 1))) v

exists alp : R, alp > 0 /\ (forall x : R, x < 1 -> 0 < x -> R_dist x 1 < alp -> dist R_met (ps_atan x) v < eps)
eps:R
eps_pos:eps > 0
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x : R, -1 < x -> x < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x)) n - ps_atan x) < eps / 3
v:R
Pv:Un_cv (fun N : nat => sum_f_R0 (tg_alt PI_tg) N) v
Pv':Un_cv (sum_f_R0 (tg_alt (Ratan_seq 1))) v
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3

exists alp : R, alp > 0 /\ (forall x : R, x < 1 -> 0 < x -> R_dist x 1 < alp -> dist R_met (ps_atan x) v < eps)
eps:R
eps_pos:eps > 0
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x : R, -1 < x -> x < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x)) n - ps_atan x) < eps / 3
v:R
Pv:Un_cv (fun N0 : nat => sum_f_R0 (tg_alt PI_tg) N0) v
Pv':Un_cv (sum_f_R0 (tg_alt (Ratan_seq 1))) v
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat

exists alp : R, alp > 0 /\ (forall x : R, x < 1 -> 0 < x -> R_dist x 1 < alp -> dist R_met (ps_atan x) v < eps)
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x : R, -1 < x -> x < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x)) n - ps_atan x) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x : Base R_met, D_x no_cond 1 x /\ dist R_met x 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3

exists alp : R, alp > 0 /\ (forall x : R, x < 1 -> 0 < x -> R_dist x 1 < alp -> dist R_met (ps_atan x) v < eps)
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x : R, -1 < x -> x < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x)) n - ps_atan x) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x : Base R_met, D_x no_cond 1 x /\ dist R_met x 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3

forall x : R, x < 1 -> 0 < x -> R_dist x 1 < alpha -> dist R_met (ps_atan x) v < eps
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha

dist R_met (ps_atan x) v < eps
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha

Rabs (ps_atan x - v) < eps
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha

Rabs (ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) N + (sum_f_R0 (tg_alt (Ratan_seq x)) N - sum_f_R0 (tg_alt (Ratan_seq 1)) N) + (sum_f_R0 (tg_alt (Ratan_seq 1)) N - v)) < eps
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) N + (sum_f_R0 (tg_alt (Ratan_seq x)) N - sum_f_R0 (tg_alt (Ratan_seq 1)) N) + (sum_f_R0 (tg_alt (Ratan_seq 1)) N - v) = ps_atan x - v
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha

Rabs (ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) N + (sum_f_R0 (tg_alt (Ratan_seq x)) N - sum_f_R0 (tg_alt (Ratan_seq 1)) N) + (sum_f_R0 (tg_alt (Ratan_seq 1)) N - v)) <= Rabs (ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) N) + Rabs (sum_f_R0 (tg_alt (Ratan_seq x)) N - sum_f_R0 (tg_alt (Ratan_seq 1)) N + (sum_f_R0 (tg_alt (Ratan_seq 1)) N - v))
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
Rabs (ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) N) + Rabs (sum_f_R0 (tg_alt (Ratan_seq x)) N - sum_f_R0 (tg_alt (Ratan_seq 1)) N + (sum_f_R0 (tg_alt (Ratan_seq 1)) N - v)) < eps
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) N + (sum_f_R0 (tg_alt (Ratan_seq x)) N - sum_f_R0 (tg_alt (Ratan_seq 1)) N) + (sum_f_R0 (tg_alt (Ratan_seq 1)) N - v) = ps_atan x - v
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha

Rabs (ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) N) + Rabs (sum_f_R0 (tg_alt (Ratan_seq x)) N - sum_f_R0 (tg_alt (Ratan_seq 1)) N + (sum_f_R0 (tg_alt (Ratan_seq 1)) N - v)) < eps
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) N + (sum_f_R0 (tg_alt (Ratan_seq x)) N - sum_f_R0 (tg_alt (Ratan_seq 1)) N) + (sum_f_R0 (tg_alt (Ratan_seq 1)) N - v) = ps_atan x - v
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha

Rabs (ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) N) + Rabs (sum_f_R0 (tg_alt (Ratan_seq x)) N - sum_f_R0 (tg_alt (Ratan_seq 1)) N + (sum_f_R0 (tg_alt (Ratan_seq 1)) N - v)) < 2 / 3 * eps + eps / 3
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
2 / 3 * eps + eps / 3 = eps
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) N + (sum_f_R0 (tg_alt (Ratan_seq x)) N - sum_f_R0 (tg_alt (Ratan_seq 1)) N) + (sum_f_R0 (tg_alt (Ratan_seq 1)) N - v) = ps_atan x - v
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha

Rabs (sum_f_R0 (tg_alt (Ratan_seq x)) N - sum_f_R0 (tg_alt (Ratan_seq 1)) N + (sum_f_R0 (tg_alt (Ratan_seq 1)) N - v)) + Rabs (ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) N) < 2 / 3 * eps + eps / 3
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
2 / 3 * eps + eps / 3 = eps
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) N + (sum_f_R0 (tg_alt (Ratan_seq x)) N - sum_f_R0 (tg_alt (Ratan_seq 1)) N) + (sum_f_R0 (tg_alt (Ratan_seq 1)) N - v) = ps_atan x - v
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha

Rabs (sum_f_R0 (tg_alt (Ratan_seq x)) N - sum_f_R0 (tg_alt (Ratan_seq 1)) N + (sum_f_R0 (tg_alt (Ratan_seq 1)) N - v)) < 2 / 3 * eps
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
Rabs (ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) N) < eps / 3
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
2 / 3 * eps + eps / 3 = eps
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) N + (sum_f_R0 (tg_alt (Ratan_seq x)) N - sum_f_R0 (tg_alt (Ratan_seq 1)) N) + (sum_f_R0 (tg_alt (Ratan_seq 1)) N - v) = ps_atan x - v
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha

Rabs (sum_f_R0 (tg_alt (Ratan_seq x)) N - sum_f_R0 (tg_alt (Ratan_seq 1)) N + (sum_f_R0 (tg_alt (Ratan_seq 1)) N - v)) <= Rabs (sum_f_R0 (tg_alt (Ratan_seq x)) N - sum_f_R0 (tg_alt (Ratan_seq 1)) N) + Rabs (sum_f_R0 (tg_alt (Ratan_seq 1)) N - v)
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
Rabs (sum_f_R0 (tg_alt (Ratan_seq x)) N - sum_f_R0 (tg_alt (Ratan_seq 1)) N) + Rabs (sum_f_R0 (tg_alt (Ratan_seq 1)) N - v) < 2 / 3 * eps
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
Rabs (ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) N) < eps / 3
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
2 / 3 * eps + eps / 3 = eps
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) N + (sum_f_R0 (tg_alt (Ratan_seq x)) N - sum_f_R0 (tg_alt (Ratan_seq 1)) N) + (sum_f_R0 (tg_alt (Ratan_seq 1)) N - v) = ps_atan x - v
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha

Rabs (sum_f_R0 (tg_alt (Ratan_seq x)) N - sum_f_R0 (tg_alt (Ratan_seq 1)) N) + Rabs (sum_f_R0 (tg_alt (Ratan_seq 1)) N - v) < 2 / 3 * eps
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
Rabs (ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) N) < eps / 3
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
2 / 3 * eps + eps / 3 = eps
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) N + (sum_f_R0 (tg_alt (Ratan_seq x)) N - sum_f_R0 (tg_alt (Ratan_seq 1)) N) + (sum_f_R0 (tg_alt (Ratan_seq 1)) N - v) = ps_atan x - v
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha

Rabs (sum_f_R0 (tg_alt (Ratan_seq x)) N - sum_f_R0 (tg_alt (Ratan_seq 1)) N) + Rabs (sum_f_R0 (tg_alt (Ratan_seq 1)) N - v) < eps / 3 + eps / 3
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
eps / 3 + eps / 3 <= 2 / 3 * eps
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
Rabs (ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) N) < eps / 3
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
2 / 3 * eps + eps / 3 = eps
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) N + (sum_f_R0 (tg_alt (Ratan_seq x)) N - sum_f_R0 (tg_alt (Ratan_seq 1)) N) + (sum_f_R0 (tg_alt (Ratan_seq 1)) N - v) = ps_atan x - v
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha

Rabs (sum_f_R0 (tg_alt (Ratan_seq x)) N - sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
Rabs (sum_f_R0 (tg_alt (Ratan_seq 1)) N - v) < eps / 3
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
eps / 3 + eps / 3 <= 2 / 3 * eps
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
Rabs (ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) N) < eps / 3
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
2 / 3 * eps + eps / 3 = eps
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) N + (sum_f_R0 (tg_alt (Ratan_seq x)) N - sum_f_R0 (tg_alt (Ratan_seq 1)) N) + (sum_f_R0 (tg_alt (Ratan_seq 1)) N - v) = ps_atan x - v
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : R, D_x no_cond 1 x0 /\ Rabs (x0 - 1) < alpha -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) N - sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha

Rabs (sum_f_R0 (tg_alt (Ratan_seq x)) N - sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
Rabs (sum_f_R0 (tg_alt (Ratan_seq 1)) N - v) < eps / 3
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
eps / 3 + eps / 3 <= 2 / 3 * eps
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
Rabs (ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) N) < eps / 3
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
2 / 3 * eps + eps / 3 = eps
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) N + (sum_f_R0 (tg_alt (Ratan_seq x)) N - sum_f_R0 (tg_alt (Ratan_seq 1)) N) + (sum_f_R0 (tg_alt (Ratan_seq 1)) N - v) = ps_atan x - v
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : R, D_x no_cond 1 x0 /\ Rabs (x0 - 1) < alpha -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) N - sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha

D_x no_cond 1 x
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : R, D_x no_cond 1 x0 /\ Rabs (x0 - 1) < alpha -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) N - sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
Rabs (x - 1) < alpha
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
Rabs (sum_f_R0 (tg_alt (Ratan_seq 1)) N - v) < eps / 3
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
eps / 3 + eps / 3 <= 2 / 3 * eps
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
Rabs (ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) N) < eps / 3
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
2 / 3 * eps + eps / 3 = eps
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) N + (sum_f_R0 (tg_alt (Ratan_seq x)) N - sum_f_R0 (tg_alt (Ratan_seq 1)) N) + (sum_f_R0 (tg_alt (Ratan_seq 1)) N - v) = ps_atan x - v
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : R, D_x no_cond 1 x0 /\ Rabs (x0 - 1) < alpha -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) N - sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha

Rabs (x - 1) < alpha
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
Rabs (sum_f_R0 (tg_alt (Ratan_seq 1)) N - v) < eps / 3
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
eps / 3 + eps / 3 <= 2 / 3 * eps
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
Rabs (ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) N) < eps / 3
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
2 / 3 * eps + eps / 3 = eps
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) N + (sum_f_R0 (tg_alt (Ratan_seq x)) N - sum_f_R0 (tg_alt (Ratan_seq 1)) N) + (sum_f_R0 (tg_alt (Ratan_seq 1)) N - v) = ps_atan x - v
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha

Rabs (sum_f_R0 (tg_alt (Ratan_seq 1)) N - v) < eps / 3
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
eps / 3 + eps / 3 <= 2 / 3 * eps
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
Rabs (ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) N) < eps / 3
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
2 / 3 * eps + eps / 3 = eps
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) N + (sum_f_R0 (tg_alt (Ratan_seq x)) N - sum_f_R0 (tg_alt (Ratan_seq 1)) N) + (sum_f_R0 (tg_alt (Ratan_seq 1)) N - v) = ps_atan x - v
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha

eps / 3 + eps / 3 <= 2 / 3 * eps
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
Rabs (ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) N) < eps / 3
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
2 / 3 * eps + eps / 3 = eps
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) N + (sum_f_R0 (tg_alt (Ratan_seq x)) N - sum_f_R0 (tg_alt (Ratan_seq 1)) N) + (sum_f_R0 (tg_alt (Ratan_seq 1)) N - v) = ps_atan x - v
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha

Rabs (ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) N) < eps / 3
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
2 / 3 * eps + eps / 3 = eps
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) N + (sum_f_R0 (tg_alt (Ratan_seq x)) N - sum_f_R0 (tg_alt (Ratan_seq 1)) N) + (sum_f_R0 (tg_alt (Ratan_seq 1)) N - v) = ps_atan x - v
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha

(N >= N1)%nat
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
-1 < x
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
x < 1
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
2 / 3 * eps + eps / 3 = eps
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) N + (sum_f_R0 (tg_alt (Ratan_seq x)) N - sum_f_R0 (tg_alt (Ratan_seq 1)) N) + (sum_f_R0 (tg_alt (Ratan_seq 1)) N - v) = ps_atan x - v
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha

-1 < x
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
x < 1
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
2 / 3 * eps + eps / 3 = eps
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) N + (sum_f_R0 (tg_alt (Ratan_seq x)) N - sum_f_R0 (tg_alt (Ratan_seq 1)) N) + (sum_f_R0 (tg_alt (Ratan_seq 1)) N - v) = ps_atan x - v
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha

x < 1
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
2 / 3 * eps + eps / 3 = eps
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) N + (sum_f_R0 (tg_alt (Ratan_seq x)) N - sum_f_R0 (tg_alt (Ratan_seq 1)) N) + (sum_f_R0 (tg_alt (Ratan_seq 1)) N - v) = ps_atan x - v
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha

2 / 3 * eps + eps / 3 = eps
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha
ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) N + (sum_f_R0 (tg_alt (Ratan_seq x)) N - sum_f_R0 (tg_alt (Ratan_seq 1)) N) + (sum_f_R0 (tg_alt (Ratan_seq 1)) N - v) = ps_atan x - v
eps:R
eps_3_pos:eps / 3 > 0
N1:nat
HN1:forall n : nat, (n >= N1)%nat -> forall x0 : R, -1 < x0 -> x0 < 1 -> Rabs (sum_f_R0 (tg_alt (Ratan_seq x0)) n - ps_atan x0) < eps / 3
v:R
N2:nat
HN2:forall n : nat, (n >= N2)%nat -> R_dist (sum_f_R0 (tg_alt (Ratan_seq 1)) n) v < eps / 3
N:=(N1 + N2)%nat:nat
alpha:R
alpha_pos:alpha > 0
Halpha:forall x0 : Base R_met, D_x no_cond 1 x0 /\ dist R_met x0 1 < alpha -> dist R_met (sum_f_R0 (tg_alt (Ratan_seq x0)) N) (sum_f_R0 (tg_alt (Ratan_seq 1)) N) < eps / 3
x:R
x_ub:x < 1
x_lb:0 < x
x_bounds:R_dist x 1 < alpha

ps_atan x - sum_f_R0 (tg_alt (Ratan_seq x)) N + (sum_f_R0 (tg_alt (Ratan_seq x)) N - sum_f_R0 (tg_alt (Ratan_seq 1)) N) + (sum_f_R0 (tg_alt (Ratan_seq 1)) N - v) = ps_atan x - v
ring. Qed.

forall x : R, -1 < x < 1 -> forall (Pratan : derivable_pt ps_atan x) (Prmymeta : derivable_pt atan x), derive_pt ps_atan x Pratan = derive_pt atan x Prmymeta

forall x : R, -1 < x < 1 -> forall (Pratan : derivable_pt ps_atan x) (Prmymeta : derivable_pt atan x), derive_pt ps_atan x Pratan = derive_pt atan x Prmymeta
freq:0 < tan 1

forall x : R, -1 < x < 1 -> forall (Pratan : derivable_pt ps_atan x) (Prmymeta : derivable_pt atan x), derive_pt ps_atan x Pratan = derive_pt atan x Prmymeta
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x

derive_pt ps_atan x Pratan = derive_pt atan x Prmymeta
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x

derive_pt ps_atan x (derivable_pt_ps_atan x x_encad) = derive_pt atan x Prmymeta
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
-1 < tan 1
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
-1 < x < tan 1
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
forall h : R, -1 < h < tan 1 -> ps_atan h = ps_atan h
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x

derive_pt ps_atan x (derivable_pt_ps_atan x x_encad) = derive_pt atan x (derivable_pt_atan x)
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
-1 < 1
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
-1 < x < 1
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
forall h : R, -1 < h < 1 -> atan h = atan h
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
-1 < tan 1
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
-1 < x < tan 1
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
forall h : R, -1 < h < tan 1 -> ps_atan h = ps_atan h
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
Temp:derivable_pt_lim ps_atan x ((fun y : R => / (1 + y ^ 2)) x)

derive_pt ps_atan x (derivable_pt_ps_atan x x_encad) = derive_pt atan x (derivable_pt_atan x)
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
-1 < 1
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
-1 < x < 1
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
forall h : R, -1 < h < 1 -> atan h = atan h
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
-1 < tan 1
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
-1 < x < tan 1
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
forall h : R, -1 < h < tan 1 -> ps_atan h = ps_atan h
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
Temp:derivable_pt_lim ps_atan x ((fun y : R => / (1 + y ^ 2)) x)

derive_pt ps_atan x (derivable_pt_ps_atan x x_encad) = / (1 + x ^ 2)
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
Temp:derivable_pt_lim ps_atan x ((fun y : R => / (1 + y ^ 2)) x)
Hrew1:derive_pt ps_atan x (derivable_pt_ps_atan x x_encad) = / (1 + x ^ 2)
derive_pt ps_atan x (derivable_pt_ps_atan x x_encad) = derive_pt atan x (derivable_pt_atan x)
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
-1 < 1
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
-1 < x < 1
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
forall h : R, -1 < h < 1 -> atan h = atan h
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
-1 < tan 1
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
-1 < x < tan 1
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
forall h : R, -1 < h < tan 1 -> ps_atan h = ps_atan h
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
Temp:derivable_pt_lim ps_atan x ((fun y : R => / (1 + y ^ 2)) x)
Hrew1:derive_pt ps_atan x (derivable_pt_ps_atan x x_encad) = / (1 + x ^ 2)

derive_pt ps_atan x (derivable_pt_ps_atan x x_encad) = derive_pt atan x (derivable_pt_atan x)
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
-1 < 1
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
-1 < x < 1
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
forall h : R, -1 < h < 1 -> atan h = atan h
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
-1 < tan 1
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
-1 < x < tan 1
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
forall h : R, -1 < h < tan 1 -> ps_atan h = ps_atan h
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
Temp:derivable_pt_lim ps_atan x ((fun y : R => / (1 + y ^ 2)) x)
Hrew1:derive_pt ps_atan x (derivable_pt_ps_atan x x_encad) = / (1 + x ^ 2)

derive_pt ps_atan x (derivable_pt_ps_atan x x_encad) = 1 / (1 + x²)
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
-1 < 1
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
-1 < x < 1
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
forall h : R, -1 < h < 1 -> atan h = atan h
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
-1 < tan 1
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
-1 < x < tan 1
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
forall h : R, -1 < h < tan 1 -> ps_atan h = ps_atan h
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
Temp:derivable_pt_lim ps_atan x ((fun y : R => / (1 + y ^ 2)) x)
Hrew1:derive_pt ps_atan x (derivable_pt_ps_atan x x_encad) = / (1 + x ^ 2)

/ (1 + x ^ 2) = 1 / (1 + x²)
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
-1 < 1
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
-1 < x < 1
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
forall h : R, -1 < h < 1 -> atan h = atan h
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
-1 < tan 1
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
-1 < x < tan 1
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
forall h : R, -1 < h < tan 1 -> ps_atan h = ps_atan h
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
Temp:derivable_pt_lim ps_atan x ((fun y : R => / (1 + y ^ 2)) x)
Hrew1:derive_pt ps_atan x (derivable_pt_ps_atan x x_encad) = / (1 + x ^ 2)

/ (1 + x ^ 2) = 1 / (1 + x ^ 2)
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
-1 < 1
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
-1 < x < 1
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
forall h : R, -1 < h < 1 -> atan h = atan h
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
-1 < tan 1
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
-1 < x < tan 1
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
forall h : R, -1 < h < tan 1 -> ps_atan h = ps_atan h
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x

-1 < 1
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
-1 < x < 1
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
forall h : R, -1 < h < 1 -> atan h = atan h
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
-1 < tan 1
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
-1 < x < tan 1
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
forall h : R, -1 < h < tan 1 -> ps_atan h = ps_atan h
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x

-1 < x < 1
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
forall h : R, -1 < h < 1 -> atan h = atan h
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
-1 < tan 1
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
-1 < x < tan 1
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
forall h : R, -1 < h < tan 1 -> ps_atan h = ps_atan h
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x

forall h : R, -1 < h < 1 -> atan h = atan h
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
-1 < tan 1
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
-1 < x < tan 1
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
forall h : R, -1 < h < tan 1 -> ps_atan h = ps_atan h
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x

-1 < tan 1
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
-1 < x < tan 1
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
forall h : R, -1 < h < tan 1 -> ps_atan h = ps_atan h
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x

-1 < x < tan 1
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x
forall h : R, -1 < h < tan 1 -> ps_atan h = ps_atan h
freq:0 < tan 1
x:R
x_encad:-1 < x < 1
Pratan:derivable_pt ps_atan x
Prmymeta:derivable_pt atan x

forall h : R, -1 < h < tan 1 -> ps_atan h = ps_atan h
intros; reflexivity. Qed.

forall x : R, 0 < x < 1 -> atan x = ps_atan x

forall x : R, 0 < x < 1 -> atan x = ps_atan x
x:R
x_encad:0 < x < 1

atan x = ps_atan x
x:R
x_encad:0 < x < 1

forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
atan x = ps_atan x
x:R
x_encad:0 < x < 1
c:R
c_encad:0 < c < x

derivable_pt (atan - ps_atan) c
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
atan x = ps_atan x
x:R
x_encad:0 < x < 1
c:R
c_encad:0 < c < x

derivable_pt atan c
x:R
x_encad:0 < x < 1
c:R
c_encad:0 < c < x
derivable_pt ps_atan c
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
atan x = ps_atan x
x:R
x_encad:0 < x < 1
c:R
c_encad:0 < c < x

derivable_pt ps_atan c
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
atan x = ps_atan x
x:R
x_encad:0 < x < 1
c:R
c_encad:0 < c < x

-1 < c < 1
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
atan x = ps_atan x
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c

atan x = ps_atan x
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c

forall c : R, 0 < c < x -> derivable_pt id c
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
atan x = ps_atan x
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c

atan x = ps_atan x
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c

forall c : R, 0 <= c <= x -> continuity_pt (atan - ps_atan) c
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
delta_cont:forall c : R, 0 <= c <= x -> continuity_pt (atan - ps_atan) c
atan x = ps_atan x
x:R
x_encad:0 < x < 1
pr1:forall c0 : R, 0 < c0 < x -> derivable_pt (atan - ps_atan) c0
pr2:forall c0 : R, 0 < c0 < x -> derivable_pt id c0
c:R
c_encad1:0 < c
c_encad2:c < x

continuity_pt atan c
x:R
x_encad:0 < x < 1
pr1:forall c0 : R, 0 < c0 < x -> derivable_pt (atan - ps_atan) c0
pr2:forall c0 : R, 0 < c0 < x -> derivable_pt id c0
c:R
c_encad1:0 < c
c_encad2:c < x
continuity_pt ps_atan c
x:R
x_encad:0 < x < 1
pr1:forall c0 : R, 0 < c0 < x -> derivable_pt (atan - ps_atan) c0
pr2:forall c0 : R, 0 < c0 < x -> derivable_pt id c0
c:R
c_encad1:0 < c
c_encad2:c = x
continuity_pt atan c
x:R
x_encad:0 < x < 1
pr1:forall c0 : R, 0 < c0 < x -> derivable_pt (atan - ps_atan) c0
pr2:forall c0 : R, 0 < c0 < x -> derivable_pt id c0
c:R
c_encad1:0 < c
c_encad2:c = x
continuity_pt ps_atan c
x:R
x_encad:0 < x < 1
pr1:forall c0 : R, 0 < c0 < x -> derivable_pt (atan - ps_atan) c0
pr2:forall c0 : R, 0 < c0 < x -> derivable_pt id c0
c:R
c_encad1:0 = c
c_encad2:c < x
continuity_pt atan c
x:R
x_encad:0 < x < 1
pr1:forall c0 : R, 0 < c0 < x -> derivable_pt (atan - ps_atan) c0
pr2:forall c0 : R, 0 < c0 < x -> derivable_pt id c0
c:R
c_encad1:0 = c
c_encad2:c < x
continuity_pt ps_atan c
x:R
x_encad:0 < x < 1
pr1:forall c0 : R, 0 < c0 < x -> derivable_pt (atan - ps_atan) c0
pr2:forall c0 : R, 0 < c0 < x -> derivable_pt id c0
c:R
c_encad1:0 = c
c_encad2:c = x
continuity_pt atan c
x:R
x_encad:0 < x < 1
pr1:forall c0 : R, 0 < c0 < x -> derivable_pt (atan - ps_atan) c0
pr2:forall c0 : R, 0 < c0 < x -> derivable_pt id c0
c:R
c_encad1:0 = c
c_encad2:c = x
continuity_pt ps_atan c
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
delta_cont:forall c : R, 0 <= c <= x -> continuity_pt (atan - ps_atan) c
atan x = ps_atan x
x:R
x_encad:0 < x < 1
pr1:forall c0 : R, 0 < c0 < x -> derivable_pt (atan - ps_atan) c0
pr2:forall c0 : R, 0 < c0 < x -> derivable_pt id c0
c:R
c_encad1:0 < c
c_encad2:c < x

continuity_pt ps_atan c
x:R
x_encad:0 < x < 1
pr1:forall c0 : R, 0 < c0 < x -> derivable_pt (atan - ps_atan) c0
pr2:forall c0 : R, 0 < c0 < x -> derivable_pt id c0
c:R
c_encad1:0 < c
c_encad2:c = x
continuity_pt atan c
x:R
x_encad:0 < x < 1
pr1:forall c0 : R, 0 < c0 < x -> derivable_pt (atan - ps_atan) c0
pr2:forall c0 : R, 0 < c0 < x -> derivable_pt id c0
c:R
c_encad1:0 < c
c_encad2:c = x
continuity_pt ps_atan c
x:R
x_encad:0 < x < 1
pr1:forall c0 : R, 0 < c0 < x -> derivable_pt (atan - ps_atan) c0
pr2:forall c0 : R, 0 < c0 < x -> derivable_pt id c0
c:R
c_encad1:0 = c
c_encad2:c < x
continuity_pt atan c
x:R
x_encad:0 < x < 1
pr1:forall c0 : R, 0 < c0 < x -> derivable_pt (atan - ps_atan) c0
pr2:forall c0 : R, 0 < c0 < x -> derivable_pt id c0
c:R
c_encad1:0 = c
c_encad2:c < x
continuity_pt ps_atan c
x:R
x_encad:0 < x < 1
pr1:forall c0 : R, 0 < c0 < x -> derivable_pt (atan - ps_atan) c0
pr2:forall c0 : R, 0 < c0 < x -> derivable_pt id c0
c:R
c_encad1:0 = c
c_encad2:c = x
continuity_pt atan c
x:R
x_encad:0 < x < 1
pr1:forall c0 : R, 0 < c0 < x -> derivable_pt (atan - ps_atan) c0
pr2:forall c0 : R, 0 < c0 < x -> derivable_pt id c0
c:R
c_encad1:0 = c
c_encad2:c = x
continuity_pt ps_atan c
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
delta_cont:forall c : R, 0 <= c <= x -> continuity_pt (atan - ps_atan) c
atan x = ps_atan x
x:R
x_encad:0 < x < 1
pr1:forall c0 : R, 0 < c0 < x -> derivable_pt (atan - ps_atan) c0
pr2:forall c0 : R, 0 < c0 < x -> derivable_pt id c0
c:R
c_encad1:0 < c
c_encad2:c < x

-1 < c < 1
x:R
x_encad:0 < x < 1
pr1:forall c0 : R, 0 < c0 < x -> derivable_pt (atan - ps_atan) c0
pr2:forall c0 : R, 0 < c0 < x -> derivable_pt id c0
c:R
c_encad1:0 < c
c_encad2:c = x
continuity_pt atan c
x:R
x_encad:0 < x < 1
pr1:forall c0 : R, 0 < c0 < x -> derivable_pt (atan - ps_atan) c0
pr2:forall c0 : R, 0 < c0 < x -> derivable_pt id c0
c:R
c_encad1:0 < c
c_encad2:c = x
continuity_pt ps_atan c
x:R
x_encad:0 < x < 1
pr1:forall c0 : R, 0 < c0 < x -> derivable_pt (atan - ps_atan) c0
pr2:forall c0 : R, 0 < c0 < x -> derivable_pt id c0
c:R
c_encad1:0 = c
c_encad2:c < x
continuity_pt atan c
x:R
x_encad:0 < x < 1
pr1:forall c0 : R, 0 < c0 < x -> derivable_pt (atan - ps_atan) c0
pr2:forall c0 : R, 0 < c0 < x -> derivable_pt id c0
c:R
c_encad1:0 = c
c_encad2:c < x
continuity_pt ps_atan c
x:R
x_encad:0 < x < 1
pr1:forall c0 : R, 0 < c0 < x -> derivable_pt (atan - ps_atan) c0
pr2:forall c0 : R, 0 < c0 < x -> derivable_pt id c0
c:R
c_encad1:0 = c
c_encad2:c = x
continuity_pt atan c
x:R
x_encad:0 < x < 1
pr1:forall c0 : R, 0 < c0 < x -> derivable_pt (atan - ps_atan) c0
pr2:forall c0 : R, 0 < c0 < x -> derivable_pt id c0
c:R
c_encad1:0 = c
c_encad2:c = x
continuity_pt ps_atan c
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
delta_cont:forall c : R, 0 <= c <= x -> continuity_pt (atan - ps_atan) c
atan x = ps_atan x
x:R
x_encad:0 < x < 1
pr1:forall c0 : R, 0 < c0 < x -> derivable_pt (atan - ps_atan) c0
pr2:forall c0 : R, 0 < c0 < x -> derivable_pt id c0
c:R
c_encad1:0 < c
c_encad2:c = x

continuity_pt atan c
x:R
x_encad:0 < x < 1
pr1:forall c0 : R, 0 < c0 < x -> derivable_pt (atan - ps_atan) c0
pr2:forall c0 : R, 0 < c0 < x -> derivable_pt id c0
c:R
c_encad1:0 < c
c_encad2:c = x
continuity_pt ps_atan c
x:R
x_encad:0 < x < 1
pr1:forall c0 : R, 0 < c0 < x -> derivable_pt (atan - ps_atan) c0
pr2:forall c0 : R, 0 < c0 < x -> derivable_pt id c0
c:R
c_encad1:0 = c
c_encad2:c < x
continuity_pt atan c
x:R
x_encad:0 < x < 1
pr1:forall c0 : R, 0 < c0 < x -> derivable_pt (atan - ps_atan) c0
pr2:forall c0 : R, 0 < c0 < x -> derivable_pt id c0
c:R
c_encad1:0 = c
c_encad2:c < x
continuity_pt ps_atan c
x:R
x_encad:0 < x < 1
pr1:forall c0 : R, 0 < c0 < x -> derivable_pt (atan - ps_atan) c0
pr2:forall c0 : R, 0 < c0 < x -> derivable_pt id c0
c:R
c_encad1:0 = c
c_encad2:c = x
continuity_pt atan c
x:R
x_encad:0 < x < 1
pr1:forall c0 : R, 0 < c0 < x -> derivable_pt (atan - ps_atan) c0
pr2:forall c0 : R, 0 < c0 < x -> derivable_pt id c0
c:R
c_encad1:0 = c
c_encad2:c = x
continuity_pt ps_atan c
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
delta_cont:forall c : R, 0 <= c <= x -> continuity_pt (atan - ps_atan) c
atan x = ps_atan x
x:R
x_encad:0 < x < 1
pr1:forall c0 : R, 0 < c0 < x -> derivable_pt (atan - ps_atan) c0
pr2:forall c0 : R, 0 < c0 < x -> derivable_pt id c0
c:R
c_encad1:0 < c
c_encad2:c = x

continuity_pt ps_atan c
x:R
x_encad:0 < x < 1
pr1:forall c0 : R, 0 < c0 < x -> derivable_pt (atan - ps_atan) c0
pr2:forall c0 : R, 0 < c0 < x -> derivable_pt id c0
c:R
c_encad1:0 = c
c_encad2:c < x
continuity_pt atan c
x:R
x_encad:0 < x < 1
pr1:forall c0 : R, 0 < c0 < x -> derivable_pt (atan - ps_atan) c0
pr2:forall c0 : R, 0 < c0 < x -> derivable_pt id c0
c:R
c_encad1:0 = c
c_encad2:c < x
continuity_pt ps_atan c
x:R
x_encad:0 < x < 1
pr1:forall c0 : R, 0 < c0 < x -> derivable_pt (atan - ps_atan) c0
pr2:forall c0 : R, 0 < c0 < x -> derivable_pt id c0
c:R
c_encad1:0 = c
c_encad2:c = x
continuity_pt atan c
x:R
x_encad:0 < x < 1
pr1:forall c0 : R, 0 < c0 < x -> derivable_pt (atan - ps_atan) c0
pr2:forall c0 : R, 0 < c0 < x -> derivable_pt id c0
c:R
c_encad1:0 = c
c_encad2:c = x
continuity_pt ps_atan c
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
delta_cont:forall c : R, 0 <= c <= x -> continuity_pt (atan - ps_atan) c
atan x = ps_atan x
x:R
x_encad:0 < x < 1
pr1:forall c0 : R, 0 < c0 < x -> derivable_pt (atan - ps_atan) c0
pr2:forall c0 : R, 0 < c0 < x -> derivable_pt id c0
c:R
c_encad1:0 < c
c_encad2:c = x

-1 < c < 1
x:R
x_encad:0 < x < 1
pr1:forall c0 : R, 0 < c0 < x -> derivable_pt (atan - ps_atan) c0
pr2:forall c0 : R, 0 < c0 < x -> derivable_pt id c0
c:R
c_encad1:0 = c
c_encad2:c < x
continuity_pt atan c
x:R
x_encad:0 < x < 1
pr1:forall c0 : R, 0 < c0 < x -> derivable_pt (atan - ps_atan) c0
pr2:forall c0 : R, 0 < c0 < x -> derivable_pt id c0
c:R
c_encad1:0 = c
c_encad2:c < x
continuity_pt ps_atan c
x:R
x_encad:0 < x < 1
pr1:forall c0 : R, 0 < c0 < x -> derivable_pt (atan - ps_atan) c0
pr2:forall c0 : R, 0 < c0 < x -> derivable_pt id c0
c:R
c_encad1:0 = c
c_encad2:c = x
continuity_pt atan c
x:R
x_encad:0 < x < 1
pr1:forall c0 : R, 0 < c0 < x -> derivable_pt (atan - ps_atan) c0
pr2:forall c0 : R, 0 < c0 < x -> derivable_pt id c0
c:R
c_encad1:0 = c
c_encad2:c = x
continuity_pt ps_atan c
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
delta_cont:forall c : R, 0 <= c <= x -> continuity_pt (atan - ps_atan) c
atan x = ps_atan x
x:R
x_encad:0 < x < 1
pr1:forall c0 : R, 0 < c0 < x -> derivable_pt (atan - ps_atan) c0
pr2:forall c0 : R, 0 < c0 < x -> derivable_pt id c0
c:R
c_encad1:0 = c
c_encad2:c < x

continuity_pt atan c
x:R
x_encad:0 < x < 1
pr1:forall c0 : R, 0 < c0 < x -> derivable_pt (atan - ps_atan) c0
pr2:forall c0 : R, 0 < c0 < x -> derivable_pt id c0
c:R
c_encad1:0 = c
c_encad2:c < x
continuity_pt ps_atan c
x:R
x_encad:0 < x < 1
pr1:forall c0 : R, 0 < c0 < x -> derivable_pt (atan - ps_atan) c0
pr2:forall c0 : R, 0 < c0 < x -> derivable_pt id c0
c:R
c_encad1:0 = c
c_encad2:c = x
continuity_pt atan c
x:R
x_encad:0 < x < 1
pr1:forall c0 : R, 0 < c0 < x -> derivable_pt (atan - ps_atan) c0
pr2:forall c0 : R, 0 < c0 < x -> derivable_pt id c0
c:R
c_encad1:0 = c
c_encad2:c = x
continuity_pt ps_atan c
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
delta_cont:forall c : R, 0 <= c <= x -> continuity_pt (atan - ps_atan) c
atan x = ps_atan x
x:R
x_encad:0 < x < 1
pr1:forall c0 : R, 0 < c0 < x -> derivable_pt (atan - ps_atan) c0
pr2:forall c0 : R, 0 < c0 < x -> derivable_pt id c0
c:R
c_encad1:0 = c
c_encad2:c < x

continuity_pt ps_atan c
x:R
x_encad:0 < x < 1
pr1:forall c0 : R, 0 < c0 < x -> derivable_pt (atan - ps_atan) c0
pr2:forall c0 : R, 0 < c0 < x -> derivable_pt id c0
c:R
c_encad1:0 = c
c_encad2:c = x
continuity_pt atan c
x:R
x_encad:0 < x < 1
pr1:forall c0 : R, 0 < c0 < x -> derivable_pt (atan - ps_atan) c0
pr2:forall c0 : R, 0 < c0 < x -> derivable_pt id c0
c:R
c_encad1:0 = c
c_encad2:c = x
continuity_pt ps_atan c
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
delta_cont:forall c : R, 0 <= c <= x -> continuity_pt (atan - ps_atan) c
atan x = ps_atan x
x:R
x_encad:0 < x < 1
pr1:forall c0 : R, 0 < c0 < x -> derivable_pt (atan - ps_atan) c0
pr2:forall c0 : R, 0 < c0 < x -> derivable_pt id c0
c:R
c_encad1:0 = c
c_encad2:c < x

-1 < c < 1
x:R
x_encad:0 < x < 1
pr1:forall c0 : R, 0 < c0 < x -> derivable_pt (atan - ps_atan) c0
pr2:forall c0 : R, 0 < c0 < x -> derivable_pt id c0
c:R
c_encad1:0 = c
c_encad2:c = x
continuity_pt atan c
x:R
x_encad:0 < x < 1
pr1:forall c0 : R, 0 < c0 < x -> derivable_pt (atan - ps_atan) c0
pr2:forall c0 : R, 0 < c0 < x -> derivable_pt id c0
c:R
c_encad1:0 = c
c_encad2:c = x
continuity_pt ps_atan c
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
delta_cont:forall c : R, 0 <= c <= x -> continuity_pt (atan - ps_atan) c
atan x = ps_atan x
x:R
x_encad:0 < x < 1
pr1:forall c0 : R, 0 < c0 < x -> derivable_pt (atan - ps_atan) c0
pr2:forall c0 : R, 0 < c0 < x -> derivable_pt id c0
c:R
c_encad1:0 = c
c_encad2:c = x

continuity_pt atan c
x:R
x_encad:0 < x < 1
pr1:forall c0 : R, 0 < c0 < x -> derivable_pt (atan - ps_atan) c0
pr2:forall c0 : R, 0 < c0 < x -> derivable_pt id c0
c:R
c_encad1:0 = c
c_encad2:c = x
continuity_pt ps_atan c
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
delta_cont:forall c : R, 0 <= c <= x -> continuity_pt (atan - ps_atan) c
atan x = ps_atan x
x:R
x_encad:0 < x < 1
pr1:forall c0 : R, 0 < c0 < x -> derivable_pt (atan - ps_atan) c0
pr2:forall c0 : R, 0 < c0 < x -> derivable_pt id c0
c:R
c_encad1:0 = c
c_encad2:c = x

continuity_pt ps_atan c
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
delta_cont:forall c : R, 0 <= c <= x -> continuity_pt (atan - ps_atan) c
atan x = ps_atan x
x:R
x_encad:0 < x < 1
pr1:forall c0 : R, 0 < c0 < x -> derivable_pt (atan - ps_atan) c0
pr2:forall c0 : R, 0 < c0 < x -> derivable_pt id c0
c:R
c_encad1:0 = c
c_encad2:c = x

-1 < c < 1
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
delta_cont:forall c : R, 0 <= c <= x -> continuity_pt (atan - ps_atan) c
atan x = ps_atan x
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
delta_cont:forall c : R, 0 <= c <= x -> continuity_pt (atan - ps_atan) c

atan x = ps_atan x
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
delta_cont:forall c : R, 0 <= c <= x -> continuity_pt (atan - ps_atan) c

forall c : R, 0 <= c <= x -> continuity_pt id c
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
delta_cont:forall c : R, 0 <= c <= x -> continuity_pt (atan - ps_atan) c
id_cont:forall c : R, 0 <= c <= x -> continuity_pt id c
atan x = ps_atan x
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
delta_cont:forall c : R, 0 <= c <= x -> continuity_pt (atan - ps_atan) c
id_cont:forall c : R, 0 <= c <= x -> continuity_pt id c

atan x = ps_atan x
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
delta_cont:forall c : R, 0 <= c <= x -> continuity_pt (atan - ps_atan) c
id_cont:forall c : R, 0 <= c <= x -> continuity_pt id c
x_lb:0 < x

atan x = ps_atan x
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
delta_cont:forall c : R, 0 <= c <= x -> continuity_pt (atan - ps_atan) c
id_cont:forall c : R, 0 <= c <= x -> continuity_pt id c
x_lb:0 < x
d:R
Temp:exists P : 0 < d < x, (id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d P) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d P)
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)

atan x = ps_atan x
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)

atan x = ps_atan x
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)

forall pr : derivable_pt (atan - ps_atan) d, derive_pt (atan - ps_atan) d pr = 0
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
Temp:forall pr : derivable_pt (atan - ps_atan) d, derive_pt (atan - ps_atan) d pr = 0
atan x = ps_atan x
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
pr:derivable_pt (atan - ps_atan) d

derive_pt (atan - ps_atan) d pr = 0
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
Temp:forall pr : derivable_pt (atan - ps_atan) d, derive_pt (atan - ps_atan) d pr = 0
atan x = ps_atan x
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
pr:derivable_pt (atan - ps_atan) d

-1 < d < 1
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
pr:derivable_pt (atan - ps_atan) d
d_encad3:-1 < d < 1
derive_pt (atan - ps_atan) d pr = 0
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
Temp:forall pr : derivable_pt (atan - ps_atan) d, derive_pt (atan - ps_atan) d pr = 0
atan x = ps_atan x
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
pr:derivable_pt (atan - ps_atan) d
d_encad3:-1 < d < 1

derive_pt (atan - ps_atan) d pr = 0
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
Temp:forall pr : derivable_pt (atan - ps_atan) d, derive_pt (atan - ps_atan) d pr = 0
atan x = ps_atan x
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
pr:derivable_pt (atan - ps_atan) d
d_encad3:-1 < d < 1
pr3:=derivable_pt_minus atan ps_atan d (derivable_pt_atan d) (derivable_pt_ps_atan d d_encad3):derivable_pt (atan - ps_atan) d

derive_pt (atan - ps_atan) d pr = 0
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
Temp:forall pr : derivable_pt (atan - ps_atan) d, derive_pt (atan - ps_atan) d pr = 0
atan x = ps_atan x
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
pr:derivable_pt (atan - ps_atan) d
d_encad3:-1 < d < 1
pr3:=derivable_pt_minus atan ps_atan d (derivable_pt_atan d) (derivable_pt_ps_atan d d_encad3):derivable_pt (atan - ps_atan) d

derive_pt (atan - ps_atan) d pr3 = 0
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
pr:derivable_pt (atan - ps_atan) d
d_encad3:-1 < d < 1
pr3:=derivable_pt_minus atan ps_atan d (derivable_pt_atan d) (derivable_pt_ps_atan d d_encad3):derivable_pt (atan - ps_atan) d
0 < x
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
pr:derivable_pt (atan - ps_atan) d
d_encad3:-1 < d < 1
pr3:=derivable_pt_minus atan ps_atan d (derivable_pt_atan d) (derivable_pt_ps_atan d d_encad3):derivable_pt (atan - ps_atan) d
0 < d < x
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
pr:derivable_pt (atan - ps_atan) d
d_encad3:-1 < d < 1
pr3:=derivable_pt_minus atan ps_atan d (derivable_pt_atan d) (derivable_pt_ps_atan d d_encad3):derivable_pt (atan - ps_atan) d
forall h : R, 0 < h < x -> (atan - ps_atan)%F h = (atan - ps_atan)%F h
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
Temp:forall pr : derivable_pt (atan - ps_atan) d, derive_pt (atan - ps_atan) d pr = 0
atan x = ps_atan x
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
pr:derivable_pt (atan - ps_atan) d
d_encad3:-1 < d < 1
pr3:=derivable_pt_minus atan ps_atan d (derivable_pt_atan d) (derivable_pt_ps_atan d d_encad3):derivable_pt (atan - ps_atan) d

derive_pt (atan - ps_atan) d (derivable_pt_minus atan ps_atan d (derivable_pt_atan d) (derivable_pt_ps_atan d d_encad3)) = 0
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
pr:derivable_pt (atan - ps_atan) d
d_encad3:-1 < d < 1
pr3:=derivable_pt_minus atan ps_atan d (derivable_pt_atan d) (derivable_pt_ps_atan d d_encad3):derivable_pt (atan - ps_atan) d
0 < x
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
pr:derivable_pt (atan - ps_atan) d
d_encad3:-1 < d < 1
pr3:=derivable_pt_minus atan ps_atan d (derivable_pt_atan d) (derivable_pt_ps_atan d d_encad3):derivable_pt (atan - ps_atan) d
0 < d < x
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
pr:derivable_pt (atan - ps_atan) d
d_encad3:-1 < d < 1
pr3:=derivable_pt_minus atan ps_atan d (derivable_pt_atan d) (derivable_pt_ps_atan d d_encad3):derivable_pt (atan - ps_atan) d
forall h : R, 0 < h < x -> (atan - ps_atan)%F h = (atan - ps_atan)%F h
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
Temp:forall pr : derivable_pt (atan - ps_atan) d, derive_pt (atan - ps_atan) d pr = 0
atan x = ps_atan x
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
pr:derivable_pt (atan - ps_atan) d
d_encad3:-1 < d < 1
pr3:=derivable_pt_minus atan ps_atan d (derivable_pt_atan d) (derivable_pt_ps_atan d d_encad3):derivable_pt (atan - ps_atan) d

derive_pt atan d (derivable_pt_atan d) - derive_pt ps_atan d (derivable_pt_ps_atan d d_encad3) = 0
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
pr:derivable_pt (atan - ps_atan) d
d_encad3:-1 < d < 1
pr3:=derivable_pt_minus atan ps_atan d (derivable_pt_atan d) (derivable_pt_ps_atan d d_encad3):derivable_pt (atan - ps_atan) d
0 < x
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
pr:derivable_pt (atan - ps_atan) d
d_encad3:-1 < d < 1
pr3:=derivable_pt_minus atan ps_atan d (derivable_pt_atan d) (derivable_pt_ps_atan d d_encad3):derivable_pt (atan - ps_atan) d
0 < d < x
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
pr:derivable_pt (atan - ps_atan) d
d_encad3:-1 < d < 1
pr3:=derivable_pt_minus atan ps_atan d (derivable_pt_atan d) (derivable_pt_ps_atan d d_encad3):derivable_pt (atan - ps_atan) d
forall h : R, 0 < h < x -> (atan - ps_atan)%F h = (atan - ps_atan)%F h
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
Temp:forall pr : derivable_pt (atan - ps_atan) d, derive_pt (atan - ps_atan) d pr = 0
atan x = ps_atan x
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
pr:derivable_pt (atan - ps_atan) d
d_encad3:-1 < d < 1
pr3:=derivable_pt_minus atan ps_atan d (derivable_pt_atan d) (derivable_pt_ps_atan d d_encad3):derivable_pt (atan - ps_atan) d

derive_pt atan d (derivable_pt_atan d) - derive_pt atan d (derivable_pt_atan d) = 0
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
pr:derivable_pt (atan - ps_atan) d
d_encad3:-1 < d < 1
pr3:=derivable_pt_minus atan ps_atan d (derivable_pt_atan d) (derivable_pt_ps_atan d d_encad3):derivable_pt (atan - ps_atan) d
-1 < d < 1
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
pr:derivable_pt (atan - ps_atan) d
d_encad3:-1 < d < 1
pr3:=derivable_pt_minus atan ps_atan d (derivable_pt_atan d) (derivable_pt_ps_atan d d_encad3):derivable_pt (atan - ps_atan) d
0 < x
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
pr:derivable_pt (atan - ps_atan) d
d_encad3:-1 < d < 1
pr3:=derivable_pt_minus atan ps_atan d (derivable_pt_atan d) (derivable_pt_ps_atan d d_encad3):derivable_pt (atan - ps_atan) d
0 < d < x
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
pr:derivable_pt (atan - ps_atan) d
d_encad3:-1 < d < 1
pr3:=derivable_pt_minus atan ps_atan d (derivable_pt_atan d) (derivable_pt_ps_atan d d_encad3):derivable_pt (atan - ps_atan) d
forall h : R, 0 < h < x -> (atan - ps_atan)%F h = (atan - ps_atan)%F h
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
Temp:forall pr : derivable_pt (atan - ps_atan) d, derive_pt (atan - ps_atan) d pr = 0
atan x = ps_atan x
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
pr:derivable_pt (atan - ps_atan) d
d_encad3:-1 < d < 1
pr3:=derivable_pt_minus atan ps_atan d (derivable_pt_atan d) (derivable_pt_ps_atan d d_encad3):derivable_pt (atan - ps_atan) d

-1 < d < 1
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
pr:derivable_pt (atan - ps_atan) d
d_encad3:-1 < d < 1
pr3:=derivable_pt_minus atan ps_atan d (derivable_pt_atan d) (derivable_pt_ps_atan d d_encad3):derivable_pt (atan - ps_atan) d
0 < x
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
pr:derivable_pt (atan - ps_atan) d
d_encad3:-1 < d < 1
pr3:=derivable_pt_minus atan ps_atan d (derivable_pt_atan d) (derivable_pt_ps_atan d d_encad3):derivable_pt (atan - ps_atan) d
0 < d < x
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
pr:derivable_pt (atan - ps_atan) d
d_encad3:-1 < d < 1
pr3:=derivable_pt_minus atan ps_atan d (derivable_pt_atan d) (derivable_pt_ps_atan d d_encad3):derivable_pt (atan - ps_atan) d
forall h : R, 0 < h < x -> (atan - ps_atan)%F h = (atan - ps_atan)%F h
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
Temp:forall pr : derivable_pt (atan - ps_atan) d, derive_pt (atan - ps_atan) d pr = 0
atan x = ps_atan x
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
pr:derivable_pt (atan - ps_atan) d
d_encad3:-1 < d < 1
pr3:=derivable_pt_minus atan ps_atan d (derivable_pt_atan d) (derivable_pt_ps_atan d d_encad3):derivable_pt (atan - ps_atan) d

0 < x
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
pr:derivable_pt (atan - ps_atan) d
d_encad3:-1 < d < 1
pr3:=derivable_pt_minus atan ps_atan d (derivable_pt_atan d) (derivable_pt_ps_atan d d_encad3):derivable_pt (atan - ps_atan) d
0 < d < x
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
pr:derivable_pt (atan - ps_atan) d
d_encad3:-1 < d < 1
pr3:=derivable_pt_minus atan ps_atan d (derivable_pt_atan d) (derivable_pt_ps_atan d d_encad3):derivable_pt (atan - ps_atan) d
forall h : R, 0 < h < x -> (atan - ps_atan)%F h = (atan - ps_atan)%F h
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
Temp:forall pr : derivable_pt (atan - ps_atan) d, derive_pt (atan - ps_atan) d pr = 0
atan x = ps_atan x
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
pr:derivable_pt (atan - ps_atan) d
d_encad3:-1 < d < 1
pr3:=derivable_pt_minus atan ps_atan d (derivable_pt_atan d) (derivable_pt_ps_atan d d_encad3):derivable_pt (atan - ps_atan) d

0 < d < x
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
pr:derivable_pt (atan - ps_atan) d
d_encad3:-1 < d < 1
pr3:=derivable_pt_minus atan ps_atan d (derivable_pt_atan d) (derivable_pt_ps_atan d d_encad3):derivable_pt (atan - ps_atan) d
forall h : R, 0 < h < x -> (atan - ps_atan)%F h = (atan - ps_atan)%F h
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
Temp:forall pr : derivable_pt (atan - ps_atan) d, derive_pt (atan - ps_atan) d pr = 0
atan x = ps_atan x
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
pr:derivable_pt (atan - ps_atan) d
d_encad3:-1 < d < 1
pr3:=derivable_pt_minus atan ps_atan d (derivable_pt_atan d) (derivable_pt_ps_atan d d_encad3):derivable_pt (atan - ps_atan) d

forall h : R, 0 < h < x -> (atan - ps_atan)%F h = (atan - ps_atan)%F h
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
Temp:forall pr : derivable_pt (atan - ps_atan) d, derive_pt (atan - ps_atan) d pr = 0
atan x = ps_atan x
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
Temp:forall pr : derivable_pt (atan - ps_atan) d, derive_pt (atan - ps_atan) d pr = 0

atan x = ps_atan x
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
Temp:forall pr : derivable_pt (atan - ps_atan) d, derive_pt (atan - ps_atan) d pr = 0

atan 0 = 0
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
Temp:forall pr : derivable_pt (atan - ps_atan) d, derive_pt (atan - ps_atan) d pr = 0
iatan0:atan 0 = 0
atan x = ps_atan x
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
Temp:forall pr : derivable_pt (atan - ps_atan) d, derive_pt (atan - ps_atan) d pr = 0

- PI / 2 < atan 0 < PI / 2
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
Temp:forall pr : derivable_pt (atan - ps_atan) d, derive_pt (atan - ps_atan) d pr = 0
- PI / 2 < 0 < PI / 2
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
Temp:forall pr : derivable_pt (atan - ps_atan) d, derive_pt (atan - ps_atan) d pr = 0
tan (atan 0) = tan 0
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
Temp:forall pr : derivable_pt (atan - ps_atan) d, derive_pt (atan - ps_atan) d pr = 0
iatan0:atan 0 = 0
atan x = ps_atan x
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
Temp:forall pr : derivable_pt (atan - ps_atan) d, derive_pt (atan - ps_atan) d pr = 0

- PI / 2 < 0 < PI / 2
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
Temp:forall pr : derivable_pt (atan - ps_atan) d, derive_pt (atan - ps_atan) d pr = 0
tan (atan 0) = tan 0
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
Temp:forall pr : derivable_pt (atan - ps_atan) d, derive_pt (atan - ps_atan) d pr = 0
iatan0:atan 0 = 0
atan x = ps_atan x
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
Temp:forall pr : derivable_pt (atan - ps_atan) d, derive_pt (atan - ps_atan) d pr = 0

tan (atan 0) = tan 0
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
Temp:forall pr : derivable_pt (atan - ps_atan) d, derive_pt (atan - ps_atan) d pr = 0
iatan0:atan 0 = 0
atan x = ps_atan x
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
Temp:forall pr : derivable_pt (atan - ps_atan) d, derive_pt (atan - ps_atan) d pr = 0
iatan0:atan 0 = 0

atan x = ps_atan x
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
Temp:forall pr : derivable_pt (atan - ps_atan) d, derive_pt (atan - ps_atan) d pr = 0
iatan0:atan 0 = 0

0 = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad) -> atan x = ps_atan x
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
Temp:forall pr : derivable_pt (atan - ps_atan) d, derive_pt (atan - ps_atan) d pr = 0
iatan0:atan 0 = 0

0 = (atan x - ps_atan x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad) -> atan x = ps_atan x
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
Temp:forall pr : derivable_pt (atan - ps_atan) d, derive_pt (atan - ps_atan) d pr = 0
iatan0:atan 0 = 0

0 = (atan x - ps_atan x - (atan 0 - ps_atan 0)) * derive_pt id d (pr2 d d_encad) -> atan x = ps_atan x
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
Temp:forall pr : derivable_pt (atan - ps_atan) d, derive_pt (atan - ps_atan) d pr = 0
iatan0:atan 0 = 0

0 = (atan x - ps_atan x) * derive_pt id d (pr2 d d_encad) -> atan x = ps_atan x
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
Temp:forall pr : derivable_pt (atan - ps_atan) d, derive_pt (atan - ps_atan) d pr = 0
iatan0:atan 0 = 0

0 = (atan x - ps_atan x) * 1 -> atan x = ps_atan x
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
Temp:forall pr : derivable_pt (atan - ps_atan) d, derive_pt (atan - ps_atan) d pr = 0
iatan0:atan 0 = 0
1 = derive_pt id d (pr2 d d_encad)
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
Temp:forall pr : derivable_pt (atan - ps_atan) d, derive_pt (atan - ps_atan) d pr = 0
iatan0:atan 0 = 0

0 = atan x - ps_atan x -> atan x = ps_atan x
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
Temp:forall pr : derivable_pt (atan - ps_atan) d, derive_pt (atan - ps_atan) d pr = 0
iatan0:atan 0 = 0
1 = derive_pt id d (pr2 d d_encad)
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
Temp:forall pr : derivable_pt (atan - ps_atan) d, derive_pt (atan - ps_atan) d pr = 0
iatan0:atan 0 = 0

1 = derive_pt id d (pr2 d d_encad)
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
Temp:forall pr : derivable_pt (atan - ps_atan) d, derive_pt (atan - ps_atan) d pr = 0
iatan0:atan 0 = 0

1 = derive_pt id d (derivable_pt_id d)
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
Temp:forall pr : derivable_pt (atan - ps_atan) d, derive_pt (atan - ps_atan) d pr = 0
iatan0:atan 0 = 0
id = id
x:R
x_encad:0 < x < 1
pr1:forall c : R, 0 < c < x -> derivable_pt (atan - ps_atan) c
pr2:forall c : R, 0 < c < x -> derivable_pt id c
d:R
d_encad:0 < d < x
Main:(id x - id 0) * derive_pt (atan - ps_atan) d (pr1 d d_encad) = ((atan - ps_atan)%F x - (atan - ps_atan)%F 0) * derive_pt id d (pr2 d d_encad)
Temp:forall pr : derivable_pt (atan - ps_atan) d, derive_pt (atan - ps_atan) d pr = 0
iatan0:atan 0 = 0

id = id
tauto. Qed.

Alt_PI = PI

Alt_PI = PI

Alt_PI / 4 = PI / 4
H:0 < PI / 6

Alt_PI / 4 = PI / 4
H:0 < PI / 6
t1:1 < PI / 2

Alt_PI / 4 = PI / 4
H:0 < PI / 6
t1:1 < PI / 2
t2:PI <= 4

Alt_PI / 4 = PI / 4
H:0 < PI / 6
t1:1 < PI / 2
t2:PI <= 4
m:0 < Alt_PI

Alt_PI / 4 = PI / 4
H:0 < PI / 6
t1:1 < PI / 2
t2:PI <= 4
m:0 < Alt_PI
H0:- PI / 2 < 1 < PI / 2

Alt_PI / 4 = PI / 4
H:0 < PI / 6
t1:1 < PI / 2
t2:PI <= 4
m:0 < Alt_PI
H0:- PI / 2 < 1 < PI / 2
eps:R
ep:0 < eps

Rabs (Alt_PI / 4 - PI / 4) < eps
H:0 < PI / 6
t1:1 < PI / 2
t2:PI <= 4
m:0 < Alt_PI
H0:- PI / 2 < 1 < PI / 2
eps:R
ep:0 < eps

R_dist (Alt_PI / 4) (PI / 4) < eps
H:0 < PI / 6
t1:1 < PI / 2
t2:PI <= 4
m:0 < Alt_PI
H0:- PI / 2 < 1 < PI / 2
eps:R
ep:0 < eps

continuity_pt atan 1
H:0 < PI / 6
t1:1 < PI / 2
t2:PI <= 4
m:0 < Alt_PI
H0:- PI / 2 < 1 < PI / 2
eps:R
ep:0 < eps
ca:continuity_pt atan 1
R_dist (Alt_PI / 4) (PI / 4) < eps
H:0 < PI / 6
t1:1 < PI / 2
t2:PI <= 4
m:0 < Alt_PI
H0:- PI / 2 < 1 < PI / 2
eps:R
ep:0 < eps
ca:continuity_pt atan 1

R_dist (Alt_PI / 4) (PI / 4) < eps
H:0 < PI / 6
t1:1 < PI / 2
t2:PI <= 4
m:0 < Alt_PI
H0:- PI / 2 < 1 < PI / 2
eps:R
ep:0 < eps
ca:continuity_pt atan 1

exists eps' eps'' : R, eps' + eps'' <= eps /\ 0 < eps' /\ 0 < eps''
H:0 < PI / 6
t1:1 < PI / 2
t2:PI <= 4
m:0 < Alt_PI
H0:- PI / 2 < 1 < PI / 2
eps:R
ep:0 < eps
ca:continuity_pt atan 1
Xe:exists eps' eps'' : R, eps' + eps'' <= eps /\ 0 < eps' /\ 0 < eps''
R_dist (Alt_PI / 4) (PI / 4) < eps
H:0 < PI / 6
t1:1 < PI / 2
t2:PI <= 4
m:0 < Alt_PI
H0:- PI / 2 < 1 < PI / 2
eps:R
ep:0 < eps
ca:continuity_pt atan 1
Xe:exists eps' eps'' : R, eps' + eps'' <= eps /\ 0 < eps' /\ 0 < eps''

R_dist (Alt_PI / 4) (PI / 4) < eps
H:0 < PI / 6
t1:1 < PI / 2
t2:PI <= 4
m:0 < Alt_PI
H0:- PI / 2 < 1 < PI / 2
eps:R
ep:0 < eps
ca:continuity_pt atan 1
eps', eps'':R
eps_ineq:eps' + eps'' <= eps
ep':0 < eps'
ep'':0 < eps''

R_dist (Alt_PI / 4) (PI / 4) < eps
H:0 < PI / 6
t1:1 < PI / 2
t2:PI <= 4
m:0 < Alt_PI
H0:- PI / 2 < 1 < PI / 2
eps:R
ep:0 < eps
ca:continuity_pt atan 1
eps', eps'':R
eps_ineq:eps' + eps'' <= eps
ep':0 < eps'
ep'':0 < eps''
alpha:R
a0:alpha > 0
Palpha:forall x : R, x < 1 -> 0 < x -> R_dist x 1 < alpha -> dist R_met (ps_atan x) (Alt_PI / 4) < eps'

R_dist (Alt_PI / 4) (PI / 4) < eps
H:0 < PI / 6
t1:1 < PI / 2
t2:PI <= 4
m:0 < Alt_PI
H0:- PI / 2 < 1 < PI / 2
eps:R
ep:0 < eps
ca:continuity_pt atan 1
eps', eps'':R
eps_ineq:eps' + eps'' <= eps
ep':0 < eps'
ep'':0 < eps''
alpha:R
a0:alpha > 0
Palpha:forall x : R, x < 1 -> 0 < x -> R_dist x 1 < alpha -> dist R_met (ps_atan x) (Alt_PI / 4) < eps'
beta:R
b0:beta > 0
Pbeta:forall x : Base R_met, D_x no_cond 1 x /\ dist R_met x 1 < beta -> dist R_met (atan x) (atan 1) < eps''

R_dist (Alt_PI / 4) (PI / 4) < eps
H:0 < PI / 6
t1:1 < PI / 2
t2:PI <= 4
m:0 < Alt_PI
H0:- PI / 2 < 1 < PI / 2
eps:R
ep:0 < eps
ca:continuity_pt atan 1
eps', eps'':R
eps_ineq:eps' + eps'' <= eps
ep':0 < eps'
ep'':0 < eps''
alpha:R
a0:alpha > 0
Palpha:forall x : R, x < 1 -> 0 < x -> R_dist x 1 < alpha -> dist R_met (ps_atan x) (Alt_PI / 4) < eps'
beta:R
b0:beta > 0
Pbeta:forall x : Base R_met, D_x no_cond 1 x /\ dist R_met x 1 < beta -> dist R_met (atan x) (atan 1) < eps''

exists a : R, 0 < a < 1 /\ R_dist a 1 < alpha /\ R_dist a 1 < beta
H:0 < PI / 6
t1:1 < PI / 2
t2:PI <= 4
m:0 < Alt_PI
H0:- PI / 2 < 1 < PI / 2
eps:R
ep:0 < eps
ca:continuity_pt atan 1
eps', eps'':R
eps_ineq:eps' + eps'' <= eps
ep':0 < eps'
ep'':0 < eps''
alpha:R
a0:alpha > 0
Palpha:forall x : R, x < 1 -> 0 < x -> R_dist x 1 < alpha -> dist R_met (ps_atan x) (Alt_PI / 4) < eps'
beta:R
b0:beta > 0
Pbeta:forall x : Base R_met, D_x no_cond 1 x /\ dist R_met x 1 < beta -> dist R_met (atan x) (atan 1) < eps''
Xa:exists a : R, 0 < a < 1 /\ R_dist a 1 < alpha /\ R_dist a 1 < beta
R_dist (Alt_PI / 4) (PI / 4) < eps
H:0 < PI / 6
t1:1 < PI / 2
t2:PI <= 4
m:0 < Alt_PI
H0:- PI / 2 < 1 < PI / 2
eps:R
ep:0 < eps
ca:continuity_pt atan 1
eps', eps'':R
eps_ineq:eps' + eps'' <= eps
ep':0 < eps'
ep'':0 < eps''
alpha:R
a0:alpha > 0
Palpha:forall x : R, x < 1 -> 0 < x -> R_dist x 1 < alpha -> dist R_met (ps_atan x) (Alt_PI / 4) < eps'
beta:R
b0:beta > 0
Pbeta:forall x : Base R_met, D_x no_cond 1 x /\ dist R_met x 1 < beta -> dist R_met (atan x) (atan 1) < eps''

0 < Rmax (/ 2) (Rmax (1 - alpha / 2) (1 - beta / 2)) < 1 /\ R_dist (Rmax (/ 2) (Rmax (1 - alpha / 2) (1 - beta / 2))) 1 < alpha /\ R_dist (Rmax (/ 2) (Rmax (1 - alpha / 2) (1 - beta / 2))) 1 < beta
H:0 < PI / 6
t1:1 < PI / 2
t2:PI <= 4
m:0 < Alt_PI
H0:- PI / 2 < 1 < PI / 2
eps:R
ep:0 < eps
ca:continuity_pt atan 1
eps', eps'':R
eps_ineq:eps' + eps'' <= eps
ep':0 < eps'
ep'':0 < eps''
alpha:R
a0:alpha > 0
Palpha:forall x : R, x < 1 -> 0 < x -> R_dist x 1 < alpha -> dist R_met (ps_atan x) (Alt_PI / 4) < eps'
beta:R
b0:beta > 0
Pbeta:forall x : Base R_met, D_x no_cond 1 x /\ dist R_met x 1 < beta -> dist R_met (atan x) (atan 1) < eps''
Xa:exists a : R, 0 < a < 1 /\ R_dist a 1 < alpha /\ R_dist a 1 < beta
R_dist (Alt_PI / 4) (PI / 4) < eps
H:0 < PI / 6
t1:1 < PI / 2
t2:PI <= 4
m:0 < Alt_PI
H0:- PI / 2 < 1 < PI / 2
eps:R
ep:0 < eps
ca:continuity_pt atan 1
eps', eps'':R
eps_ineq:eps' + eps'' <= eps
ep':0 < eps'
ep'':0 < eps''
alpha:R
a0:alpha > 0
Palpha:forall x : R, x < 1 -> 0 < x -> R_dist x 1 < alpha -> dist R_met (ps_atan x) (Alt_PI / 4) < eps'
beta:R
b0:beta > 0
Pbeta:forall x : Base R_met, D_x no_cond 1 x /\ dist R_met x 1 < beta -> dist R_met (atan x) (atan 1) < eps''
H1:/ 2 <= Rmax (/ 2) (Rmax (1 - alpha / 2) (1 - beta / 2))

0 < Rmax (/ 2) (Rmax (1 - alpha / 2) (1 - beta / 2)) < 1 /\ R_dist (Rmax (/ 2) (Rmax (1 - alpha / 2) (1 - beta / 2))) 1 < alpha /\ R_dist (Rmax (/ 2) (Rmax (1 - alpha / 2) (1 - beta / 2))) 1 < beta
H:0 < PI / 6
t1:1 < PI / 2
t2:PI <= 4
m:0 < Alt_PI
H0:- PI / 2 < 1 < PI / 2
eps:R
ep:0 < eps
ca:continuity_pt atan 1
eps', eps'':R
eps_ineq:eps' + eps'' <= eps
ep':0 < eps'
ep'':0 < eps''
alpha:R
a0:alpha > 0
Palpha:forall x : R, x < 1 -> 0 < x -> R_dist x 1 < alpha -> dist R_met (ps_atan x) (Alt_PI / 4) < eps'
beta:R
b0:beta > 0
Pbeta:forall x : Base R_met, D_x no_cond 1 x /\ dist R_met x 1 < beta -> dist R_met (atan x) (atan 1) < eps''
Xa:exists a : R, 0 < a < 1 /\ R_dist a 1 < alpha /\ R_dist a 1 < beta
R_dist (Alt_PI / 4) (PI / 4) < eps
H:0 < PI / 6
t1:1 < PI / 2
t2:PI <= 4
m:0 < Alt_PI
H0:- PI / 2 < 1 < PI / 2
eps:R
ep:0 < eps
ca:continuity_pt atan 1
eps', eps'':R
eps_ineq:eps' + eps'' <= eps
ep':0 < eps'
ep'':0 < eps''
alpha:R
a0:alpha > 0
Palpha:forall x : R, x < 1 -> 0 < x -> R_dist x 1 < alpha -> dist R_met (ps_atan x) (Alt_PI / 4) < eps'
beta:R
b0:beta > 0
Pbeta:forall x : Base R_met, D_x no_cond 1 x /\ dist R_met x 1 < beta -> dist R_met (atan x) (atan 1) < eps''
H1:/ 2 <= Rmax (/ 2) (Rmax (1 - alpha / 2) (1 - beta / 2))
H2:Rmax (1 - alpha / 2) (1 - beta / 2) <= Rmax (/ 2) (Rmax (1 - alpha / 2) (1 - beta / 2))

0 < Rmax (/ 2) (Rmax (1 - alpha / 2) (1 - beta / 2)) < 1 /\ R_dist (Rmax (/ 2) (Rmax (1 - alpha / 2) (1 - beta / 2))) 1 < alpha /\ R_dist (Rmax (/ 2) (Rmax (1 - alpha / 2) (1 - beta / 2))) 1 < beta
H:0 < PI / 6
t1:1 < PI / 2
t2:PI <= 4
m:0 < Alt_PI
H0:- PI / 2 < 1 < PI / 2
eps:R
ep:0 < eps
ca:continuity_pt atan 1
eps', eps'':R
eps_ineq:eps' + eps'' <= eps
ep':0 < eps'
ep'':0 < eps''
alpha:R
a0:alpha > 0
Palpha:forall x : R, x < 1 -> 0 < x -> R_dist x 1 < alpha -> dist R_met (ps_atan x) (Alt_PI / 4) < eps'
beta:R
b0:beta > 0
Pbeta:forall x : Base R_met, D_x no_cond 1 x /\ dist R_met x 1 < beta -> dist R_met (atan x) (atan 1) < eps''
Xa:exists a : R, 0 < a < 1 /\ R_dist a 1 < alpha /\ R_dist a 1 < beta
R_dist (Alt_PI / 4) (PI / 4) < eps
H:0 < PI / 6
t1:1 < PI / 2
t2:PI <= 4
m:0 < Alt_PI
H0:- PI / 2 < 1 < PI / 2
eps:R
ep:0 < eps
ca:continuity_pt atan 1
eps', eps'':R
eps_ineq:eps' + eps'' <= eps
ep':0 < eps'
ep'':0 < eps''
alpha:R
a0:alpha > 0
Palpha:forall x : R, x < 1 -> 0 < x -> R_dist x 1 < alpha -> dist R_met (ps_atan x) (Alt_PI / 4) < eps'
beta:R
b0:beta > 0
Pbeta:forall x : Base R_met, D_x no_cond 1 x /\ dist R_met x 1 < beta -> dist R_met (atan x) (atan 1) < eps''
H1:/ 2 <= Rmax (/ 2) (Rmax (1 - alpha / 2) (1 - beta / 2))
H2:Rmax (1 - alpha / 2) (1 - beta / 2) <= Rmax (/ 2) (Rmax (1 - alpha / 2) (1 - beta / 2))
H3:1 - alpha / 2 <= Rmax (1 - alpha / 2) (1 - beta / 2)

0 < Rmax (/ 2) (Rmax (1 - alpha / 2) (1 - beta / 2)) < 1 /\ R_dist (Rmax (/ 2) (Rmax (1 - alpha / 2) (1 - beta / 2))) 1 < alpha /\ R_dist (Rmax (/ 2) (Rmax (1 - alpha / 2) (1 - beta / 2))) 1 < beta
H:0 < PI / 6
t1:1 < PI / 2
t2:PI <= 4
m:0 < Alt_PI
H0:- PI / 2 < 1 < PI / 2
eps:R
ep:0 < eps
ca:continuity_pt atan 1
eps', eps'':R
eps_ineq:eps' + eps'' <= eps
ep':0 < eps'
ep'':0 < eps''
alpha:R
a0:alpha > 0
Palpha:forall x : R, x < 1 -> 0 < x -> R_dist x 1 < alpha -> dist R_met (ps_atan x) (Alt_PI / 4) < eps'
beta:R
b0:beta > 0
Pbeta:forall x : Base R_met, D_x no_cond 1 x /\ dist R_met x 1 < beta -> dist R_met (atan x) (atan 1) < eps''
Xa:exists a : R, 0 < a < 1 /\ R_dist a 1 < alpha /\ R_dist a 1 < beta
R_dist (Alt_PI / 4) (PI / 4) < eps
H:0 < PI / 6
t1:1 < PI / 2
t2:PI <= 4
m:0 < Alt_PI
H0:- PI / 2 < 1 < PI / 2
eps:R
ep:0 < eps
ca:continuity_pt atan 1
eps', eps'':R
eps_ineq:eps' + eps'' <= eps
ep':0 < eps'
ep'':0 < eps''
alpha:R
a0:alpha > 0
Palpha:forall x : R, x < 1 -> 0 < x -> R_dist x 1 < alpha -> dist R_met (ps_atan x) (Alt_PI / 4) < eps'
beta:R
b0:beta > 0
Pbeta:forall x : Base R_met, D_x no_cond 1 x /\ dist R_met x 1 < beta -> dist R_met (atan x) (atan 1) < eps''
H1:/ 2 <= Rmax (/ 2) (Rmax (1 - alpha / 2) (1 - beta / 2))
H2:Rmax (1 - alpha / 2) (1 - beta / 2) <= Rmax (/ 2) (Rmax (1 - alpha / 2) (1 - beta / 2))
H3:1 - alpha / 2 <= Rmax (1 - alpha / 2) (1 - beta / 2)
H4:1 - beta / 2 <= Rmax (1 - alpha / 2) (1 - beta / 2)

0 < Rmax (/ 2) (Rmax (1 - alpha / 2) (1 - beta / 2)) < 1 /\ R_dist (Rmax (/ 2) (Rmax (1 - alpha / 2) (1 - beta / 2))) 1 < alpha /\ R_dist (Rmax (/ 2) (Rmax (1 - alpha / 2) (1 - beta / 2))) 1 < beta
H:0 < PI / 6
t1:1 < PI / 2
t2:PI <= 4
m:0 < Alt_PI
H0:- PI / 2 < 1 < PI / 2
eps:R
ep:0 < eps
ca:continuity_pt atan 1
eps', eps'':R
eps_ineq:eps' + eps'' <= eps
ep':0 < eps'
ep'':0 < eps''
alpha:R
a0:alpha > 0
Palpha:forall x : R, x < 1 -> 0 < x -> R_dist x 1 < alpha -> dist R_met (ps_atan x) (Alt_PI / 4) < eps'
beta:R
b0:beta > 0
Pbeta:forall x : Base R_met, D_x no_cond 1 x /\ dist R_met x 1 < beta -> dist R_met (atan x) (atan 1) < eps''
Xa:exists a : R, 0 < a < 1 /\ R_dist a 1 < alpha /\ R_dist a 1 < beta
R_dist (Alt_PI / 4) (PI / 4) < eps
H:0 < PI / 6
t1:1 < PI / 2
t2:PI <= 4
m:0 < Alt_PI
H0:- PI / 2 < 1 < PI / 2
eps:R
ep:0 < eps
ca:continuity_pt atan 1
eps', eps'':R
eps_ineq:eps' + eps'' <= eps
ep':0 < eps'
ep'':0 < eps''
alpha:R
a0:alpha > 0
Palpha:forall x : R, x < 1 -> 0 < x -> R_dist x 1 < alpha -> dist R_met (ps_atan x) (Alt_PI / 4) < eps'
beta:R
b0:beta > 0
Pbeta:forall x : Base R_met, D_x no_cond 1 x /\ dist R_met x 1 < beta -> dist R_met (atan x) (atan 1) < eps''
H1:/ 2 <= Rmax (/ 2) (Rmax (1 - alpha / 2) (1 - beta / 2))
H2:Rmax (1 - alpha / 2) (1 - beta / 2) <= Rmax (/ 2) (Rmax (1 - alpha / 2) (1 - beta / 2))
H3:1 - alpha / 2 <= Rmax (1 - alpha / 2) (1 - beta / 2)
H4:1 - beta / 2 <= Rmax (1 - alpha / 2) (1 - beta / 2)
H5:Rmax (1 - alpha / 2) (1 - beta / 2) < 1

0 < Rmax (/ 2) (Rmax (1 - alpha / 2) (1 - beta / 2)) < 1 /\ R_dist (Rmax (/ 2) (Rmax (1 - alpha / 2) (1 - beta / 2))) 1 < alpha /\ R_dist (Rmax (/ 2) (Rmax (1 - alpha / 2) (1 - beta / 2))) 1 < beta
H:0 < PI / 6
t1:1 < PI / 2
t2:PI <= 4
m:0 < Alt_PI
H0:- PI / 2 < 1 < PI / 2
eps:R
ep:0 < eps
ca:continuity_pt atan 1
eps', eps'':R
eps_ineq:eps' + eps'' <= eps
ep':0 < eps'
ep'':0 < eps''
alpha:R
a0:alpha > 0
Palpha:forall x : R, x < 1 -> 0 < x -> R_dist x 1 < alpha -> dist R_met (ps_atan x) (Alt_PI / 4) < eps'
beta:R
b0:beta > 0
Pbeta:forall x : Base R_met, D_x no_cond 1 x /\ dist R_met x 1 < beta -> dist R_met (atan x) (atan 1) < eps''
Xa:exists a : R, 0 < a < 1 /\ R_dist a 1 < alpha /\ R_dist a 1 < beta
R_dist (Alt_PI / 4) (PI / 4) < eps
H:0 < PI / 6
t1:1 < PI / 2
t2:PI <= 4
m:0 < Alt_PI
H0:- PI / 2 < 1 < PI / 2
eps:R
ep:0 < eps
ca:continuity_pt atan 1
eps', eps'':R
eps_ineq:eps' + eps'' <= eps
ep':0 < eps'
ep'':0 < eps''
alpha:R
a0:alpha > 0
Palpha:forall x : R, x < 1 -> 0 < x -> R_dist x 1 < alpha -> dist R_met (ps_atan x) (Alt_PI / 4) < eps'
beta:R
b0:beta > 0
Pbeta:forall x : Base R_met, D_x no_cond 1 x /\ dist R_met x 1 < beta -> dist R_met (atan x) (atan 1) < eps''
H1:/ 2 <= Rmax (/ 2) (Rmax (1 - alpha / 2) (1 - beta / 2))
H2:Rmax (1 - alpha / 2) (1 - beta / 2) <= Rmax (/ 2) (Rmax (1 - alpha / 2) (1 - beta / 2))
H3:1 - alpha / 2 <= Rmax (1 - alpha / 2) (1 - beta / 2)
H4:1 - beta / 2 <= Rmax (1 - alpha / 2) (1 - beta / 2)
H5:Rmax (1 - alpha / 2) (1 - beta / 2) < 1

R_dist (Rmax (/ 2) (Rmax (1 - alpha / 2) (1 - beta / 2))) 1 < alpha /\ R_dist (Rmax (/ 2) (Rmax (1 - alpha / 2) (1 - beta / 2))) 1 < beta
H:0 < PI / 6
t1:1 < PI / 2
t2:PI <= 4
m:0 < Alt_PI
H0:- PI / 2 < 1 < PI / 2
eps:R
ep:0 < eps
ca:continuity_pt atan 1
eps', eps'':R
eps_ineq:eps' + eps'' <= eps
ep':0 < eps'
ep'':0 < eps''
alpha:R
a0:alpha > 0
Palpha:forall x : R, x < 1 -> 0 < x -> R_dist x 1 < alpha -> dist R_met (ps_atan x) (Alt_PI / 4) < eps'
beta:R
b0:beta > 0
Pbeta:forall x : Base R_met, D_x no_cond 1 x /\ dist R_met x 1 < beta -> dist R_met (atan x) (atan 1) < eps''
Xa:exists a : R, 0 < a < 1 /\ R_dist a 1 < alpha /\ R_dist a 1 < beta
R_dist (Alt_PI / 4) (PI / 4) < eps
H:0 < PI / 6
t1:1 < PI / 2
t2:PI <= 4
m:0 < Alt_PI
H0:- PI / 2 < 1 < PI / 2
eps:R
ep:0 < eps
ca:continuity_pt atan 1
eps', eps'':R
eps_ineq:eps' + eps'' <= eps
ep':0 < eps'
ep'':0 < eps''
alpha:R
a0:alpha > 0
Palpha:forall x : R, x < 1 -> 0 < x -> R_dist x 1 < alpha -> dist R_met (ps_atan x) (Alt_PI / 4) < eps'
beta:R
b0:beta > 0
Pbeta:forall x : Base R_met, D_x no_cond 1 x /\ dist R_met x 1 < beta -> dist R_met (atan x) (atan 1) < eps''
H1:/ 2 <= Rmax (/ 2) (Rmax (1 - alpha / 2) (1 - beta / 2))
H2:Rmax (1 - alpha / 2) (1 - beta / 2) <= Rmax (/ 2) (Rmax (1 - alpha / 2) (1 - beta / 2))
H3:1 - alpha / 2 <= Rmax (1 - alpha / 2) (1 - beta / 2)
H4:1 - beta / 2 <= Rmax (1 - alpha / 2) (1 - beta / 2)
H5:Rmax (1 - alpha / 2) (1 - beta / 2) < 1

0 <= 1 - Rmax (/ 2) (Rmax (1 - alpha / 2) (1 - beta / 2))
H:0 < PI / 6
t1:1 < PI / 2
t2:PI <= 4
m:0 < Alt_PI
H0:- PI / 2 < 1 < PI / 2
eps:R
ep:0 < eps
ca:continuity_pt atan 1
eps', eps'':R
eps_ineq:eps' + eps'' <= eps
ep':0 < eps'
ep'':0 < eps''
alpha:R
a0:alpha > 0
Palpha:forall x : R, x < 1 -> 0 < x -> R_dist x 1 < alpha -> dist R_met (ps_atan x) (Alt_PI / 4) < eps'
beta:R
b0:beta > 0
Pbeta:forall x : Base R_met, D_x no_cond 1 x /\ dist R_met x 1 < beta -> dist R_met (atan x) (atan 1) < eps''
H1:/ 2 <= Rmax (/ 2) (Rmax (1 - alpha / 2) (1 - beta / 2))
H2:Rmax (1 - alpha / 2) (1 - beta / 2) <= Rmax (/ 2) (Rmax (1 - alpha / 2) (1 - beta / 2))
H3:1 - alpha / 2 <= Rmax (1 - alpha / 2) (1 - beta / 2)
H4:1 - beta / 2 <= Rmax (1 - alpha / 2) (1 - beta / 2)
H5:Rmax (1 - alpha / 2) (1 - beta / 2) < 1
H6:0 <= 1 - Rmax (/ 2) (Rmax (1 - alpha / 2) (1 - beta / 2))
R_dist (Rmax (/ 2) (Rmax (1 - alpha / 2) (1 - beta / 2))) 1 < alpha /\ R_dist (Rmax (/ 2) (Rmax (1 - alpha / 2) (1 - beta / 2))) 1 < beta
H:0 < PI / 6
t1:1 < PI / 2
t2:PI <= 4
m:0 < Alt_PI
H0:- PI / 2 < 1 < PI / 2
eps:R
ep:0 < eps
ca:continuity_pt atan 1
eps', eps'':R
eps_ineq:eps' + eps'' <= eps
ep':0 < eps'
ep'':0 < eps''
alpha:R
a0:alpha > 0
Palpha:forall x : R, x < 1 -> 0 < x -> R_dist x 1 < alpha -> dist R_met (ps_atan x) (Alt_PI / 4) < eps'
beta:R
b0:beta > 0
Pbeta:forall x : Base R_met, D_x no_cond 1 x /\ dist R_met x 1 < beta -> dist R_met (atan x) (atan 1) < eps''
Xa:exists a : R, 0 < a < 1 /\ R_dist a 1 < alpha /\ R_dist a 1 < beta
R_dist (Alt_PI / 4) (PI / 4) < eps
H:0 < PI / 6
t1:1 < PI / 2
t2:PI <= 4
m:0 < Alt_PI
H0:- PI / 2 < 1 < PI / 2
eps:R
ep:0 < eps
ca:continuity_pt atan 1
eps', eps'':R
eps_ineq:eps' + eps'' <= eps
ep':0 < eps'
ep'':0 < eps''
alpha:R
a0:alpha > 0
Palpha:forall x : R, x < 1 -> 0 < x -> R_dist x 1 < alpha -> dist R_met (ps_atan x) (Alt_PI / 4) < eps'
beta:R
b0:beta > 0
Pbeta:forall x : Base R_met, D_x no_cond 1 x /\ dist R_met x 1 < beta -> dist R_met (atan x) (atan 1) < eps''
H1:/ 2 <= Rmax (/ 2) (Rmax (1 - alpha / 2) (1 - beta / 2))
H2:Rmax (1 - alpha / 2) (1 - beta / 2) <= Rmax (/ 2) (Rmax (1 - alpha / 2) (1 - beta / 2))
H3:1 - alpha / 2 <= Rmax (1 - alpha / 2) (1 - beta / 2)
H4:1 - beta / 2 <= Rmax (1 - alpha / 2) (1 - beta / 2)
H5:Rmax (1 - alpha / 2) (1 - beta / 2) < 1
H6:Rmax (/ 2) (Rmax (1 - alpha / 2) (1 - beta / 2)) <= 1

0 <= 1 - Rmax (/ 2) (Rmax (1 - alpha / 2) (1 - beta / 2))
H:0 < PI / 6
t1:1 < PI / 2
t2:PI <= 4
m:0 < Alt_PI
H0:- PI / 2 < 1 < PI / 2
eps:R
ep:0 < eps
ca:continuity_pt atan 1
eps', eps'':R
eps_ineq:eps' + eps'' <= eps
ep':0 < eps'
ep'':0 < eps''
alpha:R
a0:alpha > 0
Palpha:forall x : R, x < 1 -> 0 < x -> R_dist x 1 < alpha -> dist R_met (ps_atan x) (Alt_PI / 4) < eps'
beta:R
b0:beta > 0
Pbeta:forall x : Base R_met, D_x no_cond 1 x /\ dist R_met x 1 < beta -> dist R_met (atan x) (atan 1) < eps''
H1:/ 2 <= Rmax (/ 2) (Rmax (1 - alpha / 2) (1 - beta / 2))
H2:Rmax (1 - alpha / 2) (1 - beta / 2) <= Rmax (/ 2) (Rmax (1 - alpha / 2) (1 - beta / 2))
H3:1 - alpha / 2 <= Rmax (1 - alpha / 2) (1 - beta / 2)
H4:1 - beta / 2 <= Rmax (1 - alpha / 2) (1 - beta / 2)
H5:Rmax (1 - alpha / 2) (1 - beta / 2) < 1
H6:0 <= 1 - Rmax (/ 2) (Rmax (1 - alpha / 2) (1 - beta / 2))
R_dist (Rmax (/ 2) (Rmax (1 - alpha / 2) (1 - beta / 2))) 1 < alpha /\ R_dist (Rmax (/ 2) (Rmax (1 - alpha / 2) (1 - beta / 2))) 1 < beta
H:0 < PI / 6
t1:1 < PI / 2
t2:PI <= 4
m:0 < Alt_PI
H0:- PI / 2 < 1 < PI / 2
eps:R
ep:0 < eps
ca:continuity_pt atan 1
eps', eps'':R
eps_ineq:eps' + eps'' <= eps
ep':0 < eps'
ep'':0 < eps''
alpha:R
a0:alpha > 0
Palpha:forall x : R, x < 1 -> 0 < x -> R_dist x 1 < alpha -> dist R_met (ps_atan x) (Alt_PI / 4) < eps'
beta:R
b0:beta > 0
Pbeta:forall x : Base R_met, D_x no_cond 1 x /\ dist R_met x 1 < beta -> dist R_met (atan x) (atan 1) < eps''
Xa:exists a : R, 0 < a < 1 /\ R_dist a 1 < alpha /\ R_dist a 1 < beta
R_dist (Alt_PI / 4) (PI / 4) < eps
H:0 < PI / 6
t1:1 < PI / 2
t2:PI <= 4
m:0 < Alt_PI
H0:- PI / 2 < 1 < PI / 2
eps:R
ep:0 < eps
ca:continuity_pt atan 1
eps', eps'':R
eps_ineq:eps' + eps'' <= eps
ep':0 < eps'
ep'':0 < eps''
alpha:R
a0:alpha > 0
Palpha:forall x : R, x < 1 -> 0 < x -> R_dist x 1 < alpha -> dist R_met (ps_atan x) (Alt_PI / 4) < eps'
beta:R
b0:beta > 0
Pbeta:forall x : Base R_met, D_x no_cond 1 x /\ dist R_met x 1 < beta -> dist R_met (atan x) (atan 1) < eps''
H1:/ 2 <= Rmax (/ 2) (Rmax (1 - alpha / 2) (1 - beta / 2))
H2:Rmax (1 - alpha / 2) (1 - beta / 2) <= Rmax (/ 2) (Rmax (1 - alpha / 2) (1 - beta / 2))
H3:1 - alpha / 2 <= Rmax (1 - alpha / 2) (1 - beta / 2)
H4:1 - beta / 2 <= Rmax (1 - alpha / 2) (1 - beta / 2)
H5:Rmax (1 - alpha / 2) (1 - beta / 2) < 1
H6:0 <= 1 - Rmax (/ 2) (Rmax (1 - alpha / 2) (1 - beta / 2))

R_dist (Rmax (/ 2) (Rmax (1 - alpha / 2) (1 - beta / 2))) 1 < alpha /\ R_dist (Rmax (/ 2) (Rmax (1 - alpha / 2) (1 - beta / 2))) 1 < beta
H:0 < PI / 6
t1:1 < PI / 2
t2:PI <= 4
m:0 < Alt_PI
H0:- PI / 2 < 1 < PI / 2
eps:R
ep:0 < eps
ca:continuity_pt atan 1
eps', eps'':R
eps_ineq:eps' + eps'' <= eps
ep':0 < eps'
ep'':0 < eps''
alpha:R
a0:alpha > 0
Palpha:forall x : R, x < 1 -> 0 < x -> R_dist x 1 < alpha -> dist R_met (ps_atan x) (Alt_PI / 4) < eps'
beta:R
b0:beta > 0
Pbeta:forall x : Base R_met, D_x no_cond 1 x /\ dist R_met x 1 < beta -> dist R_met (atan x) (atan 1) < eps''
Xa:exists a : R, 0 < a < 1 /\ R_dist a 1 < alpha /\ R_dist a 1 < beta
R_dist (Alt_PI / 4) (PI / 4) < eps
H:0 < PI / 6
t1:1 < PI / 2
t2:PI <= 4
m:0 < Alt_PI
H0:- PI / 2 < 1 < PI / 2
eps:R
ep:0 < eps
ca:continuity_pt atan 1
eps', eps'':R
eps_ineq:eps' + eps'' <= eps
ep':0 < eps'
ep'':0 < eps''
alpha:R
a0:alpha > 0
Palpha:forall x : R, x < 1 -> 0 < x -> R_dist x 1 < alpha -> dist R_met (ps_atan x) (Alt_PI / 4) < eps'
beta:R
b0:beta > 0
Pbeta:forall x : Base R_met, D_x no_cond 1 x /\ dist R_met x 1 < beta -> dist R_met (atan x) (atan 1) < eps''
Xa:exists a : R, 0 < a < 1 /\ R_dist a 1 < alpha /\ R_dist a 1 < beta

R_dist (Alt_PI / 4) (PI / 4) < eps
H:0 < PI / 6
t1:1 < PI / 2
t2:PI <= 4
m:0 < Alt_PI
H0:- PI / 2 < 1 < PI / 2
eps:R
ep:0 < eps
ca:continuity_pt atan 1
eps', eps'':R
eps_ineq:eps' + eps'' <= eps
ep':0 < eps'
ep'':0 < eps''
alpha:R
a0:alpha > 0
Palpha:forall x : R, x < 1 -> 0 < x -> R_dist x 1 < alpha -> dist R_met (ps_atan x) (Alt_PI / 4) < eps'
beta:R
b0:beta > 0
Pbeta:forall x : Base R_met, D_x no_cond 1 x /\ dist R_met x 1 < beta -> dist R_met (atan x) (atan 1) < eps''
a:R
Pa0:0 < a
Pa1:a < 1
P1:R_dist a 1 < alpha
P2:R_dist a 1 < beta

R_dist (Alt_PI / 4) (PI / 4) < eps
H:0 < PI / 6
t1:1 < PI / 2
t2:PI <= 4
m:0 < Alt_PI
H0:- PI / 2 < 1 < PI / 2
eps:R
ep:0 < eps
ca:continuity_pt atan 1
eps', eps'':R
eps_ineq:eps' + eps'' <= eps
ep':0 < eps'
ep'':0 < eps''
alpha:R
a0:alpha > 0
Palpha:forall x : R, x < 1 -> 0 < x -> R_dist x 1 < alpha -> dist R_met (ps_atan x) (Alt_PI / 4) < eps'
beta:R
b0:beta > 0
Pbeta:forall x : Base R_met, D_x no_cond 1 x /\ dist R_met x 1 < beta -> dist R_met (atan x) (atan 1) < eps''
a:R
Pa0:0 < a
Pa1:a < 1
P1:R_dist a 1 < alpha
P2:R_dist a 1 < beta

R_dist (Alt_PI / 4) (ps_atan a) + R_dist (ps_atan a) (PI / 4) < eps
H:0 < PI / 6
t1:1 < PI / 2
t2:PI <= 4
m:0 < Alt_PI
H0:- PI / 2 < 1 < PI / 2
eps:R
ep:0 < eps
ca:continuity_pt atan 1
eps', eps'':R
eps_ineq:eps' + eps'' <= eps
ep':0 < eps'
ep'':0 < eps''
alpha:R
a0:alpha > 0
Palpha:forall x : R, x < 1 -> 0 < x -> R_dist x 1 < alpha -> dist R_met (ps_atan x) (Alt_PI / 4) < eps'
beta:R
b0:beta > 0
Pbeta:forall x : Base R_met, D_x no_cond 1 x /\ dist R_met x 1 < beta -> dist R_met (atan x) (atan 1) < eps''
a:R
Pa0:0 < a
Pa1:a < 1
P1:R_dist a 1 < alpha
P2:R_dist a 1 < beta

R_dist (Alt_PI / 4) (ps_atan a) + R_dist (ps_atan a) (PI / 4) < eps' + eps''
H:0 < PI / 6
t1:1 < PI / 2
t2:PI <= 4
m:0 < Alt_PI
H0:- PI / 2 < 1 < PI / 2
eps:R
ep:0 < eps
ca:continuity_pt atan 1
eps', eps'':R
eps_ineq:eps' + eps'' <= eps
ep':0 < eps'
ep'':0 < eps''
alpha:R
a0:alpha > 0
Palpha:forall x : R, x < 1 -> 0 < x -> R_dist x 1 < alpha -> dist R_met (ps_atan x) (Alt_PI / 4) < eps'
beta:R
b0:beta > 0
Pbeta:forall x : Base R_met, D_x no_cond 1 x /\ dist R_met x 1 < beta -> dist R_met (atan x) (atan 1) < eps''
a:R
Pa0:0 < a
Pa1:a < 1
P1:R_dist a 1 < alpha
P2:R_dist a 1 < beta

R_dist (Alt_PI / 4) (ps_atan a) < eps'
H:0 < PI / 6
t1:1 < PI / 2
t2:PI <= 4
m:0 < Alt_PI
H0:- PI / 2 < 1 < PI / 2
eps:R
ep:0 < eps
ca:continuity_pt atan 1
eps', eps'':R
eps_ineq:eps' + eps'' <= eps
ep':0 < eps'
ep'':0 < eps''
alpha:R
a0:alpha > 0
Palpha:forall x : R, x < 1 -> 0 < x -> R_dist x 1 < alpha -> dist R_met (ps_atan x) (Alt_PI / 4) < eps'
beta:R
b0:beta > 0
Pbeta:forall x : Base R_met, D_x no_cond 1 x /\ dist R_met x 1 < beta -> dist R_met (atan x) (atan 1) < eps''
a:R
Pa0:0 < a
Pa1:a < 1
P1:R_dist a 1 < alpha
P2:R_dist a 1 < beta
R_dist (ps_atan a) (PI / 4) < eps''
H:0 < PI / 6
t1:1 < PI / 2
t2:PI <= 4
m:0 < Alt_PI
H0:- PI / 2 < 1 < PI / 2
eps:R
ep:0 < eps
ca:continuity_pt atan 1
eps', eps'':R
eps_ineq:eps' + eps'' <= eps
ep':0 < eps'
ep'':0 < eps''
alpha:R
a0:alpha > 0
Palpha:forall x : R, x < 1 -> 0 < x -> R_dist x 1 < alpha -> dist R_met (ps_atan x) (Alt_PI / 4) < eps'
beta:R
b0:beta > 0
Pbeta:forall x : Base R_met, D_x no_cond 1 x /\ dist R_met x 1 < beta -> dist R_met (atan x) (atan 1) < eps''
a:R
Pa0:0 < a
Pa1:a < 1
P1:R_dist a 1 < alpha
P2:R_dist a 1 < beta

R_dist (ps_atan a) (PI / 4) < eps''
H:0 < PI / 6
t1:1 < PI / 2
t2:PI <= 4
m:0 < Alt_PI
H0:- PI / 2 < 1 < PI / 2
eps:R
ep:0 < eps
ca:continuity_pt atan 1
eps', eps'':R
eps_ineq:eps' + eps'' <= eps
ep':0 < eps'
ep'':0 < eps''
alpha:R
a0:alpha > 0
Palpha:forall x : R, x < 1 -> 0 < x -> R_dist x 1 < alpha -> dist R_met (ps_atan x) (Alt_PI / 4) < eps'
beta:R
b0:beta > 0
Pbeta:forall x : Base R_met, D_x no_cond 1 x /\ dist R_met x 1 < beta -> dist R_met (atan x) (atan 1) < eps''
a:R
Pa0:0 < a
Pa1:a < 1
P1:R_dist a 1 < alpha
P2:R_dist a 1 < beta

R_dist (atan a) (PI / 4) < eps''
H:0 < PI / 6
t1:1 < PI / 2
t2:PI <= 4
m:0 < Alt_PI
H0:- PI / 2 < 1 < PI / 2
eps:R
ep:0 < eps
ca:continuity_pt atan 1
eps', eps'':R
eps_ineq:eps' + eps'' <= eps
ep':0 < eps'
ep'':0 < eps''
alpha:R
a0:alpha > 0
Palpha:forall x : R, x < 1 -> 0 < x -> R_dist x 1 < alpha -> dist R_met (ps_atan x) (Alt_PI / 4) < eps'
beta:R
b0:beta > 0
Pbeta:forall x : Base R_met, D_x no_cond 1 x /\ dist R_met x 1 < beta -> dist R_met (atan x) (atan 1) < eps''
a:R
Pa0:0 < a
Pa1:a < 1
P1:R_dist a 1 < alpha
P2:R_dist a 1 < beta
0 < a < 1
H:0 < PI / 6
t1:1 < PI / 2
t2:PI <= 4
m:0 < Alt_PI
H0:- PI / 2 < 1 < PI / 2
eps:R
ep:0 < eps
ca:continuity_pt atan 1
eps', eps'':R
eps_ineq:eps' + eps'' <= eps
ep':0 < eps'
ep'':0 < eps''
alpha:R
a0:alpha > 0
Palpha:forall x : R, x < 1 -> 0 < x -> R_dist x 1 < alpha -> dist R_met (ps_atan x) (Alt_PI / 4) < eps'
beta:R
b0:beta > 0
Pbeta:forall x : Base R_met, D_x no_cond 1 x /\ dist R_met x 1 < beta -> dist R_met (atan x) (atan 1) < eps''
a:R
Pa0:0 < a
Pa1:a < 1
P1:R_dist a 1 < alpha
P2:R_dist a 1 < beta

D_x no_cond 1 a /\ dist R_met a 1 < beta
H:0 < PI / 6
t1:1 < PI / 2
t2:PI <= 4
m:0 < Alt_PI
H0:- PI / 2 < 1 < PI / 2
eps:R
ep:0 < eps
ca:continuity_pt atan 1
eps', eps'':R
eps_ineq:eps' + eps'' <= eps
ep':0 < eps'
ep'':0 < eps''
alpha:R
a0:alpha > 0
Palpha:forall x : R, x < 1 -> 0 < x -> R_dist x 1 < alpha -> dist R_met (ps_atan x) (Alt_PI / 4) < eps'
beta:R
b0:beta > 0
Pbeta:forall x : Base R_met, D_x no_cond 1 x /\ dist R_met x 1 < beta -> dist R_met (atan x) (atan 1) < eps''
a:R
Pa0:0 < a
Pa1:a < 1
P1:R_dist a 1 < alpha
P2:R_dist a 1 < beta
0 < a < 1
H:0 < PI / 6
t1:1 < PI / 2
t2:PI <= 4
m:0 < Alt_PI
H0:- PI / 2 < 1 < PI / 2
eps:R
ep:0 < eps
ca:continuity_pt atan 1
eps', eps'':R
eps_ineq:eps' + eps'' <= eps
ep':0 < eps'
ep'':0 < eps''
alpha:R
a0:alpha > 0
Palpha:forall x : R, x < 1 -> 0 < x -> R_dist x 1 < alpha -> dist R_met (ps_atan x) (Alt_PI / 4) < eps'
beta:R
b0:beta > 0
Pbeta:forall x : Base R_met, D_x no_cond 1 x /\ dist R_met x 1 < beta -> dist R_met (atan x) (atan 1) < eps''
a:R
Pa0:0 < a
Pa1:a < 1
P1:R_dist a 1 < alpha
P2:R_dist a 1 < beta

D_x no_cond 1 a
H:0 < PI / 6
t1:1 < PI / 2
t2:PI <= 4
m:0 < Alt_PI
H0:- PI / 2 < 1 < PI / 2
eps:R
ep:0 < eps
ca:continuity_pt atan 1
eps', eps'':R
eps_ineq:eps' + eps'' <= eps
ep':0 < eps'
ep'':0 < eps''
alpha:R
a0:alpha > 0
Palpha:forall x : R, x < 1 -> 0 < x -> R_dist x 1 < alpha -> dist R_met (ps_atan x) (Alt_PI / 4) < eps'
beta:R
b0:beta > 0
Pbeta:forall x : Base R_met, D_x no_cond 1 x /\ dist R_met x 1 < beta -> dist R_met (atan x) (atan 1) < eps''
a:R
Pa0:0 < a
Pa1:a < 1
P1:R_dist a 1 < alpha
P2:R_dist a 1 < beta
0 < a < 1
H:0 < PI / 6
t1:1 < PI / 2
t2:PI <= 4
m:0 < Alt_PI
H0:- PI / 2 < 1 < PI / 2
eps:R
ep:0 < eps
ca:continuity_pt atan 1
eps', eps'':R
eps_ineq:eps' + eps'' <= eps
ep':0 < eps'
ep'':0 < eps''
alpha:R
a0:alpha > 0
Palpha:forall x : R, x < 1 -> 0 < x -> R_dist x 1 < alpha -> dist R_met (ps_atan x) (Alt_PI / 4) < eps'
beta:R
b0:beta > 0
Pbeta:forall x : Base R_met, D_x no_cond 1 x /\ dist R_met x 1 < beta -> dist R_met (atan x) (atan 1) < eps''
a:R
Pa0:0 < a
Pa1:a < 1
P1:R_dist a 1 < alpha
P2:R_dist a 1 < beta

0 < a < 1
split; assumption. Qed.

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

forall N : nat, sum_f_R0 (tg_alt PI_tg) (S (2 * N)) <= PI / 4 <= sum_f_R0 (tg_alt PI_tg) (2 * N)
intros; rewrite <- Alt_PI_eq; apply Alt_PI_ineq. Qed.