begin
:: deftheorem Def1 defines C_Linear_Combination CONVEX4:def 1 :
for V being non empty 1-sorted
for b2 being Element of Funcs ( the carrier of V,COMPLEX) holds
( b2 is C_Linear_Combination of V iff ex T being finite Subset of V st
for v being Element of V st not v in T holds
b2 . v = 0 );
:: deftheorem defines Carrier CONVEX4:def 2 :
for V being non empty addLoopStr
for L being Element of Funcs ( the carrier of V,COMPLEX) holds Carrier L = { v where v is Element of V : L . v <> 0c } ;
theorem Th1:
:: deftheorem Def3 defines ZeroCLC CONVEX4:def 3 :
for V being non empty addLoopStr
for b2 being C_Linear_Combination of V holds
( b2 = ZeroCLC V iff Carrier b2 = {} );
theorem Th2:
:: deftheorem Def4 defines C_Linear_Combination CONVEX4:def 4 :
for V being non empty addLoopStr
for A being Subset of V
for b3 being C_Linear_Combination of V holds
( b3 is C_Linear_Combination of A iff Carrier b3 c= A );
theorem
theorem Th4:
theorem Th5:
:: deftheorem Def5 defines (#) CONVEX4:def 5 :
for V being non empty CLSStruct
for F being FinSequence of the carrier of V
for f being Function of the carrier of V,COMPLEX
for b4 being FinSequence of the carrier of V holds
( b4 = f (#) F iff ( len b4 = len F & ( for i being Nat st i in dom b4 holds
b4 . i = (f . (F /. i)) * (F /. i) ) ) );
theorem Th6:
theorem
theorem Th8:
theorem Th9:
theorem Th10:
:: deftheorem Def6 defines Sum CONVEX4:def 6 :
for V being non empty right_complementable Abelian add-associative right_zeroed CLSStruct
for L being C_Linear_Combination of V
for b3 being Element of V holds
( b3 = Sum L iff ex F being FinSequence of the carrier of V st
( F is one-to-one & rng F = Carrier L & b3 = Sum (L (#) F) ) );
theorem Th11:
theorem
theorem
theorem Th14:
theorem Th15:
theorem
theorem
theorem Th18:
:: deftheorem defines = CONVEX4:def 7 :
for V being non empty addLoopStr
for L1, L2 being C_Linear_Combination of V holds
( L1 = L2 iff for v being Element of V holds L1 . v = L2 . v );
:: deftheorem Def8 defines + CONVEX4:def 8 :
for V being non empty addLoopStr
for L1, L2, b4 being C_Linear_Combination of V holds
( b4 = L1 + L2 iff for v being Element of V holds b4 . v = (L1 . v) + (L2 . v) );
theorem Th19:
theorem Th20:
theorem
theorem Th22:
theorem Th23:
:: deftheorem Def9 defines * CONVEX4:def 9 :
for V being non empty CLSStruct
for a being Complex
for L, b4 being C_Linear_Combination of V holds
( b4 = a * L iff for v being VECTOR of V holds b4 . v = a * (L . v) );
theorem Th24:
theorem Th25:
theorem Th26:
theorem Th27:
theorem Th28:
theorem Th29:
theorem Th30:
:: deftheorem defines - CONVEX4:def 10 :
for V being non empty CLSStruct
for L being C_Linear_Combination of V holds - L = (- 1r) * L;
theorem Th31:
theorem
theorem
:: deftheorem defines - CONVEX4:def 11 :
for V being non empty CLSStruct
for L1, L2 being C_Linear_Combination of V holds L1 - L2 = L1 + (- L2);
theorem Th34:
theorem
theorem
theorem Th37:
:: deftheorem Def12 defines C_LinComb CONVEX4:def 12 :
for V being non empty CLSStruct
for b2 being set holds
( b2 = C_LinComb V iff for x being set holds
( x in b2 iff x is C_Linear_Combination of V ) );
:: deftheorem defines @ CONVEX4:def 13 :
for V being non empty CLSStruct
for e being Element of C_LinComb V holds @ e = e;
:: deftheorem defines @ CONVEX4:def 14 :
for V being non empty CLSStruct
for L being C_Linear_Combination of V holds @ L = L;
:: deftheorem Def15 defines C_LCAdd CONVEX4:def 15 :
for V being non empty CLSStruct
for b2 being BinOp of (C_LinComb V) holds
( b2 = C_LCAdd V iff for e1, e2 being Element of C_LinComb V holds b2 . (e1,e2) = (@ e1) + (@ e2) );
definition
let V be non
empty CLSStruct ;
func C_LCMult V -> Function of
[:COMPLEX,(C_LinComb V):],
(C_LinComb V) means :
Def16:
for
a being
Complex for
e being
Element of
C_LinComb V holds
it . [a,e] = a * (@ e);
existence
ex b1 being Function of [:COMPLEX,(C_LinComb V):],(C_LinComb V) st
for a being Complex
for e being Element of C_LinComb V holds b1 . [a,e] = a * (@ e)
uniqueness
for b1, b2 being Function of [:COMPLEX,(C_LinComb V):],(C_LinComb V) st ( for a being Complex
for e being Element of C_LinComb V holds b1 . [a,e] = a * (@ e) ) & ( for a being Complex
for e being Element of C_LinComb V holds b2 . [a,e] = a * (@ e) ) holds
b1 = b2
end;
:: deftheorem Def16 defines C_LCMult CONVEX4:def 16 :
for V being non empty CLSStruct
for b2 being Function of [:COMPLEX,(C_LinComb V):],(C_LinComb V) holds
( b2 = C_LCMult V iff for a being Complex
for e being Element of C_LinComb V holds b2 . [a,e] = a * (@ e) );
:: deftheorem defines LC_CLSpace CONVEX4:def 17 :
for V being non empty CLSStruct holds LC_CLSpace V = CLSStruct(# (C_LinComb V),(@ (ZeroCLC V)),(C_LCAdd V),(C_LCMult V) #);
theorem Th38:
theorem Th39:
theorem Th40:
theorem
:: deftheorem defines LC_CLSpace CONVEX4:def 18 :
for V being non empty CLSStruct
for A being Subset of V
for b3 being strict Subspace of LC_CLSpace V holds
( b3 = LC_CLSpace A iff the carrier of b3 = { l where l is C_Linear_Combination of A : verum } );
begin
:: deftheorem defines Up CONVEX4:def 19 :
for V being ComplexLinearSpace
for W being Subspace of V holds Up W = the carrier of W;
:: deftheorem Def20 defines Affine CONVEX4:def 20 :
for V being non empty CLSStruct
for S being Subset of V holds
( S is Affine iff for x, y being VECTOR of V
for z being Complex st z is Real & x in S & y in S holds
((1r - z) * x) + (z * y) in S );
:: deftheorem defines (Omega). CONVEX4:def 21 :
for V being ComplexLinearSpace holds (Omega). V = CLSStruct(# the carrier of V, the ZeroF of V, the addF of V, the Mult of V #);
theorem Th42:
theorem Th43:
theorem Th44:
begin
:: deftheorem defines * CONVEX4:def 22 :
for V being non empty CLSStruct
for M being Subset of V
for r being Element of COMPLEX holds r * M = { (r * v) where v is Element of V : v in M } ;
:: deftheorem Def23 defines convex CONVEX4:def 23 :
for V being non empty CLSStruct
for M being Subset of V holds
( M is convex iff for u, v being VECTOR of V
for z being Complex st ex r being Real st
( z = r & 0 < r & r < 1 ) & u in M & v in M holds
(z * u) + ((1r - z) * v) in M );
theorem
theorem
theorem
theorem Th48:
theorem
theorem
theorem Th51:
theorem Th52:
theorem
theorem Th54:
theorem Th55:
theorem
theorem
theorem
theorem Th59:
theorem Th60:
theorem Th61:
theorem Th62:
theorem Th63:
theorem Th64:
theorem Th65:
Lm2:
for V being ComplexLinearSpace
for M being Subset of V
for z1, z2 being Complex st ex r1, r2 being Real st
( z1 = r1 & z2 = r2 & r1 >= 0 & r2 >= 0 ) & M is convex holds
(z1 * M) + (z2 * M) c= (z1 + z2) * M
theorem
theorem
theorem Th68:
theorem Th69:
theorem
theorem
theorem
theorem
theorem
theorem
theorem
theorem
theorem
theorem
begin
:: deftheorem Def24 defines convex CONVEX4:def 24 :
for V being ComplexLinearSpace
for L being C_Linear_Combination of V holds
( L is convex iff ex F being FinSequence of the carrier of V st
( F is one-to-one & rng F = Carrier L & ex f being FinSequence of REAL st
( len f = len F & Sum f = 1 & ( for n being Nat st n in dom f holds
( f . n = L . (F . n) & f . n >= 0 ) ) ) ) );
theorem Th80:
theorem
theorem
theorem Th83:
theorem Th84:
Lm3:
for V being ComplexLinearSpace
for v1, v2, v3 being VECTOR of V
for L being C_Linear_Combination of {v1,v2,v3} st v1 <> v2 & v2 <> v3 & v3 <> v1 holds
Sum L = (((L . v1) * v1) + ((L . v2) * v2)) + ((L . v3) * v3)
theorem Th85:
theorem
theorem
theorem
begin
:: deftheorem Def25 defines Convex-Family CONVEX4:def 25 :
for V being non empty CLSStruct
for M being Subset of V
for b3 being Subset-Family of V holds
( b3 = Convex-Family M iff for N being Subset of V holds
( N in b3 iff ( N is convex & M c= N ) ) );
:: deftheorem defines conv CONVEX4:def 26 :
for V being non empty CLSStruct
for M being Subset of V holds conv M = meet (Convex-Family M);
theorem