Library Coq.Reals.Rseries
Require Import Rbase.
Require Import Rfunctions.
Require Import Compare.
Local Open Scope R_scope.
Implicit Type r : R.
Section sequence.
Variable Un : nat -> R.
Fixpoint Rmax_N (N:nat) : R :=
match N with
| O => Un 0
| S n => Rmax (Un (S n)) (Rmax_N n)
end.
Definition EUn r : Prop := exists i : nat, r = Un i.
Definition Un_cv (l:R) : Prop :=
forall eps:R,
eps > 0 ->
exists N : nat, (forall n:nat, (n >= N)%nat -> Rdist (Un n) l < eps).
Definition Cauchy_crit : Prop :=
forall eps:R,
eps > 0 ->
exists N : nat,
(forall n m:nat,
(n >= N)%nat -> (m >= N)%nat -> Rdist (Un n) (Un m) < eps).
Definition Un_growing : Prop := forall n:nat, Un n <= Un (S n).
Lemma EUn_noempty : exists r : R, EUn r.
Lemma Un_in_EUn : forall n:nat, EUn (Un n).
Lemma Un_bound_imp :
forall x:R, (forall n:nat, Un n <= x) -> is_upper_bound EUn x.
Lemma growing_prop :
forall n m:nat, Un_growing -> (n >= m)%nat -> Un n >= Un m.
Lemma Un_cv_crit_lub : Un_growing -> forall l, is_lub EUn l -> Un_cv l.
Lemma Un_cv_crit : Un_growing -> bound EUn -> exists l : R, Un_cv l.
Lemma finite_greater :
forall N:nat, exists M : R, (forall n:nat, (n <= N)%nat -> Un n <= M).
Lemma cauchy_bound : Cauchy_crit -> bound EUn.
End sequence.
Section Isequence.
Variable An : nat -> R.
Definition Pser (x l:R) : Prop := infinite_sum (fun n:nat => An n * x ^ n) l.
End Isequence.
Lemma GP_infinite :
forall x:R, Rabs x < 1 -> Pser (fun n:nat => 1) x (/ (1 - x)).
Lemma CV_shift :
forall f k l, Un_cv (fun n => f (n + k)%nat) l -> Un_cv f l.
Lemma CV_shift' :
forall f k l, Un_cv f l -> Un_cv (fun n => f (n + k)%nat) l.
Lemma Un_growing_shift :
forall k un, Un_growing un -> Un_growing (fun n => un (n + k)%nat).