:: The Sum and Product of Finite Sequences of Real Numbers
:: by Czes{\l}aw Byli\'nski
::
:: Received May 11, 1990
:: Copyright (c) 1990-2021 Association of Mizar Users


registration
let n be natural Number ;
cluster Relation-like omega -defined Function-like natural-valued V64() n -element FinSequence-like FinSubsequence-like for set ;
existence
ex b1 being FinSequence st
( b1 is n -element & b1 is natural-valued )
proof end;
end;

registration
cluster Relation-like omega -defined Function-like real-valued V64() FinSequence-like FinSubsequence-like for set ;
existence
ex b1 being FinSequence st b1 is real-valued
proof end;
end;

definition
let F be real-valued Relation;
:: original: rng
redefine func rng F -> Subset of REAL;
coherence
rng F is Subset of REAL
by VALUED_0:def 3;
end;

registration
let D be non empty set ;
let F be Function of REAL,D;
let F1 be real-valued FinSequence;
cluster F1 (#) F -> FinSequence-like ;
coherence
F * F1 is FinSequence-like
proof end;
end;

registration
let r be Real;
cluster <*r*> -> real-valued ;
coherence
<*r*> is real-valued
proof end;
end;

registration
let r1, r2 be Real;
cluster <*r1,r2*> -> real-valued ;
coherence
<*r1,r2*> is real-valued
proof end;
end;

registration
let r1, r2, r3 be Real;
cluster <*r1,r2,r3*> -> real-valued ;
coherence
<*r1,r2,r3*> is real-valued
proof end;
end;

registration
let r1, r2, r3, r4 be Real;
cluster <*r1,r2,r3,r4*> -> real-valued ;
coherence
<*r1,r2,r3,r4*> is real-valued
proof end;
end;

registration
let j be natural Number ;
let r be Real;
cluster j |-> r -> real-valued ;
coherence
j |-> r is real-valued
;
end;

registration
let f, g be real-valued FinSequence;
cluster f ^ g -> real-valued ;
coherence
f ^ g is real-valued
proof end;
end;

theorem :: RVSUM_1:1
0 is_a_unity_wrt addreal by BINOP_2:2, SETWISEO:14;

definition
redefine func diffreal equals :: RVSUM_1:def 1
addreal * ((id REAL),compreal);
compatibility
for b1 being Element of bool [:[:REAL,REAL:],REAL:] holds
( b1 = diffreal iff b1 = addreal * ((id REAL),compreal) )
proof end;
correctness
;
end;

:: deftheorem defines diffreal RVSUM_1:def 1 :
diffreal = addreal * ((id REAL),compreal);

definition
func sqrreal -> UnOp of REAL means :Def2: :: RVSUM_1:def 2
for r being Real holds it . r = r ^2 ;
existence
ex b1 being UnOp of REAL st
for r being Real holds b1 . r = r ^2
proof end;
uniqueness
for b1, b2 being UnOp of REAL st ( for r being Real holds b1 . r = r ^2 ) & ( for r being Real holds b2 . r = r ^2 ) holds
b1 = b2
proof end;
end;

:: deftheorem Def2 defines sqrreal RVSUM_1:def 2 :
for b1 being UnOp of REAL holds
( b1 = sqrreal iff for r being Real holds b1 . r = r ^2 );

theorem :: RVSUM_1:2
1 is_a_unity_wrt multreal by BINOP_2:7, SETWISEO:14;

theorem Th3: :: RVSUM_1:3
multreal is_distributive_wrt addreal
proof end;

theorem :: RVSUM_1:4
sqrreal is_distributive_wrt multreal
proof end;

definition
let x be Real;
func x multreal -> UnOp of REAL equals :: RVSUM_1:def 3
multreal [;] (x,(id REAL));
coherence
multreal [;] (x,(id REAL)) is UnOp of REAL
proof end;
end;

:: deftheorem defines multreal RVSUM_1:def 3 :
for x being Real holds x multreal = multreal [;] (x,(id REAL));

Lm1: for r, r1 being Real holds (multreal [;] (r,(id REAL))) . r1 = r * r1
proof end;

theorem :: RVSUM_1:5
for r, r1 being Real holds (r multreal) . r1 = r * r1 by Lm1;

theorem :: RVSUM_1:6
for r being Real holds r multreal is_distributive_wrt addreal
proof end;

theorem Th7: :: RVSUM_1:7
compreal is_an_inverseOp_wrt addreal
proof end;

theorem Th8: :: RVSUM_1:8
addreal is having_an_inverseOp by Th7;

theorem Th9: :: RVSUM_1:9
the_inverseOp_wrt addreal = compreal by Th7, Th8, FINSEQOP:def 3;

theorem :: RVSUM_1:10
compreal is_distributive_wrt addreal by Th8, Th9, FINSEQOP:63;

Lm2: for F being real-valued FinSequence holds F is FinSequence of REAL
proof end;

definition
let F1, F2 be real-valued FinSequence;
:: original: +
redefine func F1 + F2 -> FinSequence of REAL equals :: RVSUM_1:def 4
addreal .: (F1,F2);
coherence
F1 + F2 is FinSequence of REAL
by Lm2;
compatibility
for b1 being FinSequence of REAL holds
( b1 = F1 + F2 iff b1 = addreal .: (F1,F2) )
proof end;
commutativity
for b1 being FinSequence of REAL
for F1, F2 being real-valued FinSequence st b1 = addreal .: (F1,F2) holds
b1 = addreal .: (F2,F1)
proof end;
end;

:: deftheorem defines + RVSUM_1:def 4 :
for F1, F2 being real-valued FinSequence holds F1 + F2 = addreal .: (F1,F2);

definition
let i be natural Number ;
let R1, R2 be Element of i -tuples_on REAL;
:: original: +
redefine func R1 + R2 -> Element of i -tuples_on REAL;
coherence
R1 + R2 is Element of i -tuples_on REAL
by FINSEQ_2:120;
end;

theorem :: RVSUM_1:11
for s being set
for i being natural Number
for R1, R2 being Element of i -tuples_on REAL holds (R1 + R2) . s = (R1 . s) + (R2 . s)
proof end;

theorem :: RVSUM_1:12
for F being real-valued FinSequence holds (<*> REAL) + F = <*> REAL
proof end;

theorem :: RVSUM_1:13
for r1, r2 being Real holds <*r1*> + <*r2*> = <*(r1 + r2)*>
proof end;

theorem :: RVSUM_1:14
for i being natural Number
for r1, r2 being Real holds (i |-> r1) + (i |-> r2) = i |-> (r1 + r2)
proof end;

theorem :: RVSUM_1:15
for F1, F2, F3 being real-valued FinSequence holds F1 + (F2 + F3) = (F1 + F2) + F3 by RFUNCT_1:8;

theorem :: RVSUM_1:16
for i being natural Number
for R being Element of i -tuples_on REAL holds R + (i |-> 0) = R by BINOP_2:2, FINSEQOP:56;

theorem :: RVSUM_1:17
for s being set
for F being real-valued FinSequence holds (- F) . s = - (F . s) by VALUED_1:8;

definition
let F be real-valued FinSequence;
:: original: -
redefine func - F -> FinSequence of REAL equals :: RVSUM_1:def 5
compreal * F;
coherence
- F is FinSequence of REAL
by Lm2;
compatibility
for b1 being FinSequence of REAL holds
( b1 = - F iff b1 = compreal * F )
proof end;
end;

:: deftheorem defines - RVSUM_1:def 5 :
for F being real-valued FinSequence holds - F = compreal * F;

definition
let i be natural Number ;
let R be Element of i -tuples_on REAL;
:: original: -
redefine func - R -> Element of i -tuples_on REAL;
coherence
- R is Element of i -tuples_on REAL
by FINSEQ_2:113;
end;

theorem :: RVSUM_1:18
for s being set
for F being real-valued FinSequence holds (- F) . s = - (F . s) by VALUED_1:8;

theorem :: RVSUM_1:19
- (<*> REAL) = <*> REAL ;

theorem :: RVSUM_1:20
for r being Real holds - <*r*> = <*(- r)*>
proof end;

theorem Th21: :: RVSUM_1:21
for i being natural Number
for r being Real holds - (i |-> r) = i |-> (- r)
proof end;

theorem :: RVSUM_1:22
for i being natural Number
for R being Element of i -tuples_on REAL holds R + (- R) = i |-> 0 by Th8, Th9, BINOP_2:2, FINSEQOP:73;

theorem :: RVSUM_1:23
for i being natural Number
for R1, R2 being Element of i -tuples_on REAL st R1 + R2 = i |-> 0 holds
R1 = - R2 by Th8, Th9, BINOP_2:2, FINSEQOP:74;

theorem :: RVSUM_1:24
for R1, R2 being complex-valued Function st - R1 = - R2 holds
R1 = R2
proof end;

theorem :: RVSUM_1:25
for i being natural Number
for R, R1, R2 being Element of i -tuples_on REAL st R1 + R = R2 + R holds
R1 = R2
proof end;

theorem Th26: :: RVSUM_1:26
for F1, F2 being real-valued FinSequence holds - (F1 + F2) = (- F1) + (- F2)
proof end;

definition
let F1, F2 be real-valued FinSequence;
:: original: -
redefine func F1 - F2 -> FinSequence of REAL equals :: RVSUM_1:def 6
diffreal .: (F1,F2);
coherence
F1 - F2 is FinSequence of REAL
by Lm2;
compatibility
for b1 being FinSequence of REAL holds
( b1 = F1 - F2 iff b1 = diffreal .: (F1,F2) )
proof end;
end;

:: deftheorem defines - RVSUM_1:def 6 :
for F1, F2 being real-valued FinSequence holds F1 - F2 = diffreal .: (F1,F2);

definition
let i be natural Number ;
let R1, R2 be Element of i -tuples_on REAL;
:: original: -
redefine func R1 - R2 -> Element of i -tuples_on REAL;
coherence
R1 - R2 is Element of i -tuples_on REAL
by FINSEQ_2:120;
end;

theorem :: RVSUM_1:27
for s being set
for i being natural Number
for R1, R2 being Element of i -tuples_on REAL holds (R1 - R2) . s = (R1 . s) - (R2 . s)
proof end;

theorem :: RVSUM_1:28
for F being real-valued FinSequence holds
( (<*> REAL) - F = <*> REAL & F - (<*> REAL) = <*> REAL )
proof end;

theorem :: RVSUM_1:29
for r1, r2 being Real holds <*r1*> - <*r2*> = <*(r1 - r2)*>
proof end;

theorem :: RVSUM_1:30
for i being natural Number
for r1, r2 being Real holds (i |-> r1) - (i |-> r2) = i |-> (r1 - r2)
proof end;

theorem :: RVSUM_1:31
for F1, F2 being real-valued FinSequence holds F1 - F2 = F1 + (- F2) ;

theorem :: RVSUM_1:32
for i being natural Number
for R being Element of i -tuples_on REAL holds R - (i |-> 0) = R
proof end;

theorem :: RVSUM_1:33
for i being natural Number
for R being Element of i -tuples_on REAL holds (i |-> 0) - R = - R by BINOP_2:2, FINSEQOP:56;

theorem :: RVSUM_1:34
for F1, F2 being real-valued FinSequence holds F1 - (- F2) = F1 + F2 ;

theorem :: RVSUM_1:35
for F1, F2 being real-valued FinSequence holds - (F1 - F2) = F2 - F1
proof end;

theorem :: RVSUM_1:36
for F1, F2 being real-valued FinSequence holds - (F1 - F2) = (- F1) + F2
proof end;

theorem :: RVSUM_1:37
for i being natural Number
for R being Element of i -tuples_on REAL holds R - R = i |-> 0 by Th8, Th9, BINOP_2:2, FINSEQOP:73;

theorem :: RVSUM_1:38
for i being natural Number
for R1, R2 being Element of i -tuples_on REAL st R1 - R2 = i |-> 0 holds
R1 = R2
proof end;

theorem :: RVSUM_1:39
for F1, F2, F3 being real-valued FinSequence holds (F1 - F2) - F3 = F1 - (F2 + F3) by RFUNCT_1:20;

theorem :: RVSUM_1:40
for F1, F2, F3 being real-valued FinSequence holds F1 + (F2 - F3) = (F1 + F2) - F3 by RFUNCT_1:23;

theorem :: RVSUM_1:41
for F1, F2, F3 being real-valued FinSequence holds F1 - (F2 - F3) = (F1 - F2) + F3 by RFUNCT_1:22;

theorem :: RVSUM_1:42
for i being natural Number
for R, R1 being Element of i -tuples_on REAL holds R1 = (R1 + R) - R
proof end;

theorem :: RVSUM_1:43
for i being natural Number
for R, R1 being Element of i -tuples_on REAL holds R1 = (R1 - R) + R
proof end;

notation
let F be real-valued FinSequence;
let r be Real;
synonym r * F for r (#) F;
end;

theorem :: RVSUM_1:44
for s being set
for r being Real
for F being real-valued FinSequence holds (r * F) . s = r * (F . s) by VALUED_1:6;

definition
let F be real-valued FinSequence;
let r be Real;
:: original: *
redefine func r * F -> FinSequence of REAL equals :: RVSUM_1:def 7
(r multreal) * F;
coherence
r * F is FinSequence of REAL
by Lm2;
compatibility
for b1 being FinSequence of REAL holds
( b1 = r * F iff b1 = (r multreal) * F )
proof end;
end;

:: deftheorem defines * RVSUM_1:def 7 :
for F being real-valued FinSequence
for r being Real holds r * F = (r multreal) * F;

definition
let i be natural Number ;
let R be Element of i -tuples_on REAL;
let r be Real;
:: original: *
redefine func r * R -> Element of i -tuples_on REAL;
coherence
r * R is Element of i -tuples_on REAL
by FINSEQ_2:113;
end;

theorem :: RVSUM_1:45
for s being set
for r being Real
for F being real-valued FinSequence holds (r * F) . s = r * (F . s) by VALUED_1:6;

theorem :: RVSUM_1:46
for r being Real holds r * (<*> REAL) = <*> REAL ;

theorem :: RVSUM_1:47
for r, r1 being Real holds r * <*r1*> = <*(r * r1)*>
proof end;

theorem Th48: :: RVSUM_1:48
for i being natural Number
for r1, r2 being Real holds r1 * (i |-> r2) = i |-> (r1 * r2)
proof end;

theorem :: RVSUM_1:49
for r1, r2 being Real
for F being real-valued FinSequence holds (r1 * r2) * F = r1 * (r2 * F) by RFUNCT_1:17;

theorem :: RVSUM_1:50
for r1, r2 being Real
for F being real-valued FinSequence holds (r1 + r2) * F = (r1 * F) + (r2 * F)
proof end;

theorem :: RVSUM_1:51
for r being Real
for F1, F2 being real-valued FinSequence holds r * (F1 + F2) = (r * F1) + (r * F2) by RFUNCT_1:16;

theorem :: RVSUM_1:52
for F being real-valued FinSequence holds 1 * F = F by RFUNCT_1:21;

theorem :: RVSUM_1:53
for i being natural Number
for R being Element of i -tuples_on REAL holds 0 * R = i |-> 0
proof end;

theorem :: RVSUM_1:54
for F being real-valued FinSequence holds (- 1) * F = - F ;

notation
let F be real-valued FinSequence;
synonym sqr F for F ^2 ;
end;

definition
let F be real-valued FinSequence;
:: original: sqr
redefine func sqr F -> FinSequence of REAL equals :: RVSUM_1:def 8
sqrreal * F;
compatibility
for b1 being FinSequence of REAL holds
( b1 = sqr F iff b1 = sqrreal * F )
proof end;
coherence
sqr F is FinSequence of REAL
by Lm2;
end;

:: deftheorem defines sqr RVSUM_1:def 8 :
for F being real-valued FinSequence holds sqr F = sqrreal * F;

definition
let i be natural Number ;
let R be Element of i -tuples_on REAL;
:: original: sqr
redefine func sqr R -> Element of i -tuples_on REAL;
coherence
sqr R is Element of i -tuples_on REAL
by FINSEQ_2:113;
end;

theorem :: RVSUM_1:55
for r being Real holds sqr <*r*> = <*(r ^2)*>
proof end;

theorem :: RVSUM_1:56
for i being natural Number
for r being Real holds sqr (i |-> r) = i |-> (r ^2)
proof end;

theorem Th57: :: RVSUM_1:57
for F being real-valued FinSequence holds sqr (- F) = sqr F
proof end;

theorem Th58: :: RVSUM_1:58
for r being Real
for F being real-valued FinSequence holds sqr (r * F) = (r ^2) * (sqr F)
proof end;

notation
let F1, F2 be real-valued FinSequence;
synonym mlt (F1,F2) for F1 (#) F2;
end;

definition
let F1, F2 be real-valued FinSequence;
:: original: mlt
redefine func mlt (F1,F2) -> FinSequence of REAL equals :: RVSUM_1:def 9
multreal .: (F1,F2);
coherence
mlt (F1,F2) is FinSequence of REAL
by Lm2;
compatibility
for b1 being FinSequence of REAL holds
( b1 = mlt (F1,F2) iff b1 = multreal .: (F1,F2) )
proof end;
commutativity
for b1 being FinSequence of REAL
for F1, F2 being real-valued FinSequence st b1 = multreal .: (F1,F2) holds
b1 = multreal .: (F2,F1)
proof end;
end;

:: deftheorem defines mlt RVSUM_1:def 9 :
for F1, F2 being real-valued FinSequence holds mlt (F1,F2) = multreal .: (F1,F2);

theorem :: RVSUM_1:59
for s being set
for F1, F2 being real-valued FinSequence holds (mlt (F1,F2)) . s = (F1 . s) * (F2 . s) by VALUED_1:5;

definition
let i be natural Number ;
let R1, R2 be Element of i -tuples_on REAL;
:: original: mlt
redefine func mlt (R1,R2) -> Element of i -tuples_on REAL;
coherence
mlt (R1,R2) is Element of i -tuples_on REAL
by FINSEQ_2:120;
end;

theorem :: RVSUM_1:60
for s being set
for F1, F2 being real-valued FinSequence holds (mlt (F1,F2)) . s = (F1 . s) * (F2 . s) by VALUED_1:5;

theorem :: RVSUM_1:61
for F being real-valued FinSequence holds mlt ((<*> REAL),F) = <*> REAL
proof end;

theorem :: RVSUM_1:62
for r1, r2 being Real holds mlt (<*r1*>,<*r2*>) = <*(r1 * r2)*>
proof end;

theorem Th63: :: RVSUM_1:63
for i being natural Number
for r being Real
for R being Element of i -tuples_on REAL holds mlt ((i |-> r),R) = r * R
proof end;

theorem :: RVSUM_1:64
for i being natural Number
for r1, r2 being Real holds mlt ((i |-> r1),(i |-> r2)) = i |-> (r1 * r2)
proof end;

theorem :: RVSUM_1:65
for r being Real
for F1, F2 being real-valued FinSequence holds r * (mlt (F1,F2)) = mlt ((r * F1),F2) by RFUNCT_1:12;

theorem :: RVSUM_1:66
for i being natural Number
for r being Real
for R being Element of i -tuples_on REAL holds r * R = mlt ((i |-> r),R) by Th63;

theorem :: RVSUM_1:67
for F being real-valued FinSequence holds sqr F = mlt (F,F) ;

theorem Th68: :: RVSUM_1:68
for F1, F2 being real-valued FinSequence holds sqr (F1 + F2) = ((sqr F1) + (2 * (mlt (F1,F2)))) + (sqr F2)
proof end;

theorem Th69: :: RVSUM_1:69
for F1, F2 being real-valued FinSequence holds sqr (F1 - F2) = ((sqr F1) - (2 * (mlt (F1,F2)))) + (sqr F2)
proof end;

theorem :: RVSUM_1:70
for F1, F2 being real-valued FinSequence holds sqr (mlt (F1,F2)) = mlt ((sqr F1),(sqr F2))
proof end;

:: Finite sum of Finite Sequence of Real Numbers
registration
cluster -> complex-valued for FinSequence of COMPLEX ;
coherence
for b1 being FinSequence of COMPLEX holds b1 is complex-valued
;
cluster Relation-like omega -defined Function-like complex-valued real-valued V64() FinSequence-like FinSubsequence-like for set ;
existence
ex b1 being FinSequence st
( b1 is real-valued & b1 is complex-valued )
proof end;
end;

definition
let F be complex-valued FinSequence;
func Sum F -> Complex means :Def10: :: RVSUM_1:def 10
ex f being FinSequence of COMPLEX st
( f = F & it = addcomplex $$ f );
existence
ex b1 being Complex ex f being FinSequence of COMPLEX st
( f = F & b1 = addcomplex $$ f )
proof end;
uniqueness
for b1, b2 being Complex st ex f being FinSequence of COMPLEX st
( f = F & b1 = addcomplex $$ f ) & ex f being FinSequence of COMPLEX st
( f = F & b2 = addcomplex $$ f ) holds
b1 = b2
;
end;

:: deftheorem Def10 defines Sum RVSUM_1:def 10 :
for F being complex-valued FinSequence
for b2 being Complex holds
( b2 = Sum F iff ex f being FinSequence of COMPLEX st
( f = F & b2 = addcomplex $$ f ) );

registration
let F be real-valued FinSequence;
cluster Sum F -> real ;
coherence
Sum F is real
proof end;
end;

theorem Th71: :: RVSUM_1:71
for F being FinSequence of REAL holds Sum F = addreal $$ F
proof end;

definition
let F be FinSequence of COMPLEX ;
:: original: Sum
redefine func Sum F -> Element of COMPLEX equals :: RVSUM_1:def 11
addcomplex $$ F;
coherence
Sum F is Element of COMPLEX
by XCMPLX_0:def 2;
compatibility
for b1 being Element of COMPLEX holds
( b1 = Sum F iff b1 = addcomplex $$ F )
by Def10;
end;

:: deftheorem defines Sum RVSUM_1:def 11 :
for F being FinSequence of COMPLEX holds Sum F = addcomplex $$ F;

definition
let F be FinSequence of REAL ;
:: original: Sum
redefine func Sum F -> Real equals :: RVSUM_1:def 12
addreal $$ F;
coherence
Sum F is Real
;
compatibility
for b1 being Real holds
( b1 = Sum F iff b1 = addreal $$ F )
by Th71;
end;

:: deftheorem defines Sum RVSUM_1:def 12 :
for F being FinSequence of REAL holds Sum F = addreal $$ F;

theorem Th72: :: RVSUM_1:72
Sum (<*> REAL) = 0 by BINOP_2:2, FINSOP_1:10;

theorem :: RVSUM_1:73
for r being Real holds Sum <*r*> = r
proof end;

theorem Th74: :: RVSUM_1:74
for r being Real
for F being real-valued FinSequence holds Sum (F ^ <*r*>) = (Sum F) + r
proof end;

theorem Th75: :: RVSUM_1:75
for F1, F2 being real-valued FinSequence holds Sum (F1 ^ F2) = (Sum F1) + (Sum F2)
proof end;

theorem :: RVSUM_1:76
for r being Real
for F being real-valued FinSequence holds Sum (<*r*> ^ F) = r + (Sum F)
proof end;

theorem Th77: :: RVSUM_1:77
for r1, r2 being Real holds Sum <*r1,r2*> = r1 + r2
proof end;

theorem Th78: :: RVSUM_1:78
for r1, r2, r3 being Real holds Sum <*r1,r2,r3*> = (r1 + r2) + r3
proof end;

theorem :: RVSUM_1:79
for R being Element of 0 -tuples_on REAL holds Sum R = 0 by Th72;

theorem Th80: :: RVSUM_1:80
for i being natural Number
for r being Real holds Sum (i |-> r) = i * r
proof end;

theorem Th81: :: RVSUM_1:81
for i being natural Number holds Sum (i |-> 0) = 0
proof end;

Lm: for i being natural Number
for R1 being real-valued b1 -element FinSequence holds R1 is Element of i -tuples_on REAL

proof end;

theorem Th82: :: RVSUM_1:82
for i being natural Number
for R1, R2 being real-valued b1 -element FinSequence st ( for j being Nat st j in Seg i holds
R1 . j <= R2 . j ) holds
Sum R1 <= Sum R2
proof end;

theorem Th83: :: RVSUM_1:83
for i being natural Number
for R1, R2 being real-valued b1 -element FinSequence st ( for j being Nat st j in Seg i holds
R1 . j <= R2 . j ) & ex j being Nat st
( j in Seg i & R1 . j < R2 . j ) holds
Sum R1 < Sum R2
proof end;

theorem Th84: :: RVSUM_1:84
for F being real-valued FinSequence st ( for i being Nat st i in dom F holds
0 <= F . i ) holds
0 <= Sum F
proof end;

theorem Th85: :: RVSUM_1:85
for F being real-valued FinSequence st ( for i being Nat st i in dom F holds
0 <= F . i ) & ex i being Nat st
( i in dom F & 0 < F . i ) holds
0 < Sum F
proof end;

theorem Th86: :: RVSUM_1:86
for F being real-valued FinSequence holds 0 <= Sum (sqr F)
proof end;

theorem Th87: :: RVSUM_1:87
for r being Real
for F being real-valued FinSequence holds Sum (r * F) = r * (Sum F)
proof end;

theorem Th88: :: RVSUM_1:88
for F being real-valued FinSequence holds Sum (- F) = - (Sum F)
proof end;

theorem Th89: :: RVSUM_1:89
for i being natural Number
for R1, R2 being Element of i -tuples_on REAL holds Sum (R1 + R2) = (Sum R1) + (Sum R2)
proof end;

theorem Th90: :: RVSUM_1:90
for i being natural Number
for R1, R2 being Element of i -tuples_on REAL holds Sum (R1 - R2) = (Sum R1) - (Sum R2)
proof end;

theorem :: RVSUM_1:91
for i being natural Number
for R being Element of i -tuples_on REAL st Sum (sqr R) = 0 holds
R = i |-> 0
proof end;

theorem :: RVSUM_1:92
for i being natural Number
for R1, R2 being Element of i -tuples_on REAL holds (Sum (mlt (R1,R2))) ^2 <= (Sum (sqr R1)) * (Sum (sqr R2))
proof end;

:: The Product of Finite Sequences of Real Numbers
definition
let F be complex-valued FinSequence;
func Product F -> Complex means :Def13: :: RVSUM_1:def 13
ex f being FinSequence of COMPLEX st
( f = F & it = multcomplex $$ f );
existence
ex b1 being Complex ex f being FinSequence of COMPLEX st
( f = F & b1 = multcomplex $$ f )
proof end;
uniqueness
for b1, b2 being Complex st ex f being FinSequence of COMPLEX st
( f = F & b1 = multcomplex $$ f ) & ex f being FinSequence of COMPLEX st
( f = F & b2 = multcomplex $$ f ) holds
b1 = b2
;
end;

:: deftheorem Def13 defines Product RVSUM_1:def 13 :
for F being complex-valued FinSequence
for b2 being Complex holds
( b2 = Product F iff ex f being FinSequence of COMPLEX st
( f = F & b2 = multcomplex $$ f ) );

registration
let F be real-valued FinSequence;
cluster Product F -> real ;
coherence
Product F is real
proof end;
end;

theorem Th93: :: RVSUM_1:93
for F being FinSequence of REAL holds Product F = multreal $$ F
proof end;

definition
let F be FinSequence of COMPLEX ;
:: original: Product
redefine func Product F -> Element of COMPLEX equals :: RVSUM_1:def 14
multcomplex $$ F;
coherence
Product F is Element of COMPLEX
by XCMPLX_0:def 2;
compatibility
for b1 being Element of COMPLEX holds
( b1 = Product F iff b1 = multcomplex $$ F )
by Def13;
end;

:: deftheorem defines Product RVSUM_1:def 14 :
for F being FinSequence of COMPLEX holds Product F = multcomplex $$ F;

definition
let F be FinSequence of REAL ;
:: original: Product
redefine func Product F -> Real equals :: RVSUM_1:def 15
multreal $$ F;
coherence
Product F is Real
;
compatibility
for b1 being Real holds
( b1 = Product F iff b1 = multreal $$ F )
by Th93;
end;

:: deftheorem defines Product RVSUM_1:def 15 :
for F being FinSequence of REAL holds Product F = multreal $$ F;

Lm3: for F being empty FinSequence holds Product F = 1
proof end;

theorem Th94: :: RVSUM_1:94
Product (<*> REAL) = 1 by Lm3;

registration
let r be Complex;
cluster <*r*> -> complex-valued ;
coherence
<*r*> is complex-valued
proof end;
end;

registration
let r1, r2 be Complex;
cluster <*r1,r2*> -> complex-valued ;
coherence
<*r1,r2*> is complex-valued
proof end;
end;

registration
let r1, r2, r3 be Complex;
cluster <*r1,r2,r3*> -> complex-valued ;
coherence
<*r1,r2,r3*> is complex-valued
proof end;
end;

registration
let j be natural Number ;
let c be Complex;
cluster j |-> c -> complex-valued ;
coherence
j |-> c is complex-valued
;
end;

theorem Th95: :: RVSUM_1:95
for r being Complex holds Product <*r*> = r
proof end;

registration
let c be Complex;
reduce Product <*c*> to c;
reducibility
Product <*c*> = c
by Th95;
end;

registration
let f, g be complex-valued FinSequence;
cluster f ^ g -> complex-valued ;
coherence
f ^ g is complex-valued
proof end;
end;

theorem Th96: :: RVSUM_1:96
for F being complex-valued FinSequence
for r being Complex holds Product (F ^ <*r*>) = (Product F) * r
proof end;

theorem Th97: :: RVSUM_1:97
for F1, F2 being complex-valued FinSequence holds Product (F1 ^ F2) = (Product F1) * (Product F2)
proof end;

theorem :: RVSUM_1:98
for r being Real
for F being real-valued FinSequence holds Product (<*r*> ^ F) = r * (Product F)
proof end;

theorem Th99: :: RVSUM_1:99
for r1, r2 being Complex holds Product <*r1,r2*> = r1 * r2
proof end;

theorem :: RVSUM_1:100
for r1, r2, r3 being Complex holds Product <*r1,r2,r3*> = (r1 * r2) * r3
proof end;

theorem :: RVSUM_1:101
for R being Element of 0 -tuples_on REAL holds Product R = 1 by Lm3;

theorem :: RVSUM_1:102
for i being natural Number holds Product (i |-> 1) = 1
proof end;

Lm4: for p being complex-valued FinSequence st len p <> 0 holds
ex q being complex-valued FinSequence ex d being Complex st p = q ^ <*d*>

proof end;

theorem :: RVSUM_1:103
for F being complex-valued FinSequence holds
( ex k being Nat st
( k in dom F & F . k = 0 ) iff Product F = 0 )
proof end;

theorem :: RVSUM_1:104
for i, j being natural Number
for r being Real holds Product ((i + j) |-> r) = (Product (i |-> r)) * (Product (j |-> r))
proof end;

theorem :: RVSUM_1:105
for i, j being natural Number
for r being Real holds Product ((i * j) |-> r) = Product (j |-> (Product (i |-> r)))
proof end;

theorem :: RVSUM_1:106
for i being natural Number
for r1, r2 being Real holds Product (i |-> (r1 * r2)) = (Product (i |-> r1)) * (Product (i |-> r2))
proof end;

theorem Th107: :: RVSUM_1:107
for i being natural Number
for R1, R2 being Element of i -tuples_on REAL holds Product (mlt (R1,R2)) = (Product R1) * (Product R2)
proof end;

theorem :: RVSUM_1:108
for i being natural Number
for r being Real
for R being Element of i -tuples_on REAL holds Product (r * R) = (Product (i |-> r)) * (Product R)
proof end;

theorem :: RVSUM_1:109
for i being natural Number
for R being Element of i -tuples_on REAL holds Product (sqr R) = (Product R) ^2 by Th107;

theorem :: RVSUM_1:110
for F being FinSequence of COMPLEX holds Product F = multcomplex $$ F ;

theorem :: RVSUM_1:111
for i, j being natural Number
for z being Element of COMPLEX holds Product ((i + j) |-> z) = (Product (i |-> z)) * (Product (j |-> z))
proof end;

theorem :: RVSUM_1:112
for i, j being natural Number
for z being Element of COMPLEX holds Product ((i * j) |-> z) = Product (j |-> (Product (i |-> z)))
proof end;

Lm5: for F being complex-valued FinSequence holds F is FinSequence of COMPLEX
proof end;

theorem :: RVSUM_1:113
for i being natural Number
for z1, z2 being Element of COMPLEX holds Product (i |-> (z1 * z2)) = (Product (i |-> z1)) * (Product (i |-> z2))
proof end;

theorem Th114: :: RVSUM_1:114
for x being real-valued FinSequence holds len (- x) = len x
proof end;

theorem Th115: :: RVSUM_1:115
for x1, x2 being real-valued FinSequence st len x1 = len x2 holds
len (x1 + x2) = len x1
proof end;

theorem Th116: :: RVSUM_1:116
for x1, x2 being real-valued FinSequence st len x1 = len x2 holds
len (x1 - x2) = len x1
proof end;

theorem Th117: :: RVSUM_1:117
for a being Real
for x being real-valued FinSequence holds len (a * x) = len x
proof end;

theorem Th118: :: RVSUM_1:118
for x, y, z being real-valued FinSequence st len x = len y & len y = len z holds
mlt ((x + y),z) = (mlt (x,z)) + (mlt (y,z))
proof end;

definition
let x1, x2 be real-valued FinSequence;
func |(x1,x2)| -> Real equals :: RVSUM_1:def 16
Sum (mlt (x1,x2));
correctness
coherence
Sum (mlt (x1,x2)) is Real
;
;
commutativity
for b1 being Real
for x1, x2 being real-valued FinSequence st b1 = Sum (mlt (x1,x2)) holds
b1 = Sum (mlt (x2,x1))
;
end;

:: deftheorem defines |( RVSUM_1:def 16 :
for x1, x2 being real-valued FinSequence holds |(x1,x2)| = Sum (mlt (x1,x2));

theorem Th119: :: RVSUM_1:119
for x being real-valued FinSequence holds |(x,x)| >= 0
proof end;

theorem Th120: :: RVSUM_1:120
for x, y, z being real-valued FinSequence st len x = len y & len y = len z holds
|((x + y),z)| = |(x,z)| + |(y,z)|
proof end;

theorem Th121: :: RVSUM_1:121
for x, y being real-valued FinSequence
for a being Real st len x = len y holds
|((a * x),y)| = a * |(x,y)|
proof end;

theorem Th122: :: RVSUM_1:122
for x1, x2 being real-valued FinSequence st len x1 = len x2 holds
|((- x1),x2)| = - |(x1,x2)|
proof end;

theorem :: RVSUM_1:123
for x1, x2 being real-valued FinSequence st len x1 = len x2 holds
|((- x1),(- x2))| = |(x1,x2)|
proof end;

theorem Th124: :: RVSUM_1:124
for x1, x2, x3 being real-valued FinSequence st len x1 = len x2 & len x2 = len x3 holds
|((x1 - x2),x3)| = |(x1,x3)| - |(x2,x3)|
proof end;

theorem :: RVSUM_1:125
for x, y being Real
for x1, x2, x3 being real-valued FinSequence st len x1 = len x2 & len x2 = len x3 holds
|(((x * x1) + (y * x2)),x3)| = (x * |(x1,x3)|) + (y * |(x2,x3)|)
proof end;

theorem Th126: :: RVSUM_1:126
for x1, x2, y1, y2 being real-valued FinSequence st len x1 = len x2 & len x2 = len y1 & len y1 = len y2 holds
|((x1 + x2),(y1 + y2))| = ((|(x1,y1)| + |(x1,y2)|) + |(x2,y1)|) + |(x2,y2)|
proof end;

theorem Th127: :: RVSUM_1:127
for x1, x2, y1, y2 being real-valued FinSequence st len x1 = len x2 & len x2 = len y1 & len y1 = len y2 holds
|((x1 - x2),(y1 - y2))| = ((|(x1,y1)| - |(x1,y2)|) - |(x2,y1)|) + |(x2,y2)|
proof end;

theorem :: RVSUM_1:128
for x, y being real-valued FinSequence st len x = len y holds
|((x + y),(x + y))| = (|(x,x)| + (2 * |(x,y)|)) + |(y,y)|
proof end;

theorem :: RVSUM_1:129
for x, y being real-valued FinSequence st len x = len y holds
|((x - y),(x - y))| = (|(x,x)| - (2 * |(x,y)|)) + |(y,y)|
proof end;

theorem Th130: :: RVSUM_1:130
for n being Nat
for p1, p2, p3 being Element of n -tuples_on REAL holds |((p1 + p2),p3)| = |(p1,p3)| + |(p2,p3)|
proof end;

theorem Th131: :: RVSUM_1:131
for n being Nat
for p1, p2 being Element of n -tuples_on REAL
for x being Real holds |((x * p1),p2)| = x * |(p1,p2)|
proof end;

theorem Th132: :: RVSUM_1:132
for n being Nat
for p1, p2 being Element of n -tuples_on REAL holds |((- p1),p2)| = - |(p1,p2)|
proof end;

theorem :: RVSUM_1:133
for n being Nat
for p1, p2 being Element of n -tuples_on REAL holds |((- p1),(- p2))| = |(p1,p2)|
proof end;

theorem Th134: :: RVSUM_1:134
for n being Nat
for p1, p2, p3 being Element of n -tuples_on REAL holds |((p1 - p2),p3)| = |(p1,p3)| - |(p2,p3)|
proof end;

theorem :: RVSUM_1:135
for n being Nat
for x, y being Real
for p1, p2, p3 being Element of n -tuples_on REAL holds |(((x * p1) + (y * p2)),p3)| = (x * |(p1,p3)|) + (y * |(p2,p3)|)
proof end;

theorem Th136: :: RVSUM_1:136
for n being Nat
for p1, p2, q1, q2 being Element of n -tuples_on REAL holds |((p1 + p2),(q1 + q2))| = ((|(p1,q1)| + |(p1,q2)|) + |(p2,q1)|) + |(p2,q2)|
proof end;

theorem Th137: :: RVSUM_1:137
for n being Nat
for p1, p2, q1, q2 being Element of n -tuples_on REAL holds |((p1 - p2),(q1 - q2))| = ((|(p1,q1)| - |(p1,q2)|) - |(p2,q1)|) + |(p2,q2)|
proof end;

theorem :: RVSUM_1:138
for n being Nat
for p, q being Element of n -tuples_on REAL holds |((p + q),(p + q))| = (|(p,p)| + (2 * |(p,q)|)) + |(q,q)|
proof end;

theorem Th139: :: RVSUM_1:139
for n being Nat
for p, q being Element of n -tuples_on REAL holds |((p - q),(p - q))| = (|(p,p)| - (2 * |(p,q)|)) + |(q,q)|
proof end;

theorem :: RVSUM_1:140
for n being Nat
for p, q being Element of n -tuples_on REAL holds |(p,q)| <= |(p,p)| + |(q,q)|
proof end;

definition
let p, q be real-valued FinSequence;
pred p,q are_orthogonal means :: RVSUM_1:def 17
|(p,q)| = 0 ;
symmetry
for p, q being real-valued FinSequence st |(p,q)| = 0 holds
|(q,p)| = 0
;
end;

:: deftheorem defines are_orthogonal RVSUM_1:def 17 :
for p, q being real-valued FinSequence holds
( p,q are_orthogonal iff |(p,q)| = 0 );

theorem :: RVSUM_1:141
for n being Nat
for a being Real
for p, q being Element of n -tuples_on REAL st p,q are_orthogonal holds
a * p,q are_orthogonal
proof end;

:: 2009.08.15, from TURING_0, A.T.
theorem :: RVSUM_1:142
for r1, r2, r3, r4 being Real holds Sum <*r1,r2,r3,r4*> = ((r1 + r2) + r3) + r4
proof end;

theorem Th143: :: RVSUM_1:143
for f being real-valued FinSequence holds
( len f = len (sqr f) & dom f = dom (sqr f) )
proof end;

theorem :: RVSUM_1:144
for f, g being real-valued FinSequence holds sqr (f ^ g) = (sqr f) ^ (sqr g)
proof end;

theorem :: RVSUM_1:145
for F being real-valued FinSequence holds F is FinSequence of REAL
proof end;

theorem :: RVSUM_1:146
for F being complex-valued FinSequence holds F is FinSequence of COMPLEX by Lm5;

registration
let r be natural Number ;
cluster <*r*> -> natural-valued ;
coherence
<*r*> is natural-valued
proof end;
end;

registration
let r1, r2 be natural Number ;
cluster <*r1,r2*> -> natural-valued ;
coherence
<*r1,r2*> is natural-valued
proof end;
end;

registration
let r1, r2, r3 be natural Number ;
cluster <*r1,r2,r3*> -> natural-valued ;
coherence
<*r1,r2,r3*> is natural-valued
proof end;
end;

registration
let f, g be natural-valued FinSequence;
cluster f ^ g -> natural-valued ;
coherence
f ^ g is natural-valued
proof end;
end;