:: REAL_NS1 semantic presentation

Lemma1: for b1 being Nat holds
ex b2 being BinOp of REAL b1 st
( ( for b3, b4 being Element of REAL b1 holds b2 . b3,b4 = b3 + b4 ) & b2 is commutative & b2 is associative )
proof end;

Lemma2: for b1 being Nat holds
ex b2 being Function of [:REAL ,(REAL b1):], REAL b1 st
for b3 being Element of REAL
for b4 being Element of REAL b1 holds b2 . b3,b4 = b3 * b4
proof end;

definition
let c1 be Nat;
func Euclid_add c1 -> BinOp of REAL a1 means :Def1: :: REAL_NS1:def 1
for b1, b2 being Element of REAL a1 holds a2 . b1,b2 = b1 + b2;
existence
ex b1 being BinOp of REAL c1 st
for b2, b3 being Element of REAL c1 holds b1 . b2,b3 = b2 + b3
proof end;
uniqueness
for b1, b2 being BinOp of REAL c1 holds
( ( for b3, b4 being Element of REAL c1 holds b1 . b3,b4 = b3 + b4 ) & ( for b3, b4 being Element of REAL c1 holds b2 . b3,b4 = b3 + b4 ) implies b1 = b2 )
proof end;
end;

:: deftheorem Def1 defines Euclid_add REAL_NS1:def 1 :
for b1 being Nat
for b2 being BinOp of REAL b1 holds
( b2 = Euclid_add b1 iff for b3, b4 being Element of REAL b1 holds b2 . b3,b4 = b3 + b4 );

registration
let c1 be Nat;
cluster Euclid_add a1 -> commutative associative ;
coherence
( Euclid_add c1 is commutative & Euclid_add c1 is associative )
proof end;
end;

definition
let c1 be Nat;
func Euclid_mult c1 -> Function of [:REAL ,(REAL a1):], REAL a1 means :Def2: :: REAL_NS1:def 2
for b1 being Element of REAL
for b2 being Element of REAL a1 holds a2 . b1,b2 = b1 * b2;
existence
ex b1 being Function of [:REAL ,(REAL c1):], REAL c1 st
for b2 being Element of REAL
for b3 being Element of REAL c1 holds b1 . b2,b3 = b2 * b3
by Lemma2;
uniqueness
for b1, b2 being Function of [:REAL ,(REAL c1):], REAL c1 holds
( ( for b3 being Element of REAL
for b4 being Element of REAL c1 holds b1 . b3,b4 = b3 * b4 ) & ( for b3 being Element of REAL
for b4 being Element of REAL c1 holds b2 . b3,b4 = b3 * b4 ) implies b1 = b2 )
proof end;
end;

:: deftheorem Def2 defines Euclid_mult REAL_NS1:def 2 :
for b1 being Nat
for b2 being Function of [:REAL ,(REAL b1):], REAL b1 holds
( b2 = Euclid_mult b1 iff for b3 being Element of REAL
for b4 being Element of REAL b1 holds b2 . b3,b4 = b3 * b4 );

definition
let c1 be Nat;
func Euclid_norm c1 -> Function of REAL a1, REAL means :Def3: :: REAL_NS1:def 3
for b1 being Element of REAL a1 holds a2 . b1 = |.b1.|;
existence
ex b1 being Function of REAL c1, REAL st
for b2 being Element of REAL c1 holds b1 . b2 = |.b2.|
proof end;
uniqueness
for b1, b2 being Function of REAL c1, REAL holds
( ( for b3 being Element of REAL c1 holds b1 . b3 = |.b3.| ) & ( for b3 being Element of REAL c1 holds b2 . b3 = |.b3.| ) implies b1 = b2 )
proof end;
end;

:: deftheorem Def3 defines Euclid_norm REAL_NS1:def 3 :
for b1 being Nat
for b2 being Function of REAL b1, REAL holds
( b2 = Euclid_norm b1 iff for b3 being Element of REAL b1 holds b2 . b3 = |.b3.| );

definition
let c1 be Nat;
func REAL-NS c1 -> non empty strict NORMSTR means :Def4: :: REAL_NS1:def 4
( the carrier of a2 = REAL a1 & the Zero of a2 = 0* a1 & the add of a2 = Euclid_add a1 & the Mult of a2 = Euclid_mult a1 & the norm of a2 = Euclid_norm a1 );
existence
ex b1 being non empty strict NORMSTR st
( the carrier of b1 = REAL c1 & the Zero of b1 = 0* c1 & the add of b1 = Euclid_add c1 & the Mult of b1 = Euclid_mult c1 & the norm of b1 = Euclid_norm c1 )
proof end;
uniqueness
for b1, b2 being non empty strict NORMSTR holds
( the carrier of b1 = REAL c1 & the Zero of b1 = 0* c1 & the add of b1 = Euclid_add c1 & the Mult of b1 = Euclid_mult c1 & the norm of b1 = Euclid_norm c1 & the carrier of b2 = REAL c1 & the Zero of b2 = 0* c1 & the add of b2 = Euclid_add c1 & the Mult of b2 = Euclid_mult c1 & the norm of b2 = Euclid_norm c1 implies b1 = b2 )
;
end;

:: deftheorem Def4 defines REAL-NS REAL_NS1:def 4 :
for b1 being Nat
for b2 being non empty strict NORMSTR holds
( b2 = REAL-NS b1 iff ( the carrier of b2 = REAL b1 & the Zero of b2 = 0* b1 & the add of b2 = Euclid_add b1 & the Mult of b2 = Euclid_mult b1 & the norm of b2 = Euclid_norm b1 ) );

registration
let c1 be non empty Nat;
cluster REAL-NS a1 -> non empty non trivial strict ;
coherence
not REAL-NS c1 is trivial
proof end;
end;

theorem Th1: :: REAL_NS1:1
for b1 being Nat
for b2 being VECTOR of (REAL-NS b1)
for b3 being Element of REAL b1 holds
( b2 = b3 implies ||.b2.|| = |.b3.| )
proof end;

theorem Th2: :: REAL_NS1:2
for b1 being Nat
for b2, b3 being VECTOR of (REAL-NS b1)
for b4, b5 being Element of REAL b1 holds
( b2 = b4 & b3 = b5 implies b2 + b3 = b4 + b5 )
proof end;

theorem Th3: :: REAL_NS1:3
for b1 being Nat
for b2 being VECTOR of (REAL-NS b1)
for b3 being Element of REAL b1
for b4 being Real holds
( b2 = b3 implies b4 * b2 = b4 * b3 )
proof end;

registration
let c1 be Nat;
cluster REAL-NS a1 -> non empty Abelian add-associative right_zeroed right_complementable RealLinearSpace-like strict RealNormSpace-like ;
coherence
( REAL-NS c1 is RealNormSpace-like & REAL-NS c1 is RealLinearSpace-like & REAL-NS c1 is Abelian & REAL-NS c1 is add-associative & REAL-NS c1 is right_zeroed & REAL-NS c1 is right_complementable )
proof end;
end;

theorem Th4: :: REAL_NS1:4
for b1 being Nat
for b2 being VECTOR of (REAL-NS b1)
for b3 being Element of REAL b1 holds
( b2 = b3 implies - b2 = - b3 )
proof end;

theorem Th5: :: REAL_NS1:5
for b1 being Nat
for b2, b3 being VECTOR of (REAL-NS b1)
for b4, b5 being Element of REAL b1 holds
( b2 = b4 & b3 = b5 implies b2 - b3 = b4 - b5 )
proof end;

theorem Th6: :: REAL_NS1:6
for b1 being Nat
for b2 being FinSequence of REAL holds
( dom b2 = Seg b1 implies b2 is Element of REAL b1 )
proof end;

theorem Th7: :: REAL_NS1:7
for b1 being Nat
for b2 being Element of REAL b1 holds
( ( for b3 being Nat holds
( b3 in Seg b1 implies 0 <= b2 . b3 ) ) implies ( 0 <= Sum b2 & ( for b3 being Nat holds
( b3 in Seg b1 implies b2 . b3 <= Sum b2 ) ) ) )
proof end;

theorem Th8: :: REAL_NS1:8
for b1 being Nat
for b2 being Element of REAL b1
for b3 being Nat holds
( b3 in Seg b1 implies abs (b2 . b3) <= |.b2.| )
proof end;

theorem Th9: :: REAL_NS1:9
for b1 being Nat
for b2 being Point of (REAL-NS b1)
for b3 being Element of REAL b1 holds
( b2 = b3 implies for b4 being Nat holds
( b4 in Seg b1 implies abs (b3 . b4) <= ||.b2.|| ) )
proof end;

theorem Th10: :: REAL_NS1:10
for b1 being Nat
for b2 being Element of REAL (b1 + 1) holds |.b2.| ^2 = (|.(b2 | b1).| ^2 ) + ((b2 . (b1 + 1)) ^2 )
proof end;

definition
let c1 be Nat;
let c2 be Function of NAT , REAL c1;
let c3 be Nat;
redefine func . as c2 . c3 -> Element of REAL a1;
coherence
c2 . c3 is Element of REAL c1
proof end;
end;

theorem Th11: :: REAL_NS1:11
for b1 being Nat
for b2 being Point of (REAL-NS b1)
for b3 being Element of REAL b1
for b4 being sequence of (REAL-NS b1)
for b5 being Function of NAT , REAL b1 holds
( b3 = b2 & b5 = b4 implies ( ( b4 is convergent & lim b4 = b2 ) iff for b6 being Nat holds
not ( b6 in Seg b1 & ( for b7 being Real_Sequence holds
not for b8 being Nat holds
( b7 . b8 = (b5 . b8) . b6 & b7 is convergent & b3 . b6 = lim b7 ) ) ) ) )
proof end;

theorem Th12: :: REAL_NS1:12
for b1 being Nat
for b2 being sequence of (REAL-NS b1) holds
( b2 is CCauchy implies b2 is convergent )
proof end;

Lemma19: for b1 being Nat holds
REAL-NS b1 is RealBanachSpace
proof end;

registration
let c1 be Nat;
cluster REAL-NS a1 -> non empty Abelian add-associative right_zeroed right_complementable RealLinearSpace-like strict RealNormSpace-like complete ;
coherence
REAL-NS c1 is complete
by Lemma19;
end;

definition
let c1 be Nat;
func Euclid_scalar c1 -> Function of [:(REAL a1),(REAL a1):], REAL means :Def5: :: REAL_NS1:def 5
for b1, b2 being Element of REAL a1 holds a2 . b1,b2 = Sum (mlt b1,b2);
existence
ex b1 being Function of [:(REAL c1),(REAL c1):], REAL st
for b2, b3 being Element of REAL c1 holds b1 . b2,b3 = Sum (mlt b2,b3)
proof end;
uniqueness
for b1, b2 being Function of [:(REAL c1),(REAL c1):], REAL holds
( ( for b3, b4 being Element of REAL c1 holds b1 . b3,b4 = Sum (mlt b3,b4) ) & ( for b3, b4 being Element of REAL c1 holds b2 . b3,b4 = Sum (mlt b3,b4) ) implies b1 = b2 )
proof end;
end;

:: deftheorem Def5 defines Euclid_scalar REAL_NS1:def 5 :
for b1 being Nat
for b2 being Function of [:(REAL b1),(REAL b1):], REAL holds
( b2 = Euclid_scalar b1 iff for b3, b4 being Element of REAL b1 holds b2 . b3,b4 = Sum (mlt b3,b4) );

definition
let c1 be Nat;
func REAL-US c1 -> non empty strict UNITSTR means :Def6: :: REAL_NS1:def 6
( the carrier of a2 = REAL a1 & the Zero of a2 = 0* a1 & the add of a2 = Euclid_add a1 & the Mult of a2 = Euclid_mult a1 & the scalar of a2 = Euclid_scalar a1 );
existence
ex b1 being non empty strict UNITSTR st
( the carrier of b1 = REAL c1 & the Zero of b1 = 0* c1 & the add of b1 = Euclid_add c1 & the Mult of b1 = Euclid_mult c1 & the scalar of b1 = Euclid_scalar c1 )
proof end;
uniqueness
for b1, b2 being non empty strict UNITSTR holds
( the carrier of b1 = REAL c1 & the Zero of b1 = 0* c1 & the add of b1 = Euclid_add c1 & the Mult of b1 = Euclid_mult c1 & the scalar of b1 = Euclid_scalar c1 & the carrier of b2 = REAL c1 & the Zero of b2 = 0* c1 & the add of b2 = Euclid_add c1 & the Mult of b2 = Euclid_mult c1 & the scalar of b2 = Euclid_scalar c1 implies b1 = b2 )
;
end;

:: deftheorem Def6 defines REAL-US REAL_NS1:def 6 :
for b1 being Nat
for b2 being non empty strict UNITSTR holds
( b2 = REAL-US b1 iff ( the carrier of b2 = REAL b1 & the Zero of b2 = 0* b1 & the add of b2 = Euclid_add b1 & the Mult of b2 = Euclid_mult b1 & the scalar of b2 = Euclid_scalar b1 ) );

registration
let c1 be non empty Nat;
cluster REAL-US a1 -> non empty non trivial strict ;
coherence
not REAL-US c1 is trivial
proof end;
end;

registration
let c1 be Nat;
cluster REAL-US a1 -> non empty Abelian add-associative right_zeroed right_complementable RealLinearSpace-like strict RealUnitarySpace-like ;
coherence
( REAL-US c1 is RealUnitarySpace-like & REAL-US c1 is RealLinearSpace-like & REAL-US c1 is Abelian & REAL-US c1 is add-associative & REAL-US c1 is right_zeroed & REAL-US c1 is right_complementable )
proof end;
end;

theorem Th13: :: REAL_NS1:13
for b1 being Nat
for b2 being Real
for b3, b4 being Point of (REAL-NS b1)
for b5, b6 being Point of (REAL-US b1) holds
( b3 = b5 & b4 = b6 implies ( b3 + b4 = b5 + b6 & - b3 = - b5 & b2 * b3 = b2 * b5 ) )
proof end;

theorem Th14: :: REAL_NS1:14
for b1 being Nat
for b2 being Point of (REAL-NS b1)
for b3 being Point of (REAL-US b1) holds
( b2 = b3 implies ||.b2.|| ^2 = b3 .|. b3 )
proof end;

theorem Th15: :: REAL_NS1:15
for b1 being Nat
for b2 being set holds
( b2 is sequence of (REAL-NS b1) iff b2 is sequence of (REAL-US b1) )
proof end;

theorem Th16: :: REAL_NS1:16
for b1 being Nat
for b2 being sequence of (REAL-NS b1)
for b3 being sequence of (REAL-US b1) holds
( b2 = b3 implies ( ( b2 is convergent implies ( b3 is convergent & lim b2 = lim b3 ) ) & ( b3 is convergent implies ( b2 is convergent & lim b2 = lim b3 ) ) ) )
proof end;

theorem Th17: :: REAL_NS1:17
for b1 being Nat
for b2 being sequence of (REAL-NS b1)
for b3 being sequence of (REAL-US b1) holds
( b2 = b3 & b2 is CCauchy implies b3 is Cauchy )
proof end;

theorem Th18: :: REAL_NS1:18
for b1 being Nat
for b2 being sequence of (REAL-NS b1)
for b3 being sequence of (REAL-US b1) holds
( b2 = b3 & b3 is Cauchy implies b2 is CCauchy )
proof end;

Lemma26: for b1 being Nat holds REAL-US b1 is Hilbert
proof end;

registration
let c1 be Nat;
cluster REAL-US a1 -> non empty Abelian add-associative right_zeroed right_complementable RealLinearSpace-like strict RealUnitarySpace-like Hilbert ;
coherence
REAL-US c1 is Hilbert
by Lemma26;
end;