begin
:: deftheorem defines dist METRIC_1:def 1 :
for M being MetrStruct
for a, b being Element of M holds dist (a,b) = the distance of M . (a,b);
Lm1:
op2 . ({},{}) = 0
Lm2:
for x, y being Element of 1 holds
( op2 . (x,y) = 0 iff x = y )
Lm3:
for x, y being Element of 1 holds op2 . (x,y) = op2 . (y,x)
Lm4:
for x, y, z being Element of 1 holds op2 . (x,z) <= (op2 . (x,y)) + (op2 . (y,z))
definition
canceled;let A be
set ;
let f be
PartFunc of
[:A,A:],
REAL;
attr f is
Reflexive means :
Def3:
for
a being
Element of
A holds
f . (
a,
a)
= 0 ;
attr f is
discerning means :
Def4:
for
a,
b being
Element of
A st
f . (
a,
b)
= 0 holds
a = b;
attr f is
symmetric means :
Def5:
for
a,
b being
Element of
A holds
f . (
a,
b)
= f . (
b,
a);
attr f is
triangle means :
Def6:
for
a,
b,
c being
Element of
A holds
f . (
a,
c)
<= (f . (a,b)) + (f . (b,c));
end;
:: deftheorem METRIC_1:def 2 :
canceled;
:: deftheorem Def3 defines Reflexive METRIC_1:def 3 :
for A being set
for f being PartFunc of [:A,A:],REAL holds
( f is Reflexive iff for a being Element of A holds f . (a,a) = 0 );
:: deftheorem Def4 defines discerning METRIC_1:def 4 :
for A being set
for f being PartFunc of [:A,A:],REAL holds
( f is discerning iff for a, b being Element of A st f . (a,b) = 0 holds
a = b );
:: deftheorem Def5 defines symmetric METRIC_1:def 5 :
for A being set
for f being PartFunc of [:A,A:],REAL holds
( f is symmetric iff for a, b being Element of A holds f . (a,b) = f . (b,a) );
:: deftheorem Def6 defines triangle METRIC_1:def 6 :
for A being set
for f being PartFunc of [:A,A:],REAL holds
( f is triangle iff for a, b, c being Element of A holds f . (a,c) <= (f . (a,b)) + (f . (b,c)) );
:: deftheorem Def7 defines Reflexive METRIC_1:def 7 :
for M being MetrStruct holds
( M is Reflexive iff the distance of M is Reflexive );
:: deftheorem Def8 defines discerning METRIC_1:def 8 :
for M being MetrStruct holds
( M is discerning iff the distance of M is discerning );
:: deftheorem Def9 defines symmetric METRIC_1:def 9 :
for M being MetrStruct holds
( M is symmetric iff the distance of M is symmetric );
:: deftheorem Def10 defines triangle METRIC_1:def 10 :
for M being MetrStruct holds
( M is triangle iff the distance of M is triangle );
theorem Th1:
theorem Th2:
theorem Th3:
theorem Th4:
theorem Th5:
theorem Th6:
theorem Th7:
definition
let A be
set ;
func discrete_dist A -> Function of
[:A,A:],
REAL means :
Def11:
for
x,
y being
Element of
A holds
(
it . (
x,
x)
= 0 & (
x <> y implies
it . (
x,
y)
= 1 ) );
existence
ex b1 being Function of [:A,A:],REAL st
for x, y being Element of A holds
( b1 . (x,x) = 0 & ( x <> y implies b1 . (x,y) = 1 ) )
uniqueness
for b1, b2 being Function of [:A,A:],REAL st ( for x, y being Element of A holds
( b1 . (x,x) = 0 & ( x <> y implies b1 . (x,y) = 1 ) ) ) & ( for x, y being Element of A holds
( b2 . (x,x) = 0 & ( x <> y implies b2 . (x,y) = 1 ) ) ) holds
b1 = b2
end;
:: deftheorem Def11 defines discrete_dist METRIC_1:def 11 :
for A being set
for b2 being Function of [:A,A:],REAL holds
( b2 = discrete_dist A iff for x, y being Element of A holds
( b2 . (x,x) = 0 & ( x <> y implies b2 . (x,y) = 1 ) ) );
:: deftheorem defines DiscreteSpace METRIC_1:def 12 :
for A being set holds DiscreteSpace A = MetrStruct(# A,(discrete_dist A) #);
definition
func real_dist -> Function of
[:REAL,REAL:],
REAL means :
Def13:
for
x,
y being
Element of
REAL holds
it . (
x,
y)
= abs (x - y);
existence
ex b1 being Function of [:REAL,REAL:],REAL st
for x, y being Element of REAL holds b1 . (x,y) = abs (x - y)
uniqueness
for b1, b2 being Function of [:REAL,REAL:],REAL st ( for x, y being Element of REAL holds b1 . (x,y) = abs (x - y) ) & ( for x, y being Element of REAL holds b2 . (x,y) = abs (x - y) ) holds
b1 = b2
end;
:: deftheorem Def13 defines real_dist METRIC_1:def 13 :
for b1 being Function of [:REAL,REAL:],REAL holds
( b1 = real_dist iff for x, y being Element of REAL holds b1 . (x,y) = abs (x - y) );
theorem
canceled;
theorem Th9:
theorem Th10:
theorem Th11:
:: deftheorem defines RealSpace METRIC_1:def 14 :
RealSpace = MetrStruct(# REAL,real_dist #);
:: deftheorem Def15 defines Ball METRIC_1:def 15 :
for M being MetrStruct
for p being Element of M
for r being real number
for b4 being Subset of M holds
( ( not M is empty implies ( b4 = Ball (p,r) iff ex M9 being non empty MetrStruct ex p9 being Element of M9 st
( M9 = M & p9 = p & b4 = { q where q is Element of M9 : dist (p9,q) < r } ) ) ) & ( M is empty implies ( b4 = Ball (p,r) iff b4 is empty ) ) );
:: deftheorem Def16 defines cl_Ball METRIC_1:def 16 :
for M being MetrStruct
for p being Element of M
for r being real number
for b4 being Subset of M holds
( ( not M is empty implies ( b4 = cl_Ball (p,r) iff ex M9 being non empty MetrStruct ex p9 being Element of M9 st
( M9 = M & p9 = p & b4 = { q where q is Element of M9 : dist (p9,q) <= r } ) ) ) & ( M is empty implies ( b4 = cl_Ball (p,r) iff b4 is empty ) ) );
:: deftheorem Def17 defines Sphere METRIC_1:def 17 :
for M being MetrStruct
for p being Element of M
for r being real number
for b4 being Subset of M holds
( ( not M is empty implies ( b4 = Sphere (p,r) iff ex M9 being non empty MetrStruct ex p9 being Element of M9 st
( M9 = M & p9 = p & b4 = { q where q is Element of M9 : dist (p9,q) = r } ) ) ) & ( M is empty implies ( b4 = Sphere (p,r) iff b4 is empty ) ) );
Lm5:
for r being real number
for M being non empty MetrStruct
for p being Element of M holds Ball (p,r) = { q where q is Element of M : dist (p,q) < r }
Lm6:
for r being real number
for M being non empty MetrStruct
for p being Element of M holds cl_Ball (p,r) = { q where q is Element of M : dist (p,q) <= r }
Lm7:
for r being real number
for M being non empty MetrStruct
for p being Element of M holds Sphere (p,r) = { q where q is Element of M : dist (p,q) = r }
theorem Th12:
theorem Th13:
theorem Th14:
theorem Th15:
theorem Th16:
theorem
theorem
theorem
theorem
begin
Th2:
for M being MetrSpace
for x, y being Element of M st x <> y holds
0 < dist (x,y)
by Th7;
theorem
theorem Th5:
theorem Th6x:
theorem Th7:
theorem Th8:
set M0 = MetrStruct(# 1,Empty^2-to-zero #);
:: deftheorem Def1 defines Discerning METRIC_1:def 18 :
for A being non empty set
for f being Function of [:A,A:],REAL holds
( f is Discerning iff for a, b being Element of A st a <> b holds
0 < f . (a,b) );
:: deftheorem Def2 defines Discerning METRIC_1:def 19 :
for M being non empty MetrStruct holds
( M is Discerning iff the distance of M is Discerning );
theorem Th14:
:: deftheorem Def4 defines ultra METRIC_1:def 20 :
for M being non empty MetrStruct holds
( M is ultra iff for a, b, c being Element of M holds dist (a,c) <= max ((dist (a,b)),(dist (b,c))) );
theorem Th18:
:: deftheorem defines Set_to_zero METRIC_1:def 21 :
Set_to_zero = [:2,2:] --> 0;
theorem Th40:
theorem Th42:
theorem Th43:
:: deftheorem defines ZeroSpace METRIC_1:def 22 :
ZeroSpace = MetrStruct(# 2,Set_to_zero #);
:: deftheorem Def7 defines is_between METRIC_1:def 23 :
for S being MetrStruct
for p, q, r being Element of S holds
( q is_between p,r iff ( p <> q & p <> r & q <> r & dist (p,r) = (dist (p,q)) + (dist (q,r)) ) );
theorem
theorem
theorem
:: deftheorem defines open_dist_Segment METRIC_1:def 24 :
for M being non empty MetrStruct
for p, r being Element of M holds open_dist_Segment (p,r) = { q where q is Element of M : q is_between p,r } ;
theorem
:: deftheorem defines close_dist_Segment METRIC_1:def 25 :
for M being non empty MetrStruct
for p, r being Element of M holds close_dist_Segment (p,r) = { q where q is Element of M : q is_between p,r } \/ {p,r};
theorem