begin
theorem Th1:
:: deftheorem Def1 defines diagonal ROUGHS_1:def 1 :
for A being RelStr holds
( A is diagonal iff the InternalRel of A c= id the carrier of A );
theorem
Lm1:
for A being RelStr st A is reflexive & A is trivial holds
A is discrete
theorem
Lm2:
for A being RelStr st A is discrete holds
A is diagonal
theorem Th4:
theorem Th5:
theorem Th6:
theorem Th7:
begin
:: deftheorem Def2 defines with_equivalence ROUGHS_1:def 2 :
for P being RelStr holds
( P is with_equivalence iff the InternalRel of P is Equivalence_Relation of the carrier of P );
:: deftheorem Def3 defines with_tolerance ROUGHS_1:def 3 :
for P being RelStr holds
( P is with_tolerance iff the InternalRel of P is Tolerance of the carrier of P );
:: deftheorem defines LAp ROUGHS_1:def 4 :
for A being non empty RelStr
for X being Subset of A holds LAp X = { x where x is Element of A : Class ( the InternalRel of A,x) c= X } ;
:: deftheorem defines UAp ROUGHS_1:def 5 :
for A being non empty RelStr
for X being Subset of A holds UAp X = { x where x is Element of A : Class ( the InternalRel of A,x) meets X } ;
:: deftheorem defines BndAp ROUGHS_1:def 6 :
for A being non empty RelStr
for X being Subset of A holds BndAp X = (UAp X) \ (LAp X);
:: deftheorem Def7 defines rough ROUGHS_1:def 7 :
for A being non empty RelStr
for X being Subset of A holds
( X is rough iff BndAp X <> {} );
theorem Th8:
theorem
theorem Th10:
theorem
theorem Th12:
theorem Th13:
theorem Th14:
theorem Th15:
theorem Th16:
theorem
theorem Th18:
theorem Th19:
theorem Th20:
theorem
theorem
theorem
theorem Th24:
theorem Th25:
theorem
theorem
theorem Th28:
theorem Th29:
theorem
theorem
theorem
theorem
theorem Th34:
theorem
theorem Th36:
theorem Th37:
:: deftheorem defines RoughSet ROUGHS_1:def 8 :
for A being Approximation_Space
for X being Subset of A
for b3 being set holds
( b3 is RoughSet of X iff b3 = [(LAp X),(UAp X)] );
begin
definition
let A be
finite Tolerance_Space;
let X be
Subset of
A;
func MemberFunc (
X,
A)
-> Function of the
carrier of
A,
REAL means :
Def9:
for
x being
Element of
A holds
it . x = (card (X /\ (Class ( the InternalRel of A,x)))) / (card (Class ( the InternalRel of A,x)));
existence
ex b1 being Function of the carrier of A,REAL st
for x being Element of A holds b1 . x = (card (X /\ (Class ( the InternalRel of A,x)))) / (card (Class ( the InternalRel of A,x)))
uniqueness
for b1, b2 being Function of the carrier of A,REAL st ( for x being Element of A holds b1 . x = (card (X /\ (Class ( the InternalRel of A,x)))) / (card (Class ( the InternalRel of A,x))) ) & ( for x being Element of A holds b2 . x = (card (X /\ (Class ( the InternalRel of A,x)))) / (card (Class ( the InternalRel of A,x))) ) holds
b1 = b2
end;
:: deftheorem Def9 defines MemberFunc ROUGHS_1:def 9 :
for A being finite Tolerance_Space
for X being Subset of A
for b3 being Function of the carrier of A,REAL holds
( b3 = MemberFunc (X,A) iff for x being Element of A holds b3 . x = (card (X /\ (Class ( the InternalRel of A,x)))) / (card (Class ( the InternalRel of A,x))) );
theorem Th38:
theorem
theorem Th40:
theorem Th41:
theorem Th42:
theorem Th43:
theorem
theorem
theorem
theorem Th47:
theorem
theorem
theorem
theorem Th51:
theorem
:: deftheorem Def10 defines FinSeqM ROUGHS_1:def 10 :
for A being finite Tolerance_Space
for X being FinSequence of bool the carrier of A
for x being Element of A
for b4 being FinSequence of REAL holds
( b4 = FinSeqM (x,X) iff ( dom b4 = dom X & ( for n being Nat st n in dom X holds
b4 . n = (MemberFunc ((X . n),A)) . x ) ) );
theorem Th53:
theorem Th54:
theorem
theorem
theorem
theorem
begin
:: deftheorem Def11 defines _c= ROUGHS_1:def 11 :
for A being Tolerance_Space
for X, Y being Subset of A holds
( X _c= Y iff LAp X c= LAp Y );
:: deftheorem Def12 defines c=^ ROUGHS_1:def 12 :
for A being Tolerance_Space
for X, Y being Subset of A holds
( X c=^ Y iff UAp X c= UAp Y );
:: deftheorem Def13 defines _c=^ ROUGHS_1:def 13 :
for A being Tolerance_Space
for X, Y being Subset of A holds
( X _c=^ Y iff ( X _c= Y & X c=^ Y ) );
theorem Th59:
theorem Th60:
theorem
begin
:: deftheorem Def14 defines _= ROUGHS_1:def 14 :
for A being Tolerance_Space
for X, Y being Subset of A holds
( X _= Y iff LAp X = LAp Y );
:: deftheorem Def15 defines =^ ROUGHS_1:def 15 :
for A being Tolerance_Space
for X, Y being Subset of A holds
( X =^ Y iff UAp X = UAp Y );
:: deftheorem Def16 defines _=^ ROUGHS_1:def 16 :
for A being Tolerance_Space
for X, Y being Subset of A holds
( X _=^ Y iff ( LAp X = LAp Y & UAp X = UAp Y ) );
:: deftheorem defines _= ROUGHS_1:def 17 :
for A being Tolerance_Space
for X, Y being Subset of A holds
( X _= Y iff ( X _c= Y & Y _c= X ) );
:: deftheorem defines =^ ROUGHS_1:def 18 :
for A being Tolerance_Space
for X, Y being Subset of A holds
( X =^ Y iff ( X c=^ Y & Y c=^ X ) );
:: deftheorem defines _=^ ROUGHS_1:def 19 :
for A being Tolerance_Space
for X, Y being Subset of A holds
( X _=^ Y iff ( X _= Y & X =^ Y ) );