definition
let U be non
empty set ;
let A be non
empty IntervalSet of
U;
existence
ex b1, B being Subset of U st A = Inter (b1,B)
uniqueness
for b1, b2 being Subset of U st ex B being Subset of U st A = Inter (b1,B) & ex B being Subset of U st A = Inter (b2,B) holds
b1 = b2
by Th6;
existence
ex b1, B being Subset of U st A = Inter (B,b1)
uniqueness
for b1, b2 being Subset of U st ex B being Subset of U st A = Inter (B,b1) & ex B being Subset of U st A = Inter (B,b2) holds
b1 = b2
by Th6;
end;
Lm1:
for X being set
for A, B, C being Subset-Family of X holds UNION (A,(INTERSECTION (B,C))) c= INTERSECTION ((UNION (A,B)),(UNION (A,C)))
Lm2:
for X being set
for F being non empty ordered Subset-Family of X
for G being set st G in F holds
( G = min F iff for Y being set st Y in F holds
G c= Y )
by SETFAM_1:3;
Lm3:
for X being set
for F being non empty ordered Subset-Family of X
for G being set st G in F holds
( G = max F iff for Y being set st Y in F holds
Y c= G )
by ZFMISC_1:74;
Lm4:
for U being non empty set
for A being non empty IntervalSet of U holds A is non empty ordered Subset-Family of U
by Th26;
Lm5:
for X being set
for A, B, C being Subset-Family of X holds INTERSECTION (A,(UNION (B,C))) c= UNION ((INTERSECTION (A,B)),(INTERSECTION (A,C)))
definition
let U be non
empty set ;
existence
ex b1 being non empty strict LattStr st
( the carrier of b1 = IntervalSets U & ( for a, b being Element of the carrier of b1
for a9, b9 being non empty IntervalSet of U st a9 = a & b9 = b holds
( the L_join of b1 . (a,b) = a9 _\/_ b9 & the L_meet of b1 . (a,b) = a9 _/\_ b9 ) ) )
uniqueness
for b1, b2 being non empty strict LattStr st the carrier of b1 = IntervalSets U & ( for a, b being Element of the carrier of b1
for a9, b9 being non empty IntervalSet of U st a9 = a & b9 = b holds
( the L_join of b1 . (a,b) = a9 _\/_ b9 & the L_meet of b1 . (a,b) = a9 _/\_ b9 ) ) & the carrier of b2 = IntervalSets U & ( for a, b being Element of the carrier of b2
for a9, b9 being non empty IntervalSet of U st a9 = a & b9 = b holds
( the L_join of b2 . (a,b) = a9 _\/_ b9 & the L_meet of b2 . (a,b) = a9 _/\_ b9 ) ) holds
b1 = b2
end;
definition
let X be
Tolerance_Space;
existence
ex b1 being strict LattStr st
( the carrier of b1 = RoughSets X & ( for A, B being Element of RoughSets X
for A9, B9 being RoughSet of X st A = A9 & B = B9 holds
( the L_join of b1 . (A,B) = A9 _\/_ B9 & the L_meet of b1 . (A,B) = A9 _/\_ B9 ) ) )
uniqueness
for b1, b2 being strict LattStr st the carrier of b1 = RoughSets X & ( for A, B being Element of RoughSets X
for A9, B9 being RoughSet of X st A = A9 & B = B9 holds
( the L_join of b1 . (A,B) = A9 _\/_ B9 & the L_meet of b1 . (A,B) = A9 _/\_ B9 ) ) & the carrier of b2 = RoughSets X & ( for A, B being Element of RoughSets X
for A9, B9 being RoughSet of X st A = A9 & B = B9 holds
( the L_join of b2 . (A,B) = A9 _\/_ B9 & the L_meet of b2 . (A,B) = A9 _/\_ B9 ) ) holds
b1 = b2
end;
Lm6:
for X being Tolerance_Space
for a, b being Element of (RSLattice X) holds a "\/" b = b "\/" a
Lm7:
for X being Tolerance_Space
for a, b, c being Element of (RSLattice X) holds a "\/" (b "\/" c) = (a "\/" b) "\/" c
Lm8:
for X being Tolerance_Space
for K, L being Element of RoughSets X holds the L_join of (RSLattice X) . (( the L_meet of (RSLattice X) . (K,L)),L) = L
Lm9:
for X being Tolerance_Space
for a, b being Element of (RSLattice X) holds (a "/\" b) "\/" b = b
Lm10:
for X being Tolerance_Space
for a, b being Element of (RSLattice X) holds a "/\" b = b "/\" a
Lm11:
for X being Tolerance_Space
for a, b, c being Element of (RSLattice X) holds a "/\" (b "/\" c) = (a "/\" b) "/\" c
Lm12:
for X being Tolerance_Space
for K, L, M being Element of RoughSets X holds the L_meet of (RSLattice X) . (K,( the L_join of (RSLattice X) . (L,M))) = the L_join of (RSLattice X) . (( the L_meet of (RSLattice X) . (K,L)),( the L_meet of (RSLattice X) . (K,M)))
Lm13:
for X being Tolerance_Space
for a, b being Element of (RSLattice X) holds a "/\" (a "\/" b) = a
Lm14:
for X being Tolerance_Space holds ERS X in RoughSets X
by Def20;
Lm15:
for X being Tolerance_Space holds TRS X in RoughSets X
by Def20;
Lm16:
for X being Tolerance_Space holds RSLattice X is complete