begin
definition
canceled;
end;
:: deftheorem LOPBAN_3:def 1 :
canceled;
theorem Th1:
:: deftheorem Def2 defines summable LOPBAN_3:def 2 :
for X being RealNormSpace
for seq being sequence of X holds
( seq is summable iff Partial_Sums seq is convergent );
:: deftheorem defines Sum LOPBAN_3:def 3 :
for X being RealNormSpace
for seq being sequence of X holds Sum seq = lim (Partial_Sums seq);
:: deftheorem Def4 defines norm_summable LOPBAN_3:def 4 :
for X being RealNormSpace
for seq being sequence of X holds
( seq is norm_summable iff ||.seq.|| is summable );
theorem Th2:
theorem Th3:
theorem Th4:
theorem Th5:
theorem Th6:
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem Th11:
theorem Th12:
theorem
canceled;
theorem
theorem Th15:
theorem Th16:
theorem Th17:
theorem Th18:
theorem Th19:
theorem Th20:
theorem Th21:
theorem Th22:
theorem Th23:
theorem Th24:
theorem Th25:
theorem Th26:
theorem Th27:
theorem Th28:
theorem Th29:
theorem Th30:
theorem Th31:
theorem
theorem
theorem
theorem
theorem
theorem
theorem
theorem
theorem
theorem
theorem
begin
theorem Th43:
:: deftheorem LOPBAN_3:def 5 :
canceled;
:: deftheorem LOPBAN_3:def 6 :
canceled;
:: deftheorem LOPBAN_3:def 7 :
canceled;
:: deftheorem Def8 defines invertible LOPBAN_3:def 8 :
for X being non empty associative well-unital multLoopStr
for v being Element of X holds
( v is invertible iff ex w being Element of X st
( v * w = 1. X & w * v = 1. X ) );
:: deftheorem defines * LOPBAN_3:def 9 :
for X being non empty Normed_AlgebraStr
for S being sequence of X
for a being Element of X
for b4 being sequence of X holds
( b4 = a * S iff for n being Element of NAT holds b4 . n = a * (S . n) );
:: deftheorem defines * LOPBAN_3:def 10 :
for X being non empty Normed_AlgebraStr
for S being sequence of X
for a being Element of X
for b4 being sequence of X holds
( b4 = S * a iff for n being Element of NAT holds b4 . n = (S . n) * a );
:: deftheorem defines * LOPBAN_3:def 11 :
for X being non empty Normed_AlgebraStr
for seq1, seq2, b4 being sequence of X holds
( b4 = seq1 * seq2 iff for n being Element of NAT holds b4 . n = (seq1 . n) * (seq2 . n) );
:: deftheorem Def12 defines " LOPBAN_3:def 12 :
for X being non empty associative well-unital multLoopStr
for x being Element of X st x is invertible holds
for b3 being Element of the carrier of X holds
( b3 = x " iff ( x * b3 = 1. X & b3 * x = 1. X ) );
:: deftheorem Def13 defines GeoSeq LOPBAN_3:def 13 :
for X being Banach_Algebra
for z being Element of X
for b3 being sequence of X holds
( b3 = z GeoSeq iff ( b3 . 0 = 1. X & ( for n being Element of NAT holds b3 . (n + 1) = (b3 . n) * z ) ) );
:: deftheorem defines #N LOPBAN_3:def 14 :
for X being Banach_Algebra
for z being Element of X
for n being Element of NAT holds z #N n = (z GeoSeq) . n;
theorem
theorem Th45:
theorem
Lm1:
for X being RealNormSpace
for x being Point of X st ( for e being Real st e > 0 holds
||.x.|| < e ) holds
x = 0. X
Lm2:
for X being RealNormSpace
for x, y being Point of X st ( for e being Real st e > 0 holds
||.(x - y).|| < e ) holds
x = y
Lm3:
for X being RealNormSpace
for x, y being Point of X
for seq being Real_Sequence st seq is convergent & lim seq = 0 & ( for n being Element of NAT holds ||.(x - y).|| <= seq . n ) holds
x = y
theorem