Lm1:
0 in REAL
by XREAL_0:def 1;
Lm2:
op2 . (0,0) = 0
Lm3:
for x, y being Element of 1 holds
( op2 . (x,y) = 0 iff x = y )
Lm4:
for x, y being Element of 1 holds op2 . (x,y) = op2 . (y,x)
Lm5:
for x, y, z being Element of 1 holds op2 . (x,z) <= (op2 . (x,y)) + (op2 . (y,z))
definition
let A be
set ;
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;
definition
existence
ex b1 being Function of [:REAL,REAL:],REAL st
for x, y being Real holds b1 . (x,y) = |.(x - y).|
uniqueness
for b1, b2 being Function of [:REAL,REAL:],REAL st ( for x, y being Real holds b1 . (x,y) = |.(x - y).| ) & ( for x, y being Real holds b2 . (x,y) = |.(x - y).| ) holds
b1 = b2
end;
set M0 = MetrStruct(# 1,Empty^2-to-zero #);
:: let A,B be set, f be PartFunc of [:A,B:],REAL;
:: let a be Element of A;
:: let b be Element of B;
:: cluster f.(a,b) -> real;
:: coherence
:: proof
:: per cases;
:: suppose
:: [a,b] in dom f;
:: hence thesis by PARTFUN1:4;
:: end;
:: suppose
:: not [a,b] in dom f;
:: then f.(a,b) = 0 by FUNCT_1:def 2;
:: hence thesis;
:: end;
:: end;
:: end;