begin
theorem Th1:
theorem Th2:
theorem
canceled;
theorem
canceled;
theorem Th5:
theorem
canceled;
theorem
theorem
theorem
theorem
:: deftheorem PCOMPS_1:def 1 :
canceled;
:: deftheorem Def2 defines locally_finite PCOMPS_1:def 2 :
for T being TopStruct
for IT being Subset-Family of T holds
( IT is locally_finite iff for x being Point of T ex W being Subset of T st
( x in W & W is open & { V where V is Subset of T : ( V in IT & V meets W ) } is finite ) );
theorem Th11:
theorem Th12:
theorem Th13:
:: deftheorem Def3 defines clf PCOMPS_1:def 3 :
for T being TopStruct
for FX, b3 being Subset-Family of T holds
( b3 = clf FX iff for Z being Subset of T holds
( Z in b3 iff ex W being Subset of T st
( Z = Cl W & W in FX ) ) );
theorem
theorem Th15:
theorem Th16:
theorem Th17:
theorem Th18:
theorem Th19:
theorem Th20:
theorem
theorem
theorem Th23:
theorem
:: deftheorem Def4 defines paracompact PCOMPS_1:def 4 :
for IT being TopStruct holds
( IT is paracompact iff for FX being Subset-Family of IT st FX is Cover of IT & FX is open holds
ex GX being Subset-Family of IT st
( GX is open & GX is Cover of IT & GX is_finer_than FX & GX is locally_finite ) );
theorem
theorem Th26:
theorem Th27:
theorem
:: deftheorem Def5 defines Family_open_set PCOMPS_1:def 5 :
for PM being MetrStruct
for b2 being Subset-Family of PM holds
( b2 = Family_open_set PM iff for V being Subset of PM holds
( V in b2 iff for x being Element of PM st x in V holds
ex r being Real st
( r > 0 & Ball (x,r) c= V ) ) );
theorem Th29:
theorem Th30:
theorem
theorem
canceled;
theorem Th33:
theorem Th34:
theorem Th35:
theorem Th36:
theorem Th37:
:: deftheorem defines TopSpaceMetr PCOMPS_1:def 6 :
for PM being MetrStruct holds TopSpaceMetr PM = TopStruct(# the carrier of PM,(Family_open_set PM) #);
theorem Th38:
definition
let D be
set ;
let f be
Function of
[:D,D:],
REAL;
pred f is_metric_of D means :
Def7:
for
a,
b,
c being
Element of
D holds
( (
f . (
a,
b)
= 0 implies
a = b ) & (
a = b implies
f . (
a,
b)
= 0 ) &
f . (
a,
b)
= f . (
b,
a) &
f . (
a,
c)
<= (f . (a,b)) + (f . (b,c)) );
end;
:: deftheorem Def7 defines is_metric_of PCOMPS_1:def 7 :
for D being set
for f being Function of [:D,D:],REAL holds
( f is_metric_of D iff for a, b, c being Element of D holds
( ( f . (a,b) = 0 implies a = b ) & ( a = b implies f . (a,b) = 0 ) & f . (a,b) = f . (b,a) & f . (a,c) <= (f . (a,b)) + (f . (b,c)) ) );
theorem Th39:
:: deftheorem Def8 defines SpaceMetr PCOMPS_1:def 8 :
for D being set
for f being Function of [:D,D:],REAL st f is_metric_of D holds
SpaceMetr (D,f) = MetrStruct(# D,f #);
:: deftheorem defines metrizable PCOMPS_1:def 9 :
for IT being TopStruct holds
( IT is metrizable iff ex f being Function of [: the carrier of IT, the carrier of IT:],REAL st
( f is_metric_of the carrier of IT & Family_open_set (SpaceMetr ( the carrier of IT,f)) = the topology of IT ) );
theorem