:: TOPGEN_1 semantic presentation

Lemma1: for b1 being empty TopSpace
for b2 being Subset of b1 holds b2 = {}
;

theorem Th1: :: TOPGEN_1:1
for b1 being 1-sorted
for b2, b3 being Subset of b1 holds
( not ( b2 meets b3 ` & not b2 \ b3 <> {} ) & not ( b2 \ b3 <> {} & not b2 meets b3 ` ) )
proof end;

theorem Th2: :: TOPGEN_1:2
for b1 being 1-sorted holds
( b1 is countable iff [#] b1 is countable ) by ORDERS_4:def 2;

theorem Th3: :: TOPGEN_1:3
for b1 being 1-sorted holds
( b1 is countable iff Card ([#] b1) <=` alef 0 )
proof end;

registration
let c1 be finite 1-sorted ;
cluster [#] a1 -> finite ;
coherence
[#] c1 is finite
;
end;

registration
cluster finite -> countable 1-sorted ;
coherence
for b1 being 1-sorted holds
( b1 is finite implies b1 is countable )
proof end;
end;

registration
cluster non empty countable 1-sorted ;
existence
ex b1 being 1-sorted st
( b1 is countable & not b1 is empty )
proof end;
cluster non empty countable TopStruct ;
existence
ex b1 being TopSpace st
( b1 is countable & not b1 is empty )
proof end;
end;

registration
let c1 be countable 1-sorted ;
cluster [#] a1 -> countable ;
coherence
[#] c1 is countable
by ORDERS_4:def 2;
end;

registration
cluster non empty being_T1 TopStruct ;
existence
ex b1 being TopSpace st
( b1 is being_T1 & not b1 is empty )
proof end;
end;

theorem Th4: :: TOPGEN_1:4
for b1 being TopStruct
for b2 being Subset of b1 holds b2 \/ ([#] b1) = [#] b1 by XBOOLE_1:12;

theorem Th5: :: TOPGEN_1:5
for b1 being TopSpace
for b2 being Subset of b1 holds Int b2 = (Cl (b2 ` )) `
proof end;

definition
let c1 be TopSpace;
let c2 be Subset-Family of c1;
func Fr c2 -> Subset-Family of a1 means :Def1: :: TOPGEN_1:def 1
for b1 being Subset of a1 holds
( b1 in a3 iff ex b2 being Subset of a1 st
( b1 = Fr b2 & b2 in a2 ) );
existence
ex b1 being Subset-Family of c1 st
for b2 being Subset of c1 holds
( b2 in b1 iff ex b3 being Subset of c1 st
( b2 = Fr b3 & b3 in c2 ) )
proof end;
uniqueness
for b1, b2 being Subset-Family of c1 holds
( ( for b3 being Subset of c1 holds
( b3 in b1 iff ex b4 being Subset of c1 st
( b3 = Fr b4 & b4 in c2 ) ) ) & ( for b3 being Subset of c1 holds
( b3 in b2 iff ex b4 being Subset of c1 st
( b3 = Fr b4 & b4 in c2 ) ) ) implies b1 = b2 )
proof end;
end;

:: deftheorem Def1 defines Fr TOPGEN_1:def 1 :
for b1 being TopSpace
for b2, b3 being Subset-Family of b1 holds
( b3 = Fr b2 iff for b4 being Subset of b1 holds
( b4 in b3 iff ex b5 being Subset of b1 st
( b4 = Fr b5 & b5 in b2 ) ) );

theorem Th6: :: TOPGEN_1:6
for b1 being TopSpace
for b2 being Subset-Family of b1 holds
( b2 = {} implies Fr b2 = {} )
proof end;

theorem Th7: :: TOPGEN_1:7
for b1 being TopSpace
for b2 being Subset-Family of b1
for b3 being Subset of b1 holds
( b2 = {b3} implies Fr b2 = {(Fr b3)} )
proof end;

theorem Th8: :: TOPGEN_1:8
for b1 being TopSpace
for b2, b3 being Subset-Family of b1 holds
( b2 c= b3 implies Fr b2 c= Fr b3 )
proof end;

theorem Th9: :: TOPGEN_1:9
for b1 being TopSpace
for b2, b3 being Subset-Family of b1 holds Fr (b2 \/ b3) = (Fr b2) \/ (Fr b3)
proof end;

theorem Th10: :: TOPGEN_1:10
for b1 being TopStruct
for b2 being Subset of b1 holds Fr b2 = (Cl b2) \ (Int b2)
proof end;

theorem Th11: :: TOPGEN_1:11
for b1 being non empty TopSpace
for b2 being Subset of b1
for b3 being Point of b1 holds
( b3 in Fr b2 iff for b4 being Subset of b1 holds
( b4 is open & b3 in b4 implies ( b2 meets b4 & b4 \ b2 <> {} ) ) )
proof end;

theorem Th12: :: TOPGEN_1:12
for b1 being non empty TopSpace
for b2 being Subset of b1
for b3 being Point of b1 holds
( b3 in Fr b2 iff for b4 being Basis of b3
for b5 being Subset of b1 holds
( b5 in b4 implies ( b2 meets b5 & b5 \ b2 <> {} ) ) )
proof end;

theorem Th13: :: TOPGEN_1:13
for b1 being non empty TopSpace
for b2 being Subset of b1
for b3 being Point of b1 holds
( b3 in Fr b2 iff ex b4 being Basis of b3 st
for b5 being Subset of b1 holds
( b5 in b4 implies ( b2 meets b5 & b5 \ b2 <> {} ) ) )
proof end;

theorem Th14: :: TOPGEN_1:14
for b1 being TopSpace
for b2, b3 being Subset of b1 holds Fr (b2 /\ b3) c= ((Cl b2) /\ (Fr b3)) \/ ((Fr b2) /\ (Cl b3))
proof end;

theorem Th15: :: TOPGEN_1:15
for b1 being TopSpace
for b2 being Subset of b1 holds the carrier of b1 = ((Int b2) \/ (Fr b2)) \/ (Int (b2 ` ))
proof end;

theorem Th16: :: TOPGEN_1:16
for b1 being TopSpace
for b2 being Subset of b1 holds
( ( b2 is open & b2 is closed ) iff Fr b2 = {} )
proof end;

definition
let c1 be TopStruct ;
let c2 be Subset of c1;
let c3 be set ;
pred c3 is_an_accumulation_point_of c2 means :Def2: :: TOPGEN_1:def 2
a3 in Cl (a2 \ {a3});
end;

:: deftheorem Def2 defines is_an_accumulation_point_of TOPGEN_1:def 2 :
for b1 being TopStruct
for b2 being Subset of b1
for b3 being set holds
( b3 is_an_accumulation_point_of b2 iff b3 in Cl (b2 \ {b3}) );

theorem Th17: :: TOPGEN_1:17
for b1 being TopSpace
for b2 being Subset of b1
for b3 being set holds
( b3 is_an_accumulation_point_of b2 implies b3 is Point of b1 )
proof end;

definition
let c1 be TopStruct ;
let c2 be Subset of c1;
func Der c2 -> Subset of a1 means :Def3: :: TOPGEN_1:def 3
for b1 being set holds
( b1 in the carrier of a1 implies ( b1 in a3 iff b1 is_an_accumulation_point_of a2 ) );
existence
ex b1 being Subset of c1 st
for b2 being set holds
( b2 in the carrier of c1 implies ( b2 in b1 iff b2 is_an_accumulation_point_of c2 ) )
proof end;
uniqueness
for b1, b2 being Subset of c1 holds
( ( for b3 being set holds
( b3 in the carrier of c1 implies ( b3 in b1 iff b3 is_an_accumulation_point_of c2 ) ) ) & ( for b3 being set holds
( b3 in the carrier of c1 implies ( b3 in b2 iff b3 is_an_accumulation_point_of c2 ) ) ) implies b1 = b2 )
proof end;
end;

:: deftheorem Def3 defines Der TOPGEN_1:def 3 :
for b1 being TopStruct
for b2, b3 being Subset of b1 holds
( b3 = Der b2 iff for b4 being set holds
( b4 in the carrier of b1 implies ( b4 in b3 iff b4 is_an_accumulation_point_of b2 ) ) );

theorem Th18: :: TOPGEN_1:18
for b1 being non empty TopSpace
for b2 being Subset of b1
for b3 being set holds
( b3 in Der b2 iff b3 is_an_accumulation_point_of b2 )
proof end;

theorem Th19: :: TOPGEN_1:19
for b1 being non empty TopSpace
for b2 being Subset of b1
for b3 being Point of b1 holds
( b3 in Der b2 iff for b4 being open Subset of b1 holds
not ( b3 in b4 & ( for b5 being Point of b1 holds
not ( b5 in b2 /\ b4 & b3 <> b5 ) ) ) )
proof end;

theorem Th20: :: TOPGEN_1:20
for b1 being non empty TopSpace
for b2 being Subset of b1
for b3 being Point of b1 holds
( b3 in Der b2 iff for b4 being Basis of b3
for b5 being Subset of b1 holds
not ( b5 in b4 & ( for b6 being Point of b1 holds
not ( b6 in b2 /\ b5 & b3 <> b6 ) ) ) )
proof end;

theorem Th21: :: TOPGEN_1:21
for b1 being non empty TopSpace
for b2 being Subset of b1
for b3 being Point of b1 holds
( b3 in Der b2 iff ex b4 being Basis of b3 st
for b5 being Subset of b1 holds
not ( b5 in b4 & ( for b6 being Point of b1 holds
not ( b6 in b2 /\ b5 & b3 <> b6 ) ) ) )
proof end;

definition
let c1 be TopSpace;
let c2 be Subset of c1;
let c3 be set ;
pred c3 is_isolated_in c2 means :Def4: :: TOPGEN_1:def 4
( a3 in a2 & not a3 is_an_accumulation_point_of a2 );
end;

:: deftheorem Def4 defines is_isolated_in TOPGEN_1:def 4 :
for b1 being TopSpace
for b2 being Subset of b1
for b3 being set holds
( b3 is_isolated_in b2 iff ( b3 in b2 & not b3 is_an_accumulation_point_of b2 ) );

theorem Th22: :: TOPGEN_1:22
for b1 being non empty TopSpace
for b2 being Subset of b1
for b3 being set holds
( b3 in b2 \ (Der b2) iff b3 is_isolated_in b2 )
proof end;

theorem Th23: :: TOPGEN_1:23
for b1 being non empty TopSpace
for b2 being Subset of b1
for b3 being Point of b1 holds
( b3 is_an_accumulation_point_of b2 iff for b4 being open Subset of b1 holds
not ( b3 in b4 & ( for b5 being Point of b1 holds
not ( b5 <> b3 & b5 in b2 & b5 in b4 ) ) ) )
proof end;

theorem Th24: :: TOPGEN_1:24
for b1 being non empty TopSpace
for b2 being Subset of b1
for b3 being Point of b1 holds
( b3 is_isolated_in b2 iff ex b4 being open Subset of b1 st b4 /\ b2 = {b3} )
proof end;

definition
let c1 be TopSpace;
let c2 be Point of c1;
attr a2 is isolated means :Def5: :: TOPGEN_1:def 5
a2 is_isolated_in [#] a1;
end;

:: deftheorem Def5 defines isolated TOPGEN_1:def 5 :
for b1 being TopSpace
for b2 being Point of b1 holds
( b2 is isolated iff b2 is_isolated_in [#] b1 );

theorem Th25: :: TOPGEN_1:25
for b1 being non empty TopSpace
for b2 being Point of b1 holds
( b2 is isolated iff {b2} is open )
proof end;

definition
let c1 be TopSpace;
let c2 be Subset-Family of c1;
func Der c2 -> Subset-Family of a1 means :Def6: :: TOPGEN_1:def 6
for b1 being Subset of a1 holds
( b1 in a3 iff ex b2 being Subset of a1 st
( b1 = Der b2 & b2 in a2 ) );
existence
ex b1 being Subset-Family of c1 st
for b2 being Subset of c1 holds
( b2 in b1 iff ex b3 being Subset of c1 st
( b2 = Der b3 & b3 in c2 ) )
proof end;
uniqueness
for b1, b2 being Subset-Family of c1 holds
( ( for b3 being Subset of c1 holds
( b3 in b1 iff ex b4 being Subset of c1 st
( b3 = Der b4 & b4 in c2 ) ) ) & ( for b3 being Subset of c1 holds
( b3 in b2 iff ex b4 being Subset of c1 st
( b3 = Der b4 & b4 in c2 ) ) ) implies b1 = b2 )
proof end;
end;

:: deftheorem Def6 defines Der TOPGEN_1:def 6 :
for b1 being TopSpace
for b2, b3 being Subset-Family of b1 holds
( b3 = Der b2 iff for b4 being Subset of b1 holds
( b4 in b3 iff ex b5 being Subset of b1 st
( b4 = Der b5 & b5 in b2 ) ) );

theorem Th26: :: TOPGEN_1:26
for b1 being non empty TopSpace
for b2 being Subset-Family of b1 holds
( b2 = {} implies Der b2 = {} )
proof end;

theorem Th27: :: TOPGEN_1:27
for b1 being non empty TopSpace
for b2 being Subset of b1
for b3 being Subset-Family of b1 holds
( b3 = {b2} implies Der b3 = {(Der b2)} )
proof end;

theorem Th28: :: TOPGEN_1:28
for b1 being non empty TopSpace
for b2, b3 being Subset-Family of b1 holds
( b2 c= b3 implies Der b2 c= Der b3 )
proof end;

theorem Th29: :: TOPGEN_1:29
for b1 being non empty TopSpace
for b2, b3 being Subset-Family of b1 holds Der (b2 \/ b3) = (Der b2) \/ (Der b3)
proof end;

theorem Th30: :: TOPGEN_1:30
for b1 being non empty TopSpace
for b2 being Subset of b1 holds Der b2 c= Cl b2
proof end;

theorem Th31: :: TOPGEN_1:31
for b1 being TopSpace
for b2 being Subset of b1 holds Cl b2 = b2 \/ (Der b2)
proof end;

theorem Th32: :: TOPGEN_1:32
for b1 being non empty TopSpace
for b2, b3 being Subset of b1 holds
( b2 c= b3 implies Der b2 c= Der b3 )
proof end;

theorem Th33: :: TOPGEN_1:33
for b1 being non empty TopSpace
for b2, b3 being Subset of b1 holds Der (b2 \/ b3) = (Der b2) \/ (Der b3)
proof end;

theorem Th34: :: TOPGEN_1:34
for b1 being non empty TopSpace
for b2 being Subset of b1 holds
( b1 is being_T1 implies Der (Der b2) c= Der b2 )
proof end;

theorem Th35: :: TOPGEN_1:35
for b1 being TopSpace
for b2 being Subset of b1 holds
( b1 is being_T1 implies Cl (Der b2) = Der b2 )
proof end;

registration
let c1 be non empty being_T1 TopSpace;
let c2 be Subset of c1;
cluster Der a2 -> closed ;
coherence
Der c2 is closed
proof end;
end;

theorem Th36: :: TOPGEN_1:36
for b1 being non empty TopSpace
for b2 being Subset-Family of b1 holds union (Der b2) c= Der (union b2)
proof end;

theorem Th37: :: TOPGEN_1:37
for b1 being non empty TopSpace
for b2, b3 being Subset of b1
for b4 being set holds
( b2 c= b3 & b4 is_an_accumulation_point_of b2 implies b4 is_an_accumulation_point_of b3 )
proof end;

definition
let c1 be TopSpace;
let c2 be Subset of c1;
attr a2 is dense-in-itself means :Def7: :: TOPGEN_1:def 7
a2 c= Der a2;
end;

:: deftheorem Def7 defines dense-in-itself TOPGEN_1:def 7 :
for b1 being TopSpace
for b2 being Subset of b1 holds
( b2 is dense-in-itself iff b2 c= Der b2 );

definition
let c1 be non empty TopSpace;
attr a1 is dense-in-itself means :: TOPGEN_1:def 8
[#] a1 is dense-in-itself;
end;

:: deftheorem Def8 defines dense-in-itself TOPGEN_1:def 8 :
for b1 being non empty TopSpace holds
( b1 is dense-in-itself iff [#] b1 is dense-in-itself );

theorem Th38: :: TOPGEN_1:38
for b1 being non empty TopSpace
for b2 being Subset of b1 holds
( b1 is being_T1 & b2 is dense-in-itself implies Cl b2 is dense-in-itself )
proof end;

definition
let c1 be TopSpace;
let c2 be Subset-Family of c1;
attr a2 is dense-in-itself means :Def9: :: TOPGEN_1:def 9
for b1 being Subset of a1 holds
( b1 in a2 implies b1 is dense-in-itself );
end;

:: deftheorem Def9 defines dense-in-itself TOPGEN_1:def 9 :
for b1 being TopSpace
for b2 being Subset-Family of b1 holds
( b2 is dense-in-itself iff for b3 being Subset of b1 holds
( b3 in b2 implies b3 is dense-in-itself ) );

theorem Th39: :: TOPGEN_1:39
for b1 being non empty TopSpace
for b2 being Subset-Family of b1 holds
( b2 is dense-in-itself implies union b2 c= union (Der b2) )
proof end;

theorem Th40: :: TOPGEN_1:40
for b1 being non empty TopSpace
for b2 being Subset-Family of b1 holds
( b2 is dense-in-itself implies union b2 is dense-in-itself )
proof end;

theorem Th41: :: TOPGEN_1:41
for b1 being non empty TopSpace holds Fr ({} b1) = {}
proof end;

registration
let c1 be TopSpace;
let c2 be open closed Subset of c1;
cluster Fr a2 -> empty ;
coherence
Fr c2 is empty
by Th16;
end;

registration
let c1 be non empty non discrete TopSpace;
cluster non open Element of K10(the carrier of a1);
existence
not for b1 being Subset of c1 holds b1 is open
by TDLAT_3:17;
cluster non closed Element of K10(the carrier of a1);
existence
not for b1 being Subset of c1 holds b1 is closed
by TDLAT_3:18;
end;

registration
let c1 be non empty non discrete TopSpace;
let c2 be non open Subset of c1;
cluster Fr a2 -> non empty ;
coherence
not Fr c2 is empty
by Th16;
end;

registration
let c1 be non empty non discrete TopSpace;
let c2 be non closed Subset of c1;
cluster Fr a2 -> non empty ;
coherence
not Fr c2 is empty
by Th16;
end;

definition
let c1 be TopSpace;
let c2 be Subset of c1;
attr a2 is perfect means :Def10: :: TOPGEN_1:def 10
( a2 is closed & a2 is dense-in-itself );
end;

:: deftheorem Def10 defines perfect TOPGEN_1:def 10 :
for b1 being TopSpace
for b2 being Subset of b1 holds
( b2 is perfect iff ( b2 is closed & b2 is dense-in-itself ) );

registration
let c1 be TopSpace;
cluster perfect -> closed dense-in-itself Element of K10(the carrier of a1);
coherence
for b1 being Subset of c1 holds
( b1 is perfect implies ( b1 is closed & b1 is dense-in-itself ) )
by Def10;
cluster closed dense-in-itself -> perfect Element of K10(the carrier of a1);
coherence
for b1 being Subset of c1 holds
( b1 is closed & b1 is dense-in-itself implies b1 is perfect )
by Def10;
end;

theorem Th42: :: TOPGEN_1:42
for b1 being TopSpace holds Der ({} b1) = {} b1
proof end;

Lemma36: for b1 being TopSpace
for b2 being Subset of b1 holds
( b2 is closed & b2 is dense-in-itself implies Der b2 = b2 )
proof end;

theorem Th43: :: TOPGEN_1:43
for b1 being TopSpace
for b2 being Subset of b1 holds
( b2 is perfect iff Der b2 = b2 )
proof end;

theorem Th44: :: TOPGEN_1:44
for b1 being TopSpace holds {} b1 is perfect
proof end;

registration
let c1 be TopSpace;
cluster empty -> closed dense-in-itself perfect Element of K10(the carrier of a1);
coherence
for b1 being Subset of c1 holds
( b1 is empty implies b1 is perfect )
proof end;
end;

registration
let c1 be TopSpace;
cluster closed dense-in-itself perfect Element of K10(the carrier of a1);
existence
ex b1 being Subset of c1 st b1 is perfect
proof end;
end;

definition
let c1 be TopSpace;
let c2 be Subset of c1;
attr a2 is scattered means :Def11: :: TOPGEN_1:def 11
for b1 being Subset of a1 holds
not ( not b1 is empty & b1 c= a2 & b1 is dense-in-itself );
end;

:: deftheorem Def11 defines scattered TOPGEN_1:def 11 :
for b1 being TopSpace
for b2 being Subset of b1 holds
( b2 is scattered iff for b3 being Subset of b1 holds
not ( not b3 is empty & b3 c= b2 & b3 is dense-in-itself ) );

registration
let c1 be non empty TopSpace;
cluster non empty scattered -> non dense-in-itself Element of K10(the carrier of a1);
coherence
for b1 being Subset of c1 holds
not ( not b1 is empty & b1 is scattered & b1 is dense-in-itself )
by Def11;
cluster non empty dense-in-itself -> non scattered Element of K10(the carrier of a1);
coherence
for b1 being Subset of c1 holds
not ( b1 is dense-in-itself & not b1 is empty & b1 is scattered )
by Def11;
end;

theorem Th45: :: TOPGEN_1:45
for b1 being TopSpace holds {} b1 is scattered
proof end;

registration
let c1 be TopSpace;
cluster empty -> scattered Element of K10(the carrier of a1);
coherence
for b1 being Subset of c1 holds
( b1 is empty implies b1 is scattered )
proof end;
end;

theorem Th46: :: TOPGEN_1:46
for b1 being non empty TopSpace holds
not ( b1 is being_T1 & ( for b2, b3 being Subset of b1 holds
not ( b2 \/ b3 = [#] b1 & b2 misses b3 & b2 is perfect & b3 is scattered ) ) )
proof end;

registration
let c1 be discrete TopSpace;
let c2 be Subset of c1;
cluster Fr a2 -> empty dense-in-itself perfect scattered ;
coherence
Fr c2 is empty
proof end;
end;

registration
let c1 be discrete TopSpace;
cluster -> open closed Element of K10(the carrier of a1);
coherence
for b1 being Subset of c1 holds
( b1 is closed & b1 is open )
by TDLAT_3:18, TDLAT_3:17;
end;

theorem Th47: :: TOPGEN_1:47
for b1 being discrete TopSpace
for b2 being Subset of b1 holds Der b2 = {}
proof end;

registration
let c1 be non empty discrete TopSpace;
let c2 be Subset of c1;
cluster Der a2 -> empty open closed dense-in-itself perfect scattered ;
coherence
Der c2 is empty
by Th47;
end;

theorem Th48: :: TOPGEN_1:48
for b1 being non empty discrete TopSpace
for b2 being Subset of b1 holds
( b2 is dense implies b2 = [#] b1 ) by TOPS_3:19;

definition
let c1 be TopSpace;
func density c1 -> cardinal number means :Def12: :: TOPGEN_1:def 12
( ex b1 being Subset of a1 st
( b1 is dense & a2 = Card b1 ) & ( for b1 being Subset of a1 holds
( b1 is dense implies a2 <=` Card b1 ) ) );
existence
ex b1 being cardinal number st
( ex b2 being Subset of c1 st
( b2 is dense & b1 = Card b2 ) & ( for b2 being Subset of c1 holds
( b2 is dense implies b1 <=` Card b2 ) ) )
proof end;
uniqueness
for b1, b2 being cardinal number holds
( ex b3 being Subset of c1 st
( b3 is dense & b1 = Card b3 ) & ( for b3 being Subset of c1 holds
( b3 is dense implies b1 <=` Card b3 ) ) & ex b3 being Subset of c1 st
( b3 is dense & b2 = Card b3 ) & ( for b3 being Subset of c1 holds
( b3 is dense implies b2 <=` Card b3 ) ) implies b1 = b2 )
proof end;
end;

:: deftheorem Def12 defines density TOPGEN_1:def 12 :
for b1 being TopSpace
for b2 being cardinal number holds
( b2 = density b1 iff ( ex b3 being Subset of b1 st
( b3 is dense & b2 = Card b3 ) & ( for b3 being Subset of b1 holds
( b3 is dense implies b2 <=` Card b3 ) ) ) );

definition
let c1 be TopSpace;
attr a1 is separable means :Def13: :: TOPGEN_1:def 13
density a1 <=` alef 0;
end;

:: deftheorem Def13 defines separable TOPGEN_1:def 13 :
for b1 being TopSpace holds
( b1 is separable iff density b1 <=` alef 0 );

theorem Th49: :: TOPGEN_1:49
for b1 being countable TopSpace holds b1 is separable
proof end;

registration
cluster countable -> separable TopStruct ;
coherence
for b1 being TopSpace holds
( b1 is countable implies b1 is separable )
by Th49;
end;

theorem Th50: :: TOPGEN_1:50
for b1 being Subset of R^1 holds
( b1 = RAT implies b1 ` = IRRAT ) by BORSUK_5:def 3, BORSUK_5:13;

Lemma45: RAT = REAL \ IRRAT
proof end;

theorem Th51: :: TOPGEN_1:51
for b1 being Subset of R^1 holds
( b1 = IRRAT implies b1 ` = RAT ) by Lemma45, BORSUK_5:13;

theorem Th52: :: TOPGEN_1:52
for b1 being Subset of R^1 holds
( b1 = RAT implies Int b1 = {} )
proof end;

theorem Th53: :: TOPGEN_1:53
for b1 being Subset of R^1 holds
( b1 = IRRAT implies Int b1 = {} )
proof end;

theorem Th54: :: TOPGEN_1:54
for b1 being Subset of R^1 holds
( b1 = RAT implies b1 is dense )
proof end;

theorem Th55: :: TOPGEN_1:55
for b1 being Subset of R^1 holds
( b1 = IRRAT implies b1 is dense )
proof end;

theorem Th56: :: TOPGEN_1:56
for b1 being Subset of R^1 holds
( b1 = RAT implies b1 is boundary )
proof end;

theorem Th57: :: TOPGEN_1:57
for b1 being Subset of R^1 holds
( b1 = IRRAT implies b1 is boundary )
proof end;

theorem Th58: :: TOPGEN_1:58
for b1 being Subset of R^1 holds
not ( b1 = REAL & b1 is boundary )
proof end;

theorem Th59: :: TOPGEN_1:59
ex b1, b2 being Subset of R^1 st
( b1 is boundary & b2 is boundary & not b1 \/ b2 is boundary )
proof end;