begin
:: deftheorem Def1 defines || MESFUN9C:def 1 :
for X, Y being set
for F being Functional_Sequence of X,Y
for D being set
for b5 being Functional_Sequence of X,Y holds
( b5 = F || D iff for n being Nat holds b5 . n = (F . n) | D );
theorem Th1:
theorem Th2:
theorem Th3:
theorem Th4:
theorem Th5:
theorem Th6:
:: deftheorem Def2 defines Partial_Sums MESFUN9C:def 2 :
for X being non empty set
for F, b3 being Functional_Sequence of X,REAL holds
( b3 = Partial_Sums F iff ( b3 . 0 = F . 0 & ( for n being Nat holds b3 . (n + 1) = (b3 . n) + (F . (n + 1)) ) ) );
theorem Th7:
theorem Th8:
theorem Th9:
theorem Th10:
theorem Th11:
theorem Th12:
theorem Th13:
theorem Th14:
theorem Th15:
theorem Th16:
theorem Th17:
theorem Th18:
theorem Th19:
begin
theorem Th20:
Lm1:
for X being non empty set
for D being set
for F being Functional_Sequence of X,COMPLEX
for n being Nat holds
( (Re (F || D)) . n = ((Re F) . n) | D & (Im (F || D)) . n = ((Im F) . n) | D )
theorem Th21:
theorem Th22:
theorem Th23:
theorem Th24:
theorem Th25:
theorem
theorem
theorem
:: deftheorem Def3 defines Partial_Sums MESFUN9C:def 3 :
for X being non empty set
for F, b3 being Functional_Sequence of X,COMPLEX holds
( b3 = Partial_Sums F iff ( b3 . 0 = F . 0 & ( for n being Nat holds b3 . (n + 1) = (b3 . n) + (F . (n + 1)) ) ) );
theorem Th29:
theorem
theorem
theorem Th32:
theorem
theorem Th34:
theorem Th35:
theorem
theorem
Lm2:
for X being non empty set
for S being SigmaField of X
for E being Element of S
for m being Nat
for F being Functional_Sequence of X,COMPLEX st ( for n being Nat holds F . n is_measurable_on E ) holds
( (Re F) . m is_measurable_on E & (Im F) . m is_measurable_on E )
theorem
theorem
theorem
begin
theorem
theorem
theorem
theorem
theorem
begin
theorem Th46:
theorem Th47:
theorem Th48:
:: deftheorem Def4 defines uniformly_bounded MESFUN9C:def 4 :
for X being set
for F being Functional_Sequence of X,REAL holds
( F is uniformly_bounded iff ex K being real number st
for n being Nat
for x being Element of X st x in dom (F . 0) holds
|.((F . n) . x).| <= K );
theorem Th49:
:: deftheorem Def5 defines is_uniformly_convergent_to MESFUN9C:def 5 :
for X being set
for F being Functional_Sequence of X,REAL
for f being PartFunc of X,REAL holds
( F is_uniformly_convergent_to f iff ( F is with_the_same_dom & dom (F . 0) = dom f & ( for e being real number st e > 0 holds
ex N being Nat st
for n being Nat
for x being Element of X st n >= N & x in dom (F . 0) holds
|.(((F . n) . x) - (f . x)).| < e ) ) );
theorem Th50:
theorem Th51:
theorem
:: deftheorem Def6 defines uniformly_bounded MESFUN9C:def 6 :
for X being set
for F being Functional_Sequence of X,COMPLEX holds
( F is uniformly_bounded iff ex K being real number st
for n being Nat
for x being Element of X st x in dom (F . 0) holds
|.((F . n) . x).| <= K );
theorem
:: deftheorem Def7 defines is_uniformly_convergent_to MESFUN9C:def 7 :
for X being set
for F being Functional_Sequence of X,COMPLEX
for f being PartFunc of X,COMPLEX holds
( F is_uniformly_convergent_to f iff ( F is with_the_same_dom & dom (F . 0) = dom f & ( for e being real number st e > 0 holds
ex N being Nat st
for n being Nat
for x being Element of X st n >= N & x in dom (F . 0) holds
|.(((F . n) . x) - (f . x)).| < e ) ) );
theorem