:: RSSPACE3 semantic presentation

definition
func the_set_of_l1RealSequences -> Subset of Linear_Space_of_RealSequences means :Def1: :: RSSPACE3:def 1
for b1 being set holds
( b1 in a1 iff ( b1 in the_set_of_RealSequences & seq_id b1 is absolutely_summable ) );
existence
ex b1 being Subset of Linear_Space_of_RealSequences st
for b2 being set holds
( b2 in b1 iff ( b2 in the_set_of_RealSequences & seq_id b2 is absolutely_summable ) )
proof end;
uniqueness
for b1, b2 being Subset of Linear_Space_of_RealSequences holds
( ( for b3 being set holds
( b3 in b1 iff ( b3 in the_set_of_RealSequences & seq_id b3 is absolutely_summable ) ) ) & ( for b3 being set holds
( b3 in b2 iff ( b3 in the_set_of_RealSequences & seq_id b3 is absolutely_summable ) ) ) implies b1 = b2 )
proof end;
end;

:: deftheorem Def1 defines the_set_of_l1RealSequences RSSPACE3:def 1 :
for b1 being Subset of Linear_Space_of_RealSequences holds
( b1 = the_set_of_l1RealSequences iff for b2 being set holds
( b2 in b1 iff ( b2 in the_set_of_RealSequences & seq_id b2 is absolutely_summable ) ) );

theorem Th1: :: RSSPACE3:1
for b1 being Real
for b2 being Real_Sequence holds
( b2 is convergent implies for b3 being Real_Sequence holds
( ( for b4 being Nat holds b3 . b4 = abs ((b2 . b4) - b1) ) implies ( b3 is convergent & lim b3 = abs ((lim b2) - b1) ) ) )
proof end;

registration
cluster the_set_of_l1RealSequences -> non empty ;
coherence
not the_set_of_l1RealSequences is empty
proof end;
end;

registration
cluster the_set_of_l1RealSequences -> non empty lineary-closed ;
coherence
the_set_of_l1RealSequences is lineary-closed
proof end;
end;

Lemma3: RLSStruct(# the_set_of_l1RealSequences ,(Zero_ the_set_of_l1RealSequences ,Linear_Space_of_RealSequences ),(Add_ the_set_of_l1RealSequences ,Linear_Space_of_RealSequences ),(Mult_ the_set_of_l1RealSequences ,Linear_Space_of_RealSequences ) #) is Subspace of Linear_Space_of_RealSequences
by RSSPACE:13;

registration
cluster RLSStruct(# the_set_of_l1RealSequences ,(Zero_ the_set_of_l1RealSequences ,Linear_Space_of_RealSequences ),(Add_ the_set_of_l1RealSequences ,Linear_Space_of_RealSequences ),(Mult_ the_set_of_l1RealSequences ,Linear_Space_of_RealSequences ) #) -> Abelian add-associative right_zeroed right_complementable RealLinearSpace-like ;
coherence
( RLSStruct(# the_set_of_l1RealSequences ,(Zero_ the_set_of_l1RealSequences ,Linear_Space_of_RealSequences ),(Add_ the_set_of_l1RealSequences ,Linear_Space_of_RealSequences ),(Mult_ the_set_of_l1RealSequences ,Linear_Space_of_RealSequences ) #) is Abelian & RLSStruct(# the_set_of_l1RealSequences ,(Zero_ the_set_of_l1RealSequences ,Linear_Space_of_RealSequences ),(Add_ the_set_of_l1RealSequences ,Linear_Space_of_RealSequences ),(Mult_ the_set_of_l1RealSequences ,Linear_Space_of_RealSequences ) #) is add-associative & RLSStruct(# the_set_of_l1RealSequences ,(Zero_ the_set_of_l1RealSequences ,Linear_Space_of_RealSequences ),(Add_ the_set_of_l1RealSequences ,Linear_Space_of_RealSequences ),(Mult_ the_set_of_l1RealSequences ,Linear_Space_of_RealSequences ) #) is right_zeroed & RLSStruct(# the_set_of_l1RealSequences ,(Zero_ the_set_of_l1RealSequences ,Linear_Space_of_RealSequences ),(Add_ the_set_of_l1RealSequences ,Linear_Space_of_RealSequences ),(Mult_ the_set_of_l1RealSequences ,Linear_Space_of_RealSequences ) #) is right_complementable & RLSStruct(# the_set_of_l1RealSequences ,(Zero_ the_set_of_l1RealSequences ,Linear_Space_of_RealSequences ),(Add_ the_set_of_l1RealSequences ,Linear_Space_of_RealSequences ),(Mult_ the_set_of_l1RealSequences ,Linear_Space_of_RealSequences ) #) is RealLinearSpace-like )
by RSSPACE:13;
end;

Lemma4: RLSStruct(# the_set_of_l1RealSequences ,(Zero_ the_set_of_l1RealSequences ,Linear_Space_of_RealSequences ),(Add_ the_set_of_l1RealSequences ,Linear_Space_of_RealSequences ),(Mult_ the_set_of_l1RealSequences ,Linear_Space_of_RealSequences ) #) is RealLinearSpace
;

Lemma5: ex b1 being Function of the_set_of_l1RealSequences , REAL st
for b2 being set holds
( b2 in the_set_of_l1RealSequences implies b1 . b2 = Sum (abs (seq_id b2)) )
proof end;

definition
func l_norm -> Function of the_set_of_l1RealSequences , REAL means :Def2: :: RSSPACE3:def 2
for b1 being set holds
( b1 in the_set_of_l1RealSequences implies a1 . b1 = Sum (abs (seq_id b1)) );
existence
ex b1 being Function of the_set_of_l1RealSequences , REAL st
for b2 being set holds
( b2 in the_set_of_l1RealSequences implies b1 . b2 = Sum (abs (seq_id b2)) )
by Lemma5;
uniqueness
for b1, b2 being Function of the_set_of_l1RealSequences , REAL holds
( ( for b3 being set holds
( b3 in the_set_of_l1RealSequences implies b1 . b3 = Sum (abs (seq_id b3)) ) ) & ( for b3 being set holds
( b3 in the_set_of_l1RealSequences implies b2 . b3 = Sum (abs (seq_id b3)) ) ) implies b1 = b2 )
proof end;
end;

:: deftheorem Def2 defines l_norm RSSPACE3:def 2 :
for b1 being Function of the_set_of_l1RealSequences , REAL holds
( b1 = l_norm iff for b2 being set holds
( b2 in the_set_of_l1RealSequences implies b1 . b2 = Sum (abs (seq_id b2)) ) );

registration
let c1 be non empty set ;
let c2 be Element of c1;
let c3 be BinOp of c1;
let c4 be Function of [:REAL ,c1:],c1;
let c5 be Function of c1, REAL ;
cluster NORMSTR(# a1,a2,a3,a4,a5 #) -> non empty ;
coherence
not NORMSTR(# c1,c2,c3,c4,c5 #) is empty
by STRUCT_0:def 1;
end;

theorem Th2: :: RSSPACE3:2
canceled;

theorem Th3: :: RSSPACE3:3
canceled;

theorem Th4: :: RSSPACE3:4
for b1 being NORMSTR holds
( RLSStruct(# the carrier of b1,the Zero of b1,the add of b1,the Mult of b1 #) is RealLinearSpace implies b1 is RealLinearSpace )
proof end;

theorem Th5: :: RSSPACE3:5
for b1 being Real_Sequence holds
( ( for b2 being Nat holds b1 . b2 = 0 ) implies ( b1 is absolutely_summable & Sum (abs b1) = 0 ) )
proof end;

theorem Th6: :: RSSPACE3:6
for b1 being Real_Sequence holds
( b1 is absolutely_summable & Sum (abs b1) = 0 implies for b2 being Nat holds b1 . b2 = 0 )
proof end;

theorem Th7: :: RSSPACE3:7
NORMSTR(# the_set_of_l1RealSequences ,(Zero_ the_set_of_l1RealSequences ,Linear_Space_of_RealSequences ),(Add_ the_set_of_l1RealSequences ,Linear_Space_of_RealSequences ),(Mult_ the_set_of_l1RealSequences ,Linear_Space_of_RealSequences ),l_norm #) is RealLinearSpace by Lemma4, Th4;

definition
func l1_Space -> non empty NORMSTR equals :: RSSPACE3:def 3
NORMSTR(# the_set_of_l1RealSequences ,(Zero_ the_set_of_l1RealSequences ,Linear_Space_of_RealSequences ),(Add_ the_set_of_l1RealSequences ,Linear_Space_of_RealSequences ),(Mult_ the_set_of_l1RealSequences ,Linear_Space_of_RealSequences ),l_norm #);
coherence
NORMSTR(# the_set_of_l1RealSequences ,(Zero_ the_set_of_l1RealSequences ,Linear_Space_of_RealSequences ),(Add_ the_set_of_l1RealSequences ,Linear_Space_of_RealSequences ),(Mult_ the_set_of_l1RealSequences ,Linear_Space_of_RealSequences ),l_norm #) is non empty NORMSTR
;
end;

:: deftheorem Def3 defines l1_Space RSSPACE3:def 3 :
l1_Space = NORMSTR(# the_set_of_l1RealSequences ,(Zero_ the_set_of_l1RealSequences ,Linear_Space_of_RealSequences ),(Add_ the_set_of_l1RealSequences ,Linear_Space_of_RealSequences ),(Mult_ the_set_of_l1RealSequences ,Linear_Space_of_RealSequences ),l_norm #);

theorem Th8: :: RSSPACE3:8
( the carrier of l1_Space = the_set_of_l1RealSequences & ( for b1 being set holds
( b1 is VECTOR of l1_Space iff ( b1 is Real_Sequence & seq_id b1 is absolutely_summable ) ) ) & 0. l1_Space = Zeroseq & ( for b1 being VECTOR of l1_Space holds b1 = seq_id b1 ) & ( for b1, b2 being VECTOR of l1_Space holds b1 + b2 = (seq_id b1) + (seq_id b2) ) & ( for b1 being Real
for b2 being VECTOR of l1_Space holds b1 * b2 = b1 (#) (seq_id b2) ) & ( for b1 being VECTOR of l1_Space holds
( - b1 = - (seq_id b1) & seq_id (- b1) = - (seq_id b1) ) ) & ( for b1, b2 being VECTOR of l1_Space holds b1 - b2 = (seq_id b1) - (seq_id b2) ) & ( for b1 being VECTOR of l1_Space holds seq_id b1 is absolutely_summable ) & ( for b1 being VECTOR of l1_Space holds ||.b1.|| = Sum (abs (seq_id b1)) ) )
proof end;

theorem Th9: :: RSSPACE3:9
for b1, b2 being Point of l1_Space
for b3 being Real holds
( ( ||.b1.|| = 0 implies b1 = 0. l1_Space ) & ( b1 = 0. l1_Space implies ||.b1.|| = 0 ) & 0 <= ||.b1.|| & ||.(b1 + b2).|| <= ||.b1.|| + ||.b2.|| & ||.(b3 * b1).|| = (abs b3) * ||.b1.|| )
proof end;

registration
cluster l1_Space -> non empty Abelian add-associative right_zeroed right_complementable RealLinearSpace-like RealNormSpace-like ;
coherence
( l1_Space is RealNormSpace-like & l1_Space is RealLinearSpace-like & l1_Space is Abelian & l1_Space is add-associative & l1_Space is right_zeroed & l1_Space is right_complementable )
by Lemma4, Th4, Th9, NORMSP_1:def 2;
end;

Lemma13: for b1 being Real
for b2, b3 being Real_Sequence holds
( b2 is convergent & b3 is convergent implies for b4 being Real_Sequence holds
( ( for b5 being Nat holds b4 . b5 = (abs ((b2 . b5) - b1)) + (b3 . b5) ) implies ( b4 is convergent & lim b4 = (abs ((lim b2) - b1)) + (lim b3) ) ) )
proof end;

definition
let c1 be non empty NORMSTR ;
let c2, c3 be Point of c1;
func dist c2,c3 -> Real equals :: RSSPACE3:def 4
||.(a2 - a3).||;
coherence
||.(c2 - c3).|| is Real
;
end;

:: deftheorem Def4 defines dist RSSPACE3:def 4 :
for b1 being non empty NORMSTR
for b2, b3 being Point of b1 holds dist b2,b3 = ||.(b2 - b3).||;

definition
let c1 be non empty NORMSTR ;
let c2 be sequence of c1;
attr a2 is CCauchy means :Def5: :: RSSPACE3:def 5
for b1 being Real holds
not ( b1 > 0 & ( for b2 being Nat holds
ex b3, b4 being Nat st
( b3 >= b2 & b4 >= b2 & not dist (a2 . b3),(a2 . b4) < b1 ) ) );
end;

:: deftheorem Def5 defines CCauchy RSSPACE3:def 5 :
for b1 being non empty NORMSTR
for b2 being sequence of b1 holds
( b2 is CCauchy iff for b3 being Real holds
not ( b3 > 0 & ( for b4 being Nat holds
ex b5, b6 being Nat st
( b5 >= b4 & b6 >= b4 & not dist (b2 . b5),(b2 . b6) < b3 ) ) ) );

notation
let c1 be non empty NORMSTR ;
let c2 be sequence of c1;
synonym Cauchy_sequence_by_Norm c2 for CCauchy c2;
end;

theorem Th10: :: RSSPACE3:10
for b1 being non empty RealNormSpace
for b2 being sequence of b1 holds
( b2 is CCauchy iff for b3 being Real holds
not ( b3 > 0 & ( for b4 being Nat holds
ex b5, b6 being Nat st
( b5 >= b4 & b6 >= b4 & not ||.((b2 . b5) - (b2 . b6)).|| < b3 ) ) ) )
proof end;

theorem Th11: :: RSSPACE3:11
for b1 being sequence of l1_Space holds
( b1 is CCauchy implies b1 is convergent )
proof end;