begin
theorem Th1:
theorem Th2:
theorem Th3:
theorem
theorem Th5:
theorem Th6:
theorem Th7:
theorem Th8:
theorem Th9:
theorem Th10:
theorem Th11:
begin
:: deftheorem Def1 defines Classification TAXONOM1:def 1 :
for A being non empty set
for b2 being Subset of (PARTITIONS A) holds
( b2 is Classification of A iff for X, Y being a_partition of A st X in b2 & Y in b2 & not X is_finer_than Y holds
Y is_finer_than X );
theorem
theorem
theorem Th14:
:: deftheorem Def2 defines Strong_Classification TAXONOM1:def 2 :
for A being non empty set
for b2 being Subset of (PARTITIONS A) holds
( b2 is Strong_Classification of A iff ( b2 is Classification of A & {A} in b2 & SmallestPartition A in b2 ) );
theorem
begin
definition
let X be non
empty set ;
let f be
PartFunc of
[:X,X:],
REAL;
let a be
real number ;
func low_toler (
f,
a)
-> Relation of
X means :
Def3:
for
x,
y being
Element of
X holds
(
[x,y] in it iff
f . (
x,
y)
<= a );
existence
ex b1 being Relation of X st
for x, y being Element of X holds
( [x,y] in b1 iff f . (x,y) <= a )
uniqueness
for b1, b2 being Relation of X st ( for x, y being Element of X holds
( [x,y] in b1 iff f . (x,y) <= a ) ) & ( for x, y being Element of X holds
( [x,y] in b2 iff f . (x,y) <= a ) ) holds
b1 = b2
end;
:: deftheorem Def3 defines low_toler TAXONOM1:def 3 :
for X being non empty set
for f being PartFunc of [:X,X:],REAL
for a being real number
for b4 being Relation of X holds
( b4 = low_toler (f,a) iff for x, y being Element of X holds
( [x,y] in b4 iff f . (x,y) <= a ) );
theorem Th16:
theorem Th17:
theorem Th18:
theorem Th19:
:: deftheorem Def4 defines nonnegative TAXONOM1:def 4 :
for X being set
for f being PartFunc of [:X,X:],REAL holds
( f is nonnegative iff for x, y being Element of X holds f . (x,y) >= 0 );
theorem Th20:
theorem Th21:
theorem Th22:
Lm2:
for x being set
for X being non empty set
for a1, a2 being non negative real number st a1 <= a2 holds
for f being PartFunc of [:X,X:],REAL
for R1, R2 being Equivalence_Relation of X st R1 = (low_toler (f,a1)) [*] & R2 = (low_toler (f,a2)) [*] holds
Class (R1,x) c= Class (R2,x)
begin
theorem Th23:
theorem Th24:
theorem
theorem Th26:
theorem Th27:
theorem
begin
:: deftheorem Def5 defines fam_class TAXONOM1:def 5 :
for X being non empty set
for f being PartFunc of [:X,X:],REAL
for b3 being Subset of (PARTITIONS X) holds
( b3 = fam_class f iff for x being set holds
( x in b3 iff ex a being non negative real number ex R being Equivalence_Relation of X st
( R = (low_toler (f,a)) [*] & Class R = x ) ) );
theorem
theorem Th30:
theorem Th31:
theorem
begin
:: deftheorem Def6 defines are_in_tolerance_wrt TAXONOM1:def 6 :
for M being MetrStruct
for a being real number
for x, y being Element of M holds
( x,y are_in_tolerance_wrt a iff dist (x,y) <= a );
definition
let M be non
empty MetrStruct ;
let a be
real number ;
func dist_toler (
M,
a)
-> Relation of
M means :
Def7:
for
x,
y being
Element of
M holds
(
[x,y] in it iff
x,
y are_in_tolerance_wrt a );
existence
ex b1 being Relation of M st
for x, y being Element of M holds
( [x,y] in b1 iff x,y are_in_tolerance_wrt a )
uniqueness
for b1, b2 being Relation of M st ( for x, y being Element of M holds
( [x,y] in b1 iff x,y are_in_tolerance_wrt a ) ) & ( for x, y being Element of M holds
( [x,y] in b2 iff x,y are_in_tolerance_wrt a ) ) holds
b1 = b2
end;
:: deftheorem Def7 defines dist_toler TAXONOM1:def 7 :
for M being non empty MetrStruct
for a being real number
for b3 being Relation of M holds
( b3 = dist_toler (M,a) iff for x, y being Element of M holds
( [x,y] in b3 iff x,y are_in_tolerance_wrt a ) );
theorem Th33:
theorem
:: deftheorem Def8 defines fam_class_metr TAXONOM1:def 8 :
for M being non empty Reflexive symmetric MetrStruct
for b2 being Subset of (PARTITIONS the carrier of M) holds
( b2 = fam_class_metr M iff for x being set holds
( x in b2 iff ex a being non negative real number ex R being Equivalence_Relation of M st
( R = (dist_toler (M,a)) [*] & Class R = x ) ) );
theorem Th35:
theorem Th36:
theorem Th37:
theorem Th38:
theorem Th39:
theorem Th40:
theorem Th41:
theorem Th42:
theorem