begin
Lm1:
1 / 1 = 1
;
theorem Th1:
theorem Th2:
theorem Th3:
:: deftheorem Def1 defines Partial_Sums SERIES_1:def 1 :
for X being complex-membered set
for s being sequence of X st not X is empty & X is add-closed holds
for b3 being sequence of X holds
( b3 = Partial_Sums s iff ( b3 . 0 = s . 0 & ( for n being Element of NAT holds b3 . (n + 1) = (b3 . n) + (s . (n + 1)) ) ) );
:: deftheorem Def2 defines summable SERIES_1:def 2 :
for s being Real_Sequence holds
( s is summable iff Partial_Sums s is convergent );
:: deftheorem defines Sum SERIES_1:def 3 :
for s being Real_Sequence holds Sum s = lim (Partial_Sums s);
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem Th7:
Lm2:
for X being non empty complex-membered add-closed set
for seq, seq1, seq2 being sequence of X holds
( seq = seq1 + seq2 iff for n being Element of NAT holds seq . n = (seq1 . n) + (seq2 . n) )
theorem Th8:
theorem Th9:
theorem
theorem
theorem Th12:
theorem Th13:
theorem Th14:
theorem Th15:
theorem Th16:
theorem Th17:
theorem
theorem Th19:
theorem Th20:
theorem
theorem Th22:
theorem
canceled;
theorem
theorem Th25:
theorem Th26:
theorem
theorem Th28:
theorem
theorem Th30:
theorem
theorem Th32:
theorem Th33:
theorem
theorem Th35:
theorem
theorem
:: deftheorem SERIES_1:def 4 :
canceled;
:: deftheorem Def5 defines absolutely_summable SERIES_1:def 5 :
for s being Real_Sequence holds
( s is absolutely_summable iff abs s is summable );
theorem
canceled;
theorem Th39:
theorem
theorem
theorem
theorem Th43:
theorem
theorem
theorem
theorem
begin
:: deftheorem defines Sum SERIES_1:def 6 :
for s being Real_Sequence
for n being Nat holds Sum (s,n) = (Partial_Sums s) . n;
:: deftheorem defines Sum SERIES_1:def 7 :
for s being Real_Sequence
for n, m being Nat holds Sum (s,n,m) = (Sum (s,n)) - (Sum (s,m));