:: FRECHET2 semantic presentation

Lemma1: for b1 being non empty TopSpace holds
( ( for b2 being Point of b1 holds Cl {b2} = {b2} ) implies b1 is_T1 )
proof end;

Lemma2: for b1 being non empty TopSpace holds
not ( not b1 is_T1 & ( for b2, b3 being Point of b1 holds
not ( b2 <> b3 & b3 in Cl {b2} ) ) )
proof end;

Lemma3: for b1 being non empty TopSpace holds
not ( not b1 is_T1 & ( for b2, b3 being Point of b1
for b4 being sequence of b1 holds
not ( b4 = NAT --> b2 & b2 <> b3 & b4 is_convergent_to b3 ) ) )
proof end;

Lemma4: for b1 being non empty TopSpace holds
( b1 is_T2 implies b1 is_T1 )
proof end;

Lemma5: for b1 being non empty 1-sorted
for b2 being sequence of b1
for b3 being Function of NAT , NAT holds
b2 * b3 is sequence of b1
;

definition
let c1 be non empty 1-sorted ;
let c2 be Function of NAT , NAT ;
let c3 be sequence of c1;
redefine func * as c3 * c2 -> sequence of a1;
coherence
c2 * c3 is sequence of c1
by Lemma5;
end;

theorem Th1: :: FRECHET2:1
for b1 being non empty 1-sorted
for b2 being sequence of b1
for b3 being increasing Seq_of_Nat holds
b2 * b3 is sequence of b1
proof end;

Lemma7: id NAT is Real_Sequence
proof end;

Lemma8: for b1 being Real_Sequence holds
( b1 = id NAT implies b1 is natural-yielding )
proof end;

Lemma9: for b1 being Real_Sequence holds
( b1 = id NAT implies b1 is increasing )
proof end;

theorem Th2: :: FRECHET2:2
for b1 being Real_Sequence holds
( b1 = id NAT implies b1 is increasing Seq_of_Nat ) by Lemma8, Lemma9;

Lemma10: for b1 being non empty 1-sorted
for b2 being sequence of b1 holds
ex b3 being increasing Seq_of_Nat st b2 = b2 * b3
proof end;

theorem Th3: :: FRECHET2:3
for b1 being non empty 1-sorted
for b2 being sequence of b1 holds
b2 is subsequence of b2
proof end;

theorem Th4: :: FRECHET2:4
for b1 being non empty 1-sorted
for b2 being sequence of b1
for b3 being subsequence of b2 holds rng b3 c= rng b2
proof end;

Lemma13: for b1 being non empty 1-sorted
for b2 being sequence of b1
for b3 being increasing Seq_of_Nat holds
b2 * b3 is subsequence of b2
proof end;

definition
let c1 be non empty 1-sorted ;
let c2 be increasing Seq_of_Nat;
let c3 be sequence of c1;
redefine func * as c3 * c2 -> subsequence of a3;
correctness
coherence
c2 * c3 is subsequence of c3
;
by Lemma13;
end;

theorem Th5: :: FRECHET2:5
for b1 being non empty 1-sorted
for b2 being sequence of b1
for b3 being subsequence of b2
for b4 being subsequence of b3 holds
b4 is subsequence of b2
proof end;

scheme :: FRECHET2:sch 1
s1{ F1() -> non empty 1-sorted , F2() -> sequence of F1(), P1[ set ] } :
ex b1 being subsequence of F2() st
for b2 being Nat holds P1[b1 . b2]
provided
E15: for b1 being Nat holds
ex b2 being Natex b3 being Point of F1() st
( b1 <= b2 & b3 = F2() . b2 & P1[b3] )
proof end;

scheme :: FRECHET2:sch 2
s2{ F1() -> non empty TopStruct , F2() -> sequence of F1(), P1[ set ] } :
ex b1 being subsequence of F2() st
for b2 being Nat holds P1[b1 . b2]
provided
E15: for b1 being Nat holds
ex b2 being Natex b3 being Point of F1() st
( b1 <= b2 & b3 = F2() . b2 & P1[b3] )
proof end;

theorem Th6: :: FRECHET2:6
for b1 being non empty 1-sorted
for b2 being sequence of b1
for b3 being Subset of b1 holds
not ( ( for b4 being subsequence of b2 holds
not rng b4 c= b3 ) & ( for b4 being Nat holds
ex b5 being Nat st
( b4 <= b5 & b2 . b5 in b3 ) ) )
proof end;

theorem Th7: :: FRECHET2:7
for b1 being non empty 1-sorted
for b2 being sequence of b1
for b3, b4 being Subset of b1 holds
not ( rng b2 c= b3 \/ b4 & ( for b5 being subsequence of b2 holds
( not rng b5 c= b3 & not rng b5 c= b4 ) ) )
proof end;

Lemma17: for b1 being non empty TopSpace holds
( b1 is first-countable implies for b2 being Point of b1 holds
ex b3 being Basis of b2ex b4 being Function st
( dom b4 = NAT & rng b4 = b3 & ( for b5, b6 being Nat holds
( b6 >= b5 implies b4 . b6 c= b4 . b5 ) ) ) )
proof end;

theorem Th8: :: FRECHET2:8
for b1 being non empty TopSpace holds
( ( for b2 being sequence of b1
for b3, b4 being Point of b1 holds
( b3 in Lim b2 & b4 in Lim b2 implies b3 = b4 ) ) implies b1 is_T1 )
proof end;

theorem Th9: :: FRECHET2:9
for b1 being non empty TopSpace holds
( b1 is_T2 implies for b2 being sequence of b1
for b3, b4 being Point of b1 holds
( b3 in Lim b2 & b4 in Lim b2 implies b3 = b4 ) )
proof end;

theorem Th10: :: FRECHET2:10
for b1 being non empty TopSpace holds
( b1 is first-countable implies ( b1 is_T2 iff for b2 being sequence of b1
for b3, b4 being Point of b1 holds
( b3 in Lim b2 & b4 in Lim b2 implies b3 = b4 ) ) )
proof end;

theorem Th11: :: FRECHET2:11
for b1 being non empty TopStruct
for b2 being sequence of b1 holds
( not b2 is convergent implies Lim b2 = {} )
proof end;

theorem Th12: :: FRECHET2:12
for b1 being non empty TopSpace
for b2 being Subset of b1 holds
( b2 is closed implies for b3 being sequence of b1 holds
( rng b3 c= b2 implies Lim b3 c= b2 ) )
proof end;

theorem Th13: :: FRECHET2:13
for b1 being non empty TopStruct
for b2 being sequence of b1
for b3 being Point of b1 holds
not ( not b2 is_convergent_to b3 & ( for b4 being subsequence of b2 holds
ex b5 being subsequence of b4 st b5 is_convergent_to b3 ) )
proof end;

theorem Th14: :: FRECHET2:14
for b1, b2 being non empty TopSpace
for b3 being Function of b1,b2 holds
( b3 is continuous implies for b4 being sequence of b1
for b5 being sequence of b2 holds
( b5 = b3 * b4 implies b3 .: (Lim b4) c= Lim b5 ) )
proof end;

theorem Th15: :: FRECHET2:15
for b1, b2 being non empty TopSpace
for b3 being Function of b1,b2 holds
( b1 is sequential implies ( b3 is continuous iff for b4 being sequence of b1
for b5 being sequence of b2 holds
( b5 = b3 * b4 implies b3 .: (Lim b4) c= Lim b5 ) ) )
proof end;

definition
let c1 be non empty TopStruct ;
let c2 be Subset of c1;
canceled;
func Cl_Seq c2 -> Subset of a1 means :Def2: :: FRECHET2:def 2
for b1 being Point of a1 holds
( b1 in a3 iff ex b2 being sequence of a1 st
( rng b2 c= a2 & b1 in Lim b2 ) );
existence
ex b1 being Subset of c1 st
for b2 being Point of c1 holds
( b2 in b1 iff ex b3 being sequence of c1 st
( rng b3 c= c2 & b2 in Lim b3 ) )
proof end;
uniqueness
for b1, b2 being Subset of c1 holds
( ( for b3 being Point of c1 holds
( b3 in b1 iff ex b4 being sequence of c1 st
( rng b4 c= c2 & b3 in Lim b4 ) ) ) & ( for b3 being Point of c1 holds
( b3 in b2 iff ex b4 being sequence of c1 st
( rng b4 c= c2 & b3 in Lim b4 ) ) ) implies b1 = b2 )
proof end;
end;

:: deftheorem Def1 FRECHET2:def 1 :
canceled;

:: deftheorem Def2 defines Cl_Seq FRECHET2:def 2 :
for b1 being non empty TopStruct
for b2, b3 being Subset of b1 holds
( b3 = Cl_Seq b2 iff for b4 being Point of b1 holds
( b4 in b3 iff ex b5 being sequence of b1 st
( rng b5 c= b2 & b4 in Lim b5 ) ) );

theorem Th16: :: FRECHET2:16
for b1 being non empty TopStruct
for b2 being Subset of b1
for b3 being sequence of b1
for b4 being Point of b1 holds
( rng b3 c= b2 & b4 in Lim b3 implies b4 in Cl b2 )
proof end;

theorem Th17: :: FRECHET2:17
for b1 being non empty TopStruct
for b2 being Subset of b1 holds Cl_Seq b2 c= Cl b2
proof end;

theorem Th18: :: FRECHET2:18
for b1 being non empty TopStruct
for b2 being sequence of b1
for b3 being subsequence of b2
for b4 being Point of b1 holds
( b2 is_convergent_to b4 implies b3 is_convergent_to b4 )
proof end;

theorem Th19: :: FRECHET2:19
for b1 being non empty TopStruct
for b2 being sequence of b1
for b3 being subsequence of b2 holds Lim b2 c= Lim b3
proof end;

theorem Th20: :: FRECHET2:20
for b1 being non empty TopStruct holds Cl_Seq ({} b1) = {}
proof end;

theorem Th21: :: FRECHET2:21
for b1 being non empty TopStruct
for b2 being Subset of b1 holds b2 c= Cl_Seq b2
proof end;

theorem Th22: :: FRECHET2:22
for b1 being non empty TopStruct
for b2, b3 being Subset of b1 holds (Cl_Seq b2) \/ (Cl_Seq b3) = Cl_Seq (b2 \/ b3)
proof end;

theorem Th23: :: FRECHET2:23
for b1 being non empty TopStruct holds
( b1 is Frechet iff for b2 being Subset of b1 holds Cl b2 = Cl_Seq b2 )
proof end;

theorem Th24: :: FRECHET2:24
for b1 being non empty TopSpace holds
( b1 is Frechet implies for b2, b3 being Subset of b1 holds
( Cl_Seq ({} b1) = {} & b2 c= Cl_Seq b2 & Cl_Seq (b2 \/ b3) = (Cl_Seq b2) \/ (Cl_Seq b3) & Cl_Seq (Cl_Seq b2) = Cl_Seq b2 ) )
proof end;

theorem Th25: :: FRECHET2:25
for b1 being non empty TopSpace holds
( b1 is sequential & ( for b2 being Subset of b1 holds Cl_Seq (Cl_Seq b2) = Cl_Seq b2 ) implies b1 is Frechet )
proof end;

theorem Th26: :: FRECHET2:26
for b1 being non empty TopSpace holds
( b1 is sequential implies ( b1 is Frechet iff for b2, b3 being Subset of b1 holds
( Cl_Seq ({} b1) = {} & b2 c= Cl_Seq b2 & Cl_Seq (b2 \/ b3) = (Cl_Seq b2) \/ (Cl_Seq b3) & Cl_Seq (Cl_Seq b2) = Cl_Seq b2 ) ) ) by Th24, Th25;

definition
let c1 be non empty TopSpace;
let c2 be sequence of c1;
assume E33: ex b1 being Point of c1 st Lim c2 = {b1} ;
func lim c2 -> Point of a1 means :Def3: :: FRECHET2:def 3
a2 is_convergent_to a3;
existence
ex b1 being Point of c1 st c2 is_convergent_to b1
proof end;
uniqueness
for b1, b2 being Point of c1 holds
( c2 is_convergent_to b1 & c2 is_convergent_to b2 implies b1 = b2 )
proof end;
end;

:: deftheorem Def3 defines lim FRECHET2:def 3 :
for b1 being non empty TopSpace
for b2 being sequence of b1 holds
( ex b3 being Point of b1 st Lim b2 = {b3} implies for b3 being Point of b1 holds
( b3 = lim b2 iff b2 is_convergent_to b3 ) );

theorem Th27: :: FRECHET2:27
for b1 being non empty TopSpace holds
( b1 is_T2 implies for b2 being sequence of b1 holds
not ( b2 is convergent & ( for b3 being Point of b1 holds
not Lim b2 = {b3} ) ) )
proof end;

theorem Th28: :: FRECHET2:28
for b1 being non empty TopSpace holds
( b1 is_T2 implies for b2 being sequence of b1
for b3 being Point of b1 holds
( b2 is_convergent_to b3 iff ( b2 is convergent & b3 = lim b2 ) ) )
proof end;

theorem Th29: :: FRECHET2:29
for b1 being MetrStruct
for b2 being sequence of b1 holds
b2 is sequence of (TopSpaceMetr b1)
proof end;

theorem Th30: :: FRECHET2:30
for b1 being non empty MetrStruct
for b2 being sequence of (TopSpaceMetr b1) holds
b2 is sequence of b1
proof end;

theorem Th31: :: FRECHET2:31
for b1 being non empty MetrSpace
for b2 being sequence of b1
for b3 being Point of b1
for b4 being sequence of (TopSpaceMetr b1)
for b5 being Point of (TopSpaceMetr b1) holds
( b2 = b4 & b3 = b5 implies ( b2 is_convergent_in_metrspace_to b3 iff b4 is_convergent_to b5 ) )
proof end;

theorem Th32: :: FRECHET2:32
for b1 being non empty MetrSpace
for b2 being sequence of b1
for b3 being sequence of (TopSpaceMetr b1) holds
( b2 = b3 implies ( b2 is convergent iff b3 is convergent ) )
proof end;

theorem Th33: :: FRECHET2:33
for b1 being non empty MetrSpace
for b2 being sequence of b1
for b3 being sequence of (TopSpaceMetr b1) holds
( b2 = b3 & b2 is convergent implies lim b2 = lim b3 )
proof end;

definition
let c1 be TopStruct ;
let c2 be sequence of c1;
let c3 be Point of c1;
pred c3 is_a_cluster_point_of c2 means :Def4: :: FRECHET2:def 4
for b1 being Subset of a1
for b2 being Nat holds
not ( b1 is open & a3 in b1 & ( for b3 being Nat holds
not ( b2 <= b3 & a2 . b3 in b1 ) ) );
end;

:: deftheorem Def4 defines is_a_cluster_point_of FRECHET2:def 4 :
for b1 being TopStruct
for b2 being sequence of b1
for b3 being Point of b1 holds
( b3 is_a_cluster_point_of b2 iff for b4 being Subset of b1
for b5 being Nat holds
not ( b4 is open & b3 in b4 & ( for b6 being Nat holds
not ( b5 <= b6 & b2 . b6 in b4 ) ) ) );

theorem Th34: :: FRECHET2:34
for b1 being non empty TopStruct
for b2 being sequence of b1
for b3 being Point of b1 holds
( ex b4 being subsequence of b2 st b4 is_convergent_to b3 implies b3 is_a_cluster_point_of b2 )
proof end;

theorem Th35: :: FRECHET2:35
for b1 being non empty TopStruct
for b2 being sequence of b1
for b3 being Point of b1 holds
( b2 is_convergent_to b3 implies b3 is_a_cluster_point_of b2 )
proof end;

theorem Th36: :: FRECHET2:36
for b1 being non empty TopStruct
for b2 being sequence of b1
for b3 being Point of b1
for b4 being Subset of b1 holds
( b4 = { b5 where B is Point of b1 : b3 in Cl {b5} } & rng b2 c= b4 implies b2 is_convergent_to b3 )
proof end;

theorem Th37: :: FRECHET2:37
for b1 being non empty TopStruct
for b2 being sequence of b1
for b3, b4 being Point of b1 holds
( ( for b5 being Nat holds
( b2 . b5 = b4 & b2 is_convergent_to b3 ) ) implies b3 in Cl {b4} )
proof end;

theorem Th38: :: FRECHET2:38
for b1 being non empty TopStruct
for b2 being Point of b1
for b3 being Subset of b1
for b4 being sequence of b1 holds
not ( b3 = { b5 where B is Point of b1 : b2 in Cl {b5} } & rng b4 misses b3 & b4 is_convergent_to b2 & ( for b5 being subsequence of b4 holds
not b5 is one-to-one ) )
proof end;

Lemma42: for b1 being Function holds
not ( not dom b1 is finite & b1 is one-to-one & rng b1 is finite )
proof end;

theorem Th39: :: FRECHET2:39
for b1 being non empty TopStruct
for b2, b3 being sequence of b1 holds
not ( rng b3 c= rng b2 & b3 is one-to-one & ( for b4 being Permutation of NAT holds
not b3 * b4 is subsequence of b2 ) )
proof end;

scheme :: FRECHET2:sch 3
s3{ F1() -> non empty 1-sorted , F2() -> sequence of F1(), F3() -> Permutation of NAT , P1[ set ] } :
ex b1 being Nat st
for b2 being Nat holds
( b1 <= b2 implies P1[(F2() * F3()) . b2] )
provided
E44: ex b1 being Nat st
for b2 being Nat
for b3 being Point of F1() holds
( b1 <= b2 & b3 = F2() . b2 implies P1[b3] )
proof end;

scheme :: FRECHET2:sch 4
s4{ F1() -> non empty TopStruct , F2() -> sequence of F1(), F3() -> Permutation of NAT , P1[ set ] } :
ex b1 being Nat st
for b2 being Nat holds
( b1 <= b2 implies P1[(F2() * F3()) . b2] )
provided
E44: ex b1 being Nat st
for b2 being Nat
for b3 being Point of F1() holds
( b1 <= b2 & b3 = F2() . b2 implies P1[b3] )
proof end;

theorem Th40: :: FRECHET2:40
for b1 being non empty TopStruct
for b2 being sequence of b1
for b3 being Permutation of NAT
for b4 being Point of b1 holds
( b2 is_convergent_to b4 implies b2 * b3 is_convergent_to b4 )
proof end;

theorem Th41: :: FRECHET2:41
for b1 being Nat holds
ex b2 being increasing Seq_of_Nat st
for b3 being Nat holds b2 . b3 = b3 + b1
proof end;

theorem Th42: :: FRECHET2:42
for b1 being non empty 1-sorted
for b2 being sequence of b1
for b3 being Nat holds
ex b4 being subsequence of b2 st
for b5 being Nat holds b4 . b5 = b2 . (b5 + b3)
proof end;

theorem Th43: :: FRECHET2:43
for b1 being non empty TopStruct
for b2 being sequence of b1
for b3 being Point of b1
for b4 being subsequence of b2 holds
( b3 is_a_cluster_point_of b2 & ex b5 being Nat st
for b6 being Nat holds b4 . b6 = b2 . (b6 + b5) implies b3 is_a_cluster_point_of b4 )
proof end;

theorem Th44: :: FRECHET2:44
for b1 being non empty TopStruct
for b2 being sequence of b1
for b3 being Point of b1 holds
( b3 is_a_cluster_point_of b2 implies b3 in Cl (rng b2) )
proof end;

theorem Th45: :: FRECHET2:45
for b1 being non empty TopStruct holds
( b1 is Frechet implies for b2 being sequence of b1
for b3 being Point of b1 holds
not ( b3 is_a_cluster_point_of b2 & ( for b4 being subsequence of b2 holds
not b4 is_convergent_to b3 ) ) )
proof end;

theorem Th46: :: FRECHET2:46
for b1 being non empty TopSpace holds
( b1 is first-countable implies for b2 being Point of b1 holds
ex b3 being Basis of b2ex b4 being Function st
( dom b4 = NAT & rng b4 = b3 & ( for b5, b6 being Nat holds
( b6 >= b5 implies b4 . b6 c= b4 . b5 ) ) ) ) by Lemma17;

theorem Th47: :: FRECHET2:47
for b1 being non empty TopSpace holds
( ( for b2 being Point of b1 holds Cl {b2} = {b2} ) implies b1 is_T1 ) by Lemma1;

theorem Th48: :: FRECHET2:48
for b1 being non empty TopSpace holds
( b1 is_T2 implies b1 is_T1 ) by Lemma4;

theorem Th49: :: FRECHET2:49
for b1 being non empty TopSpace holds
not ( not b1 is_T1 & ( for b2, b3 being Point of b1
for b4 being sequence of b1 holds
not ( b4 = NAT --> b2 & b2 <> b3 & b4 is_convergent_to b3 ) ) ) by Lemma3;

theorem Th50: :: FRECHET2:50
for b1 being Function holds
not ( not dom b1 is finite & b1 is one-to-one & rng b1 is finite ) by Lemma42;