begin
Lm1:
for x, y, z, e being Real st abs (x - y) < e / 2 & abs (y - z) < e / 2 holds
abs (x - z) < e
Lm2:
for p being real number st p > 0 holds
ex k being Element of NAT st 1 / (k + 1) < p
Lm3:
for p being real number
for m being Element of NAT st p > 0 holds
ex i being Element of NAT st
( 1 / (i + 1) < p & i >= m )
theorem Th1:
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)
theorem Th3:
theorem Th4:
theorem Th5:
begin
theorem Th6:
theorem Th7:
Lm4:
for p1, p2 being real number st ( for e being Real st 0 < e holds
abs (p1 - p2) < e ) holds
p1 = p2
theorem