:: On Paracompactness of Metrizable Spaces
:: by Leszek Borys
::
:: Received July 23, 1992
:: Copyright (c) 1992-2021 Association of Mizar Users


theorem Th1: :: PCOMPS_2:1
for R being Relation
for A being set st R well_orders A holds
( R |_2 A well_orders A & A = field (R |_2 A) )
proof end;

scheme :: PCOMPS_2:sch 1
MinSet{ F1() -> set , F2() -> Relation, P1[ object ] } :
ex X being set st
( X in F1() & P1[X] & ( for Y being set st Y in F1() & P1[Y] holds
[X,Y] in F2() ) )
provided
A1: F2() well_orders F1() and
A2: ex X being set st
( X in F1() & P1[X] )
proof end;

definition
let FX be set ;
let R be Relation;
let B be Element of FX;
func PartUnion (B,R) -> set equals :: PCOMPS_2:def 1
union (R -Seg B);
coherence
union (R -Seg B) is set
;
end;

:: deftheorem defines PartUnion PCOMPS_2:def 1 :
for FX being set
for R being Relation
for B being Element of FX holds PartUnion (B,R) = union (R -Seg B);

definition
let FX be set ;
let R be Relation;
func DisjointFam (FX,R) -> set means :: PCOMPS_2:def 2
for A being set holds
( A in it iff ex B being Element of FX st
( B in FX & A = B \ (PartUnion (B,R)) ) );
existence
ex b1 being set st
for A being set holds
( A in b1 iff ex B being Element of FX st
( B in FX & A = B \ (PartUnion (B,R)) ) )
proof end;
uniqueness
for b1, b2 being set st ( for A being set holds
( A in b1 iff ex B being Element of FX st
( B in FX & A = B \ (PartUnion (B,R)) ) ) ) & ( for A being set holds
( A in b2 iff ex B being Element of FX st
( B in FX & A = B \ (PartUnion (B,R)) ) ) ) holds
b1 = b2
proof end;
end;

:: deftheorem defines DisjointFam PCOMPS_2:def 2 :
for FX being set
for R being Relation
for b3 being set holds
( b3 = DisjointFam (FX,R) iff for A being set holds
( A in b3 iff ex B being Element of FX st
( B in FX & A = B \ (PartUnion (B,R)) ) ) );

definition
let X be set ;
let n be Element of NAT ;
let f be sequence of (bool X);
func PartUnionNat (n,f) -> set equals :: PCOMPS_2:def 3
union (f .: ((Seg n) \ {n}));
coherence
union (f .: ((Seg n) \ {n})) is set
;
end;

:: deftheorem defines PartUnionNat PCOMPS_2:def 3 :
for X being set
for n being Element of NAT
for f being sequence of (bool X) holds PartUnionNat (n,f) = union (f .: ((Seg n) \ {n}));

theorem Th2: :: PCOMPS_2:2
for PT being non empty TopSpace st PT is regular holds
for FX being Subset-Family of PT st FX is Cover of PT & FX is open holds
ex HX being Subset-Family of PT st
( HX is open & HX is Cover of PT & ( for V being Subset of PT st V in HX holds
ex W being Subset of PT st
( W in FX & Cl V c= W ) ) )
proof end;

theorem :: PCOMPS_2:3
for PT being non empty TopSpace
for FX being Subset-Family of PT st PT is T_2 & PT is paracompact & FX is Cover of PT & FX is open holds
ex GX being Subset-Family of PT st
( GX is open & GX is Cover of PT & clf GX is_finer_than FX & GX is locally_finite )
proof end;

theorem Th4: :: PCOMPS_2:4
for PT being non empty TopSpace
for PM being MetrSpace
for f being Function of [: the carrier of PT, the carrier of PT:],REAL st f is_metric_of the carrier of PT & PM = SpaceMetr ( the carrier of PT,f) holds
the carrier of PM = the carrier of PT
proof end;

theorem :: PCOMPS_2:5
for PT being non empty TopSpace
for PM being MetrSpace
for FX being Subset-Family of PT
for f being Function of [: the carrier of PT, the carrier of PT:],REAL st f is_metric_of the carrier of PT & PM = SpaceMetr ( the carrier of PT,f) holds
( FX is Subset-Family of PT iff FX is Subset-Family of PM ) by Th4;

scheme :: PCOMPS_2:sch 2
XXX1{ F1() -> non empty set , F2() -> Subset-Family of F1(), F3( object , object ) -> Subset of F1(), P1[ object ], P2[ object , object , object , object ] } :
ex f being sequence of (bool (bool F1())) st
( f . 0 = F2() & ( for n being Nat holds f . (n + 1) = { (union { F3(c,n) where c is Element of F1() : for fq being Subset-Family of F1()
for q being Nat st q <= n & fq = f . q holds
P2[c,V,n,fq]
}
)
where V is Subset of F1() : P1[V]
}
) )
proof end;

reconsider jd = 1 / 2 as Real ;

scheme :: PCOMPS_2:sch 3
XXX{ F1() -> non empty set , F2() -> Subset-Family of F1(), F3( object , object ) -> Subset of F1(), P1[ object ], P2[ object , object , object ] } :
ex f being sequence of (bool (bool F1())) st
( f . 0 = F2() & ( for n being Nat holds f . (n + 1) = { (union { F3(c,n) where c is Element of F1() : ( P2[c,V,n] & not c in union { (union (f . q)) where q is Nat : q <= n } ) } ) where V is Subset of F1() : P1[V] } ) )
proof end;

theorem Th6: :: PCOMPS_2:6
for PT being non empty TopSpace st PT is metrizable holds
for FX being Subset-Family of PT st FX is Cover of PT & FX is open holds
ex GX being Subset-Family of PT st
( GX is open & GX is Cover of PT & GX is_finer_than FX & GX is locally_finite )
proof end;

theorem :: PCOMPS_2:7
for PT being non empty TopSpace st PT is metrizable holds
PT is paracompact
proof end;