:: RLVECT_1 semantic presentation

definition
attr c1 is strict;
struct LoopStr -> ZeroStr ;
aggr LoopStr(# carrier, add, ZeroF #) -> LoopStr ;
sel add c1 -> BinOp of the carrier of c1;
end;

definition
attr c1 is strict;
struct RLSStruct -> LoopStr ;
aggr RLSStruct(# carrier, ZeroF, add, Mult #) -> RLSStruct ;
sel Mult c1 -> Function of [:REAL ,the carrier of c1:],the carrier of c1;
end;

registration
cluster non empty RLSStruct ;
existence
not for b1 being RLSStruct holds b1 is empty
proof end;
end;

definition
let V be RLSStruct ;
mode VECTOR of V is Element of V;
end;

theorem :: RLVECT_1:1
canceled;

theorem :: RLVECT_1:2
canceled;

theorem :: RLVECT_1:3
for V being non empty 1-sorted
for v being Element of V holds v in V by STRUCT_0:def 5;

registration
cluster non empty strict LoopStr ;
existence
ex b1 being LoopStr st
( b1 is strict & not b1 is empty )
proof end;
end;

definition
let V be non empty LoopStr ;
let v, w be Element of V;
canceled;
canceled;
func v + w -> Element of V equals :: RLVECT_1:def 3
the add of V . v,w;
coherence
the add of V . v,w is Element of V
;
end;

:: deftheorem RLVECT_1:def 1 :
canceled;

:: deftheorem RLVECT_1:def 2 :
canceled;

:: deftheorem defines + RLVECT_1:def 3 :
for V being non empty LoopStr
for v, w being Element of V holds v + w = the add of V . v,w;

definition
let V be non empty RLSStruct ;
let v be VECTOR of V;
let a be Real;
func a * v -> Element of V equals :: RLVECT_1:def 4
the Mult of V . a,v;
coherence
the Mult of V . a,v is Element of V
;
end;

:: deftheorem defines * RLVECT_1:def 4 :
for V being non empty RLSStruct
for v being VECTOR of V
for a being Real holds a * v = the Mult of V . a,v;

theorem :: RLVECT_1:4
canceled;

theorem :: RLVECT_1:5
for V being non empty LoopStr
for v, w being Element of V holds v + w = the add of V . v,w ;

registration
let ZS be non empty set ;
let O be Element of ZS;
let F be BinOp of ZS;
let G be Function of [:REAL ,ZS:],ZS;
cluster RLSStruct(# ZS,O,F,G #) -> non empty ;
coherence
not RLSStruct(# ZS,O,F,G #) is empty
by STRUCT_0:def 1;
end;

Lm1: now
take ZS = {0 };
reconsider O = 0 as Element of ZS by TARSKI:def 1;
take O = O;
deffunc H1( Element of ZS, Element of ZS) -> Element of ZS = O;
consider F being BinOp of ZS such that
A1: for x, y being Element of ZS holds F . x,y = H1(x,y) from BINOP_1:sch 4();
deffunc H2( Element of REAL , Element of ZS) -> Element of ZS = O;
consider G being Function of [:REAL ,ZS:],ZS such that
A2: for a being Element of REAL
for x being Element of ZS holds G . a,x = H2(a,x) from BINOP_1:sch 4();
take F = F;
take G = G;
set W = RLSStruct(# ZS,O,F,G #);
thus for x, y being VECTOR of RLSStruct(# ZS,O,F,G #) holds x + y = y + x
proof
let x, y be VECTOR of RLSStruct(# ZS,O,F,G #);
reconsider X = x, Y = y as Element of ZS ;
( x + y = H1(X,Y) & y + x = H1(Y,X) ) by A1;
hence x + y = y + x ;
end;
thus for x, y, z being VECTOR of RLSStruct(# ZS,O,F,G #) holds (x + y) + z = x + (y + z)
proof
let x, y, z be VECTOR of RLSStruct(# ZS,O,F,G #);
reconsider X = x, Y = y, Z = z as Element of ZS ;
( (x + y) + z = H1(H1(X,Y),Z) & x + (y + z) = H1(X,H1(Y,Z)) ) by A1;
hence (x + y) + z = x + (y + z) ;
end;
thus for x being VECTOR of RLSStruct(# ZS,O,F,G #) holds x + (0. RLSStruct(# ZS,O,F,G #)) = x
proof
let x be VECTOR of RLSStruct(# ZS,O,F,G #);
reconsider X = x as Element of ZS ;
x + (0. RLSStruct(# ZS,O,F,G #)) = H1(X,O) by A1;
hence x + (0. RLSStruct(# ZS,O,F,G #)) = x by TARSKI:def 1;
end;
thus for x being VECTOR of RLSStruct(# ZS,O,F,G #) ex y being VECTOR of RLSStruct(# ZS,O,F,G #) st x + y = 0. RLSStruct(# ZS,O,F,G #)
proof
let x be VECTOR of RLSStruct(# ZS,O,F,G #);
reconsider y = O as VECTOR of RLSStruct(# ZS,O,F,G #) ;
take y ;
thus x + y = 0. RLSStruct(# ZS,O,F,G #) by A1;
end;
thus for a being Real
for x, y being VECTOR of RLSStruct(# ZS,O,F,G #) holds a * (x + y) = (a * x) + (a * y)
proof
let a be Real;
let x, y be VECTOR of RLSStruct(# ZS,O,F,G #);
reconsider X = x, Y = y as Element of ZS ;
(a * x) + (a * y) = O by A1
.= G . a,(F . X,Y) by A2 ;
hence a * (x + y) = (a * x) + (a * y) ;
end;
thus for a, b being Real
for x being VECTOR of RLSStruct(# ZS,O,F,G #) holds (a + b) * x = (a * x) + (b * x)
proof
let a, b be Real;
let x be VECTOR of RLSStruct(# ZS,O,F,G #);
set c = a + b;
reconsider X = x as Element of ZS ;
(a + b) * x = H2(a + b,X) by A2;
hence (a + b) * x = (a * x) + (b * x) by A1;
end;
thus for a, b being Real
for x being VECTOR of RLSStruct(# ZS,O,F,G #) holds (a * b) * x = a * (b * x)
proof
let a, b be Real;
let x be VECTOR of RLSStruct(# ZS,O,F,G #);
set c = a * b;
reconsider X = x as Element of ZS ;
(a * b) * x = H2(a * b,X) by A2;
hence (a * b) * x = a * (b * x) by A2;
end;
thus for x being VECTOR of RLSStruct(# ZS,O,F,G #) holds 1 * x = x
proof
let x be VECTOR of RLSStruct(# ZS,O,F,G #);
reconsider X = x as Element of ZS ;
reconsider A' = 1 as Element of REAL ;
1 * x = H2(A',X) by A2;
hence 1 * x = x by TARSKI:def 1;
end;
end;

definition
let IT be non empty LoopStr ;
attr IT is Abelian means :Def5: :: RLVECT_1:def 5
for v, w being Element of IT holds v + w = w + v;
attr IT is add-associative means :Def6: :: RLVECT_1:def 6
for u, v, w being Element of IT holds (u + v) + w = u + (v + w);
attr IT is right_zeroed means :Def7: :: RLVECT_1:def 7
for v being Element of IT holds v + (0. IT) = v;
attr IT is right_complementable means :Def8: :: RLVECT_1:def 8
for v being Element of IT ex w being Element of IT st v + w = 0. IT;
end;

:: deftheorem Def5 defines Abelian RLVECT_1:def 5 :
for IT being non empty LoopStr holds
( IT is Abelian iff for v, w being Element of IT holds v + w = w + v );

:: deftheorem Def6 defines add-associative RLVECT_1:def 6 :
for IT being non empty LoopStr holds
( IT is add-associative iff for u, v, w being Element of IT holds (u + v) + w = u + (v + w) );

:: deftheorem Def7 defines right_zeroed RLVECT_1:def 7 :
for IT being non empty LoopStr holds
( IT is right_zeroed iff for v being Element of IT holds v + (0. IT) = v );

:: deftheorem Def8 defines right_complementable RLVECT_1:def 8 :
for IT being non empty LoopStr holds
( IT is right_complementable iff for v being Element of IT ex w being Element of IT st v + w = 0. IT );

definition
let IT be non empty RLSStruct ;
attr IT is RealLinearSpace-like means :Def9: :: RLVECT_1:def 9
( ( for a being Real
for v, w being VECTOR of IT holds a * (v + w) = (a * v) + (a * w) ) & ( for a, b being Real
for v being VECTOR of IT holds (a + b) * v = (a * v) + (b * v) ) & ( for a, b being Real
for v being VECTOR of IT holds (a * b) * v = a * (b * v) ) & ( for v being VECTOR of IT holds 1 * v = v ) );
end;

:: deftheorem Def9 defines RealLinearSpace-like RLVECT_1:def 9 :
for IT being non empty RLSStruct holds
( IT is RealLinearSpace-like iff ( ( for a being Real
for v, w being VECTOR of IT holds a * (v + w) = (a * v) + (a * w) ) & ( for a, b being Real
for v being VECTOR of IT holds (a + b) * v = (a * v) + (b * v) ) & ( for a, b being Real
for v being VECTOR of IT holds (a * b) * v = a * (b * v) ) & ( for v being VECTOR of IT holds 1 * v = v ) ) );

registration
cluster non empty strict Abelian add-associative right_zeroed right_complementable LoopStr ;
existence
ex b1 being non empty LoopStr st
( b1 is strict & b1 is Abelian & b1 is add-associative & b1 is right_zeroed & b1 is right_complementable )
proof end;
end;

registration
cluster non empty strict Abelian add-associative right_zeroed right_complementable RealLinearSpace-like RLSStruct ;
existence
ex b1 being non empty RLSStruct st
( not b1 is empty & b1 is strict & b1 is Abelian & b1 is add-associative & b1 is right_zeroed & b1 is right_complementable & b1 is RealLinearSpace-like )
proof end;
end;

definition
mode RealLinearSpace is non empty Abelian add-associative right_zeroed right_complementable RealLinearSpace-like RLSStruct ;
end;

definition
let V be non empty Abelian LoopStr ;
let v, w be Element of V;
:: original: +
redefine func v + w -> Element of V;
commutativity
for v, w being Element of V holds v + w = w + v
by Def5;
end;

theorem :: RLVECT_1:6
canceled;

theorem :: RLVECT_1:7
for V being non empty RLSStruct st ( for v, w being VECTOR of V holds v + w = w + v ) & ( for u, v, w being VECTOR of V holds (u + v) + w = u + (v + w) ) & ( for v being VECTOR of V holds v + (0. V) = v ) & ( for v being VECTOR of V ex w being VECTOR of V st v + w = 0. V ) & ( for a being Real
for v, w being VECTOR of V holds a * (v + w) = (a * v) + (a * w) ) & ( for a, b being Real
for v being VECTOR of V holds (a + b) * v = (a * v) + (b * v) ) & ( for a, b being Real
for v being VECTOR of V holds (a * b) * v = a * (b * v) ) & ( for v being VECTOR of V holds 1 * v = v ) holds
V is RealLinearSpace by Def5, Def6, Def7, Def8, Def9;

Lm2: for V being non empty add-associative right_zeroed right_complementable LoopStr
for v, w being Element of V st v + w = 0. V holds
w + v = 0. V
proof end;

theorem :: RLVECT_1:8
canceled;

theorem :: RLVECT_1:9
canceled;

theorem Th10: :: RLVECT_1:10
for V being non empty add-associative right_zeroed right_complementable LoopStr
for v being Element of V holds
( v + (0. V) = v & (0. V) + v = v )
proof end;

definition
let V be non empty LoopStr ;
let v be Element of V;
assume A1: ( V is add-associative & V is right_zeroed & V is right_complementable ) ;
func - v -> Element of V means :Def10: :: RLVECT_1:def 10
v + it = 0. V;
existence
ex b1 being Element of V st v + b1 = 0. V
by A1, Def8;
uniqueness
for b1, b2 being Element of V st v + b1 = 0. V & v + b2 = 0. V holds
b1 = b2
proof end;
end;

:: deftheorem Def10 defines - RLVECT_1:def 10 :
for V being non empty LoopStr
for v being Element of V st V is add-associative & V is right_zeroed & V is right_complementable holds
for b3 being Element of V holds
( b3 = - v iff v + b3 = 0. V );

Lm3: for V being non empty add-associative right_zeroed right_complementable LoopStr
for v, u being Element of V ex w being Element of V st v + w = u
proof end;

definition
let V be non empty LoopStr ;
let v, w be Element of V;
func v - w -> Element of V equals :: RLVECT_1:def 11
v + (- w);
correctness
coherence
v + (- w) is Element of V
;
;
end;

:: deftheorem defines - RLVECT_1:def 11 :
for V being non empty LoopStr
for v, w being Element of V holds v - w = v + (- w);

theorem :: RLVECT_1:11
canceled;

theorem :: RLVECT_1:12
canceled;

theorem :: RLVECT_1:13
canceled;

theorem :: RLVECT_1:14
canceled;

theorem :: RLVECT_1:15
canceled;

theorem Th16: :: RLVECT_1:16
for V being non empty add-associative right_zeroed right_complementable LoopStr
for v being Element of V holds
( v + (- v) = 0. V & (- v) + v = 0. V )
proof end;

theorem :: RLVECT_1:17
canceled;

theorem :: RLVECT_1:18
canceled;

theorem Th19: :: RLVECT_1:19
for V being non empty add-associative right_zeroed right_complementable LoopStr
for v, w being Element of V st v + w = 0. V holds
v = - w
proof end;

theorem :: RLVECT_1:20
for V being non empty add-associative right_zeroed right_complementable LoopStr
for v, u being Element of V ex w being Element of V st v + w = u by Lm3;

theorem Th21: :: RLVECT_1:21
for V being non empty add-associative right_zeroed right_complementable LoopStr
for w, u, v1, v2 being Element of V st ( w + v1 = w + v2 or v1 + w = v2 + w ) holds
v1 = v2
proof end;

theorem :: RLVECT_1:22
for V being non empty add-associative right_zeroed right_complementable LoopStr
for v, w being Element of V st ( v + w = v or w + v = v ) holds
w = 0. V
proof end;

theorem Th23: :: RLVECT_1:23
for a being Real
for V being RealLinearSpace
for v being VECTOR of V st ( a = 0 or v = 0. V ) holds
a * v = 0. V
proof end;

theorem Th24: :: RLVECT_1:24
for a being Real
for V being RealLinearSpace
for v being VECTOR of V holds
( not a * v = 0. V or a = 0 or v = 0. V )
proof end;

theorem Th25: :: RLVECT_1:25
for V being non empty add-associative right_zeroed right_complementable LoopStr holds - (0. V) = 0. V
proof end;

theorem :: RLVECT_1:26
for V being non empty add-associative right_zeroed right_complementable LoopStr
for v being Element of V holds v - (0. V) = v
proof end;

theorem :: RLVECT_1:27
for V being non empty add-associative right_zeroed right_complementable LoopStr
for v being Element of V holds (0. V) - v = - v by Th10;

theorem :: RLVECT_1:28
for V being non empty add-associative right_zeroed right_complementable LoopStr
for v being Element of V holds v - v = 0. V by Def10;

theorem Th29: :: RLVECT_1:29
for V being RealLinearSpace
for v being VECTOR of V holds - v = (- 1) * v
proof end;

theorem Th30: :: RLVECT_1:30
for V being non empty add-associative right_zeroed right_complementable LoopStr
for v being Element of V holds - (- v) = v
proof end;

theorem Th31: :: RLVECT_1:31
for V being non empty add-associative right_zeroed right_complementable LoopStr
for v, w being Element of V st - v = - w holds
v = w
proof end;

theorem :: RLVECT_1:32
canceled;

theorem Th33: :: RLVECT_1:33
for V being RealLinearSpace
for v being VECTOR of V st v = - v holds
v = 0. V
proof end;

theorem :: RLVECT_1:34
for V being RealLinearSpace
for v being VECTOR of V st v + v = 0. V holds
v = 0. V
proof end;

theorem Th35: :: RLVECT_1:35
for V being non empty add-associative right_zeroed right_complementable LoopStr
for v, w being Element of V st v - w = 0. V holds
v = w
proof end;

theorem :: RLVECT_1:36
for V being non empty add-associative right_zeroed right_complementable LoopStr
for u, v being Element of V ex w being Element of V st v - w = u
proof end;

theorem :: RLVECT_1:37
for V being non empty add-associative right_zeroed right_complementable LoopStr
for w, v1, v2 being Element of V st w - v1 = w - v2 holds
v1 = v2
proof end;

theorem Th38: :: RLVECT_1:38
for a being Real
for V being RealLinearSpace
for v being VECTOR of V holds a * (- v) = (- a) * v
proof end;

theorem Th39: :: RLVECT_1:39
for a being Real
for V being RealLinearSpace
for v being VECTOR of V holds a * (- v) = - (a * v)
proof end;

theorem :: RLVECT_1:40
for a being Real
for V being RealLinearSpace
for v being VECTOR of V holds (- a) * (- v) = a * v
proof end;

Lm4: for V being non empty add-associative right_zeroed right_complementable LoopStr
for u, w being Element of V holds - (u + w) = (- w) + (- u)
proof end;

theorem Th41: :: RLVECT_1:41
for V being non empty add-associative right_zeroed right_complementable LoopStr
for v, u, w being Element of V holds v - (u + w) = (v - w) - u
proof end;

theorem :: RLVECT_1:42
for V being non empty add-associative LoopStr
for v, u, w being Element of V holds (v + u) - w = v + (u - w) by Def6;

theorem :: RLVECT_1:43
for V being non empty Abelian add-associative right_zeroed right_complementable LoopStr
for v, u, w being Element of V holds v - (u - w) = (v - u) + w
proof end;

theorem Th44: :: RLVECT_1:44
for V being non empty add-associative right_zeroed right_complementable LoopStr
for v, w being Element of V holds - (v + w) = (- w) - v
proof end;

theorem :: RLVECT_1:45
for V being non empty add-associative right_zeroed right_complementable LoopStr
for v, w being Element of V holds - (v + w) = (- w) + (- v) by Lm4;

theorem :: RLVECT_1:46
for V being non empty Abelian add-associative right_zeroed right_complementable LoopStr
for v, w being Element of V holds (- v) - w = (- w) - v
proof end;

theorem :: RLVECT_1:47
for V being non empty add-associative right_zeroed right_complementable LoopStr
for v, w being Element of V holds - (v - w) = w + (- v)
proof end;

theorem Th48: :: RLVECT_1:48
for a being Real
for V being RealLinearSpace
for v, w being VECTOR of V holds a * (v - w) = (a * v) - (a * w)
proof end;

theorem Th49: :: RLVECT_1:49
for a, b being Real
for V being RealLinearSpace
for v being VECTOR of V holds (a - b) * v = (a * v) - (b * v)
proof end;

theorem :: RLVECT_1:50
for a being Real
for V being RealLinearSpace
for v, w being VECTOR of V st a <> 0 & a * v = a * w holds
v = w
proof end;

theorem :: RLVECT_1:51
for a, b being Real
for V being RealLinearSpace
for v being VECTOR of V st v <> 0. V & a * v = b * v holds
a = b
proof end;

definition
let V be non empty 1-sorted ;
let v, u be Element of V;
:: original: <*
redefine func <*v,u*> -> FinSequence of the carrier of V;
coherence
<*v,u*> is FinSequence of the carrier of V
proof end;
end;

definition
let V be non empty 1-sorted ;
let v, u, w be Element of V;
:: original: <*
redefine func <*v,u,w*> -> FinSequence of the carrier of V;
coherence
<*v,u,w*> is FinSequence of the carrier of V
proof end;
end;

definition
let V be non empty LoopStr ;
let F be FinSequence of the carrier of V;
func Sum F -> Element of V means :Def12: :: RLVECT_1:def 12
ex f being Function of NAT ,the carrier of V st
( it = f . (len F) & f . 0 = 0. V & ( for j being Element of NAT
for v being Element of V st j < len F & v = F . (j + 1) holds
f . (j + 1) = (f . j) + v ) );
existence
ex b1 being Element of V ex f being Function of NAT ,the carrier of V st
( b1 = f . (len F) & f . 0 = 0. V & ( for j being Element of NAT
for v being Element of V st j < len F & v = F . (j + 1) holds
f . (j + 1) = (f . j) + v ) )
proof end;
uniqueness
for b1, b2 being Element of V st ex f being Function of NAT ,the carrier of V st
( b1 = f . (len F) & f . 0 = 0. V & ( for j being Element of NAT
for v being Element of V st j < len F & v = F . (j + 1) holds
f . (j + 1) = (f . j) + v ) ) & ex f being Function of NAT ,the carrier of V st
( b2 = f . (len F) & f . 0 = 0. V & ( for j being Element of NAT
for v being Element of V st j < len F & v = F . (j + 1) holds
f . (j + 1) = (f . j) + v ) ) holds
b1 = b2
proof end;
end;

:: deftheorem Def12 defines Sum RLVECT_1:def 12 :
for V being non empty LoopStr
for F being FinSequence of the carrier of V
for b3 being Element of V holds
( b3 = Sum F iff ex f being Function of NAT ,the carrier of V st
( b3 = f . (len F) & f . 0 = 0. V & ( for j being Element of NAT
for v being Element of V st j < len F & v = F . (j + 1) holds
f . (j + 1) = (f . j) + v ) ) );

Lm5: for V being non empty LoopStr holds Sum (<*> the carrier of V) = 0. V
proof end;

Lm6: for V being non empty LoopStr
for F being FinSequence of the carrier of V st len F = 0 holds
Sum F = 0. V
proof end;

theorem :: RLVECT_1:52
canceled;

theorem :: RLVECT_1:53
canceled;

theorem Th54: :: RLVECT_1:54
for V being non empty LoopStr
for F being FinSequence of the carrier of V
for k, n being Element of NAT st k in Seg n & len F = n holds
F . k is Element of V
proof end;

theorem Th55: :: RLVECT_1:55
for V being non empty LoopStr
for F, G being FinSequence of the carrier of V
for v being Element of V st len F = (len G) + 1 & G = F | (dom G) & v = F . (len F) holds
Sum F = (Sum G) + v
proof end;

theorem :: RLVECT_1:56
for a being Real
for V being RealLinearSpace
for F, G being FinSequence of the carrier of V st len F = len G & ( for k being Element of NAT
for v being VECTOR of V st k in dom F & v = G . k holds
F . k = a * v ) holds
Sum F = a * (Sum G)
proof end;

theorem :: RLVECT_1:57
for V being non empty Abelian add-associative right_zeroed right_complementable LoopStr
for F, G being FinSequence of the carrier of V st len F = len G & ( for k being Element of NAT
for v being Element of V st k in dom F & v = G . k holds
F . k = - v ) holds
Sum F = - (Sum G)
proof end;

theorem Th58: :: RLVECT_1:58
for V being non empty add-associative right_zeroed LoopStr
for F, G being FinSequence of the carrier of V holds Sum (F ^ G) = (Sum F) + (Sum G)
proof end;

Lm7: for V being non empty add-associative right_zeroed right_complementable LoopStr
for v being Element of V holds Sum <*v*> = v
proof end;

theorem :: RLVECT_1:59
for V being non empty Abelian add-associative right_zeroed LoopStr
for F, G being FinSequence of the carrier of V st rng F = rng G & F is one-to-one & G is one-to-one holds
Sum F = Sum G
proof end;

theorem :: RLVECT_1:60
for V being non empty LoopStr holds Sum (<*> the carrier of V) = 0. V by Lm5;

theorem :: RLVECT_1:61
for V being non empty add-associative right_zeroed right_complementable LoopStr
for v being Element of V holds Sum <*v*> = v by Lm7;

theorem Th62: :: RLVECT_1:62
for V being non empty add-associative right_zeroed right_complementable LoopStr
for v, u being Element of V holds Sum <*v,u*> = v + u
proof end;

theorem Th63: :: RLVECT_1:63
for V being non empty add-associative right_zeroed right_complementable LoopStr
for v, u, w being Element of V holds Sum <*v,u,w*> = (v + u) + w
proof end;

theorem :: RLVECT_1:64
for V being RealLinearSpace
for a being Real holds a * (Sum (<*> the carrier of V)) = 0. V
proof end;

theorem :: RLVECT_1:65
canceled;

theorem :: RLVECT_1:66
for V being RealLinearSpace
for a being Real
for v, u being VECTOR of V holds a * (Sum <*v,u*>) = (a * v) + (a * u)
proof end;

theorem :: RLVECT_1:67
for V being RealLinearSpace
for a being Real
for v, u, w being VECTOR of V holds a * (Sum <*v,u,w*>) = ((a * v) + (a * u)) + (a * w)
proof end;

theorem :: RLVECT_1:68
for V being non empty add-associative right_zeroed right_complementable LoopStr holds - (Sum (<*> the carrier of V)) = 0. V
proof end;

theorem :: RLVECT_1:69
for V being non empty add-associative right_zeroed right_complementable LoopStr
for v being Element of V holds - (Sum <*v*>) = - v by Lm7;

theorem :: RLVECT_1:70
for V being non empty Abelian add-associative right_zeroed right_complementable LoopStr
for v, u being Element of V holds - (Sum <*v,u*>) = (- v) - u
proof end;

theorem :: RLVECT_1:71
for V being non empty Abelian add-associative right_zeroed right_complementable LoopStr
for v, u, w being Element of V holds - (Sum <*v,u,w*>) = ((- v) - u) - w
proof end;

theorem :: RLVECT_1:72
for V being non empty Abelian add-associative right_zeroed right_complementable LoopStr
for v, w being Element of V holds Sum <*v,w*> = Sum <*w,v*>
proof end;

theorem :: RLVECT_1:73
for V being non empty add-associative right_zeroed right_complementable LoopStr
for v, w being Element of V holds Sum <*v,w*> = (Sum <*v*>) + (Sum <*w*>)
proof end;

theorem :: RLVECT_1:74
for V being non empty add-associative right_zeroed right_complementable LoopStr holds Sum <*(0. V),(0. V)*> = 0. V
proof end;

theorem :: RLVECT_1:75
for V being non empty add-associative right_zeroed right_complementable LoopStr
for v being Element of V holds
( Sum <*(0. V),v*> = v & Sum <*v,(0. V)*> = v )
proof end;

theorem :: RLVECT_1:76
for V being non empty add-associative right_zeroed right_complementable LoopStr
for v being Element of V holds
( Sum <*v,(- v)*> = 0. V & Sum <*(- v),v*> = 0. V )
proof end;

theorem :: RLVECT_1:77
for V being non empty add-associative right_zeroed right_complementable LoopStr
for v, w being Element of V holds Sum <*v,(- w)*> = v - w by Th62;

theorem Th78: :: RLVECT_1:78
for V being non empty add-associative right_zeroed right_complementable LoopStr
for v, w being Element of V holds Sum <*(- v),(- w)*> = - (w + v)
proof end;

theorem Th79: :: RLVECT_1:79
for V being RealLinearSpace
for v being VECTOR of V holds Sum <*v,v*> = 2 * v
proof end;

theorem :: RLVECT_1:80
for V being RealLinearSpace
for v being VECTOR of V holds Sum <*(- v),(- v)*> = (- 2) * v
proof end;

theorem :: RLVECT_1:81
for V being non empty add-associative right_zeroed right_complementable LoopStr
for u, v, w being Element of V holds Sum <*u,v,w*> = ((Sum <*u*>) + (Sum <*v*>)) + (Sum <*w*>)
proof end;

theorem :: RLVECT_1:82
for V being non empty add-associative right_zeroed right_complementable LoopStr
for u, v, w being Element of V holds Sum <*u,v,w*> = (Sum <*u,v*>) + w
proof end;

theorem :: RLVECT_1:83
for V being non empty Abelian add-associative right_zeroed right_complementable LoopStr
for v, u, w being Element of V holds Sum <*u,v,w*> = (Sum <*v,w*>) + u
proof end;

theorem Th84: :: RLVECT_1:84
for V being non empty Abelian add-associative right_zeroed right_complementable LoopStr
for v, u, w being Element of V holds Sum <*u,v,w*> = (Sum <*u,w*>) + v
proof end;

theorem Th85: :: RLVECT_1:85
for V being non empty Abelian add-associative right_zeroed right_complementable LoopStr
for v, u, w being Element of V holds Sum <*u,v,w*> = Sum <*u,w,v*>
proof end;

theorem Th86: :: RLVECT_1:86
for V being non empty Abelian add-associative right_zeroed right_complementable LoopStr
for v, u, w being Element of V holds Sum <*u,v,w*> = Sum <*v,u,w*>
proof end;

theorem Th87: :: RLVECT_1:87
for V being non empty Abelian add-associative right_zeroed right_complementable LoopStr
for v, u, w being Element of V holds Sum <*u,v,w*> = Sum <*v,w,u*>
proof end;

theorem :: RLVECT_1:88
canceled;

theorem :: RLVECT_1:89
for V being non empty Abelian add-associative right_zeroed right_complementable LoopStr
for v, u, w being Element of V holds Sum <*u,v,w*> = Sum <*w,v,u*>
proof end;

theorem :: RLVECT_1:90
for V being non empty add-associative right_zeroed right_complementable LoopStr holds Sum <*(0. V),(0. V),(0. V)*> = 0. V
proof end;

theorem :: RLVECT_1:91
for V being non empty add-associative right_zeroed right_complementable LoopStr
for v being Element of V holds
( Sum <*(0. V),(0. V),v*> = v & Sum <*(0. V),v,(0. V)*> = v & Sum <*v,(0. V),(0. V)*> = v )
proof end;

theorem :: RLVECT_1:92
for V being non empty add-associative right_zeroed right_complementable LoopStr
for u, v being Element of V holds
( Sum <*(0. V),u,v*> = u + v & Sum <*u,v,(0. V)*> = u + v & Sum <*u,(0. V),v*> = u + v )
proof end;

theorem :: RLVECT_1:93
for V being RealLinearSpace
for v being VECTOR of V holds Sum <*v,v,v*> = 3 * v
proof end;

theorem :: RLVECT_1:94
for V being non empty add-associative right_zeroed right_complementable LoopStr
for F being FinSequence of the carrier of V st len F = 0 holds
Sum F = 0. V by Lm6;

theorem :: RLVECT_1:95
for V being non empty add-associative right_zeroed right_complementable LoopStr
for F being FinSequence of the carrier of V st len F = 1 holds
Sum F = F . 1
proof end;

theorem :: RLVECT_1:96
for V being non empty add-associative right_zeroed right_complementable LoopStr
for F being FinSequence of the carrier of V
for v1, v2 being Element of V st len F = 2 & v1 = F . 1 & v2 = F . 2 holds
Sum F = v1 + v2
proof end;

theorem :: RLVECT_1:97
for V being non empty add-associative right_zeroed right_complementable LoopStr
for F being FinSequence of the carrier of V
for v1, v2, v being Element of V st len F = 3 & v1 = F . 1 & v2 = F . 2 & v = F . 3 holds
Sum F = (v1 + v2) + v
proof end;

definition
let R be non empty ZeroStr ;
let a be Element of R;
attr a is non-zero means :: RLVECT_1:def 13
a <> 0. R;
end;

:: deftheorem defines non-zero RLVECT_1:def 13 :
for R being non empty ZeroStr
for a being Element of R holds
( a is non-zero iff a <> 0. R );

definition
let L be non empty LoopStr ;
attr L is zeroed means :Def14: :: RLVECT_1:def 14
for a being Element of L holds
( a + (0. L) = a & (0. L) + a = a );
attr L is complementable means :Def15: :: RLVECT_1:def 15
for a being Element of L ex b being Element of L st
( a + b = 0. L & b + a = 0. L );
end;

:: deftheorem Def14 defines zeroed RLVECT_1:def 14 :
for L being non empty LoopStr holds
( L is zeroed iff for a being Element of L holds
( a + (0. L) = a & (0. L) + a = a ) );

:: deftheorem Def15 defines complementable RLVECT_1:def 15 :
for L being non empty LoopStr holds
( L is complementable iff for a being Element of L ex b being Element of L st
( a + b = 0. L & b + a = 0. L ) );

registration
cluster non empty zeroed -> non empty right_zeroed LoopStr ;
coherence
for b1 being non empty LoopStr st b1 is zeroed holds
b1 is right_zeroed
proof end;
cluster non empty complementable -> non empty right_complementable LoopStr ;
coherence
for b1 being non empty LoopStr st b1 is complementable holds
b1 is right_complementable
proof end;
end;

registration
cluster non empty Abelian right_zeroed -> non empty zeroed LoopStr ;
coherence
for b1 being non empty LoopStr st b1 is Abelian & b1 is right_zeroed holds
b1 is zeroed
proof end;
cluster non empty Abelian right_complementable -> non empty complementable LoopStr ;
coherence
for b1 being non empty LoopStr st b1 is Abelian & b1 is right_complementable holds
b1 is complementable
proof end;
end;