Lm1:
for x, y, z, e being Real st |.(x - y).| < e / 2 & |.(y - z).| < e / 2 holds
|.(x - z).| < e
Lm2:
for p being Real st p > 0 holds
ex k being Nat st 1 / (k + 1) < p
Lm3:
for p being Real
for m being Nat st p > 0 holds
ex i being Nat st
( 1 / (i + 1) < p & i >= m )
theorem Th2:
for
X being
RealUnitarySpace st the
addF of
X is
commutative & the
addF of
X is
associative & the
addF of
X is
having_a_unity holds
for
S being
finite OrthogonalFamily of
X st not
S is
empty holds
for
I being
Function of the
carrier of
X, the
carrier of
X st
S c= dom I & ( for
y being
Point of
X st
y in S holds
I . y = y ) holds
for
H being
Function of the
carrier of
X,
REAL st
S c= dom H & ( for
y being
Point of
X st
y in S holds
H . y = y .|. y ) holds
(setopfunc (S, the carrier of X, the carrier of X,I, the addF of X)) .|. (setopfunc (S, the carrier of X, the carrier of X,I, the addF of X)) = setopfunc (
S, the
carrier of
X,
REAL,
H,
addreal)
Lm4:
for p1, p2 being Real st ( for e being Real st 0 < e holds
|.(p1 - p2).| < e ) holds
p1 = p2