begin
:: deftheorem Def1 defines vector RLVECT_2:def 1 :
for S being 1-sorted
for x being set st x in S holds
vector (S,x) = x;
theorem
canceled;
theorem
canceled;
theorem
theorem Th4:
theorem
theorem Th6:
theorem
theorem Th8:
theorem
:: deftheorem RLVECT_2:def 2 :
canceled;
:: deftheorem RLVECT_2:def 3 :
canceled;
:: deftheorem Def4 defines Sum RLVECT_2:def 4 :
for V being non empty addLoopStr
for T being finite Subset of V st V is Abelian & V is add-associative & V is right_zeroed holds
for b3 being Element of V holds
( b3 = Sum T iff ex F being FinSequence of the carrier of V st
( rng F = T & F is one-to-one & b3 = Sum F ) );
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem Th14:
theorem
theorem
theorem
theorem Th18:
theorem Th19:
theorem
theorem Th21:
theorem Th22:
theorem
theorem
:: deftheorem Def5 defines Linear_Combination RLVECT_2:def 5 :
for V being non empty ZeroStr
for b2 being Element of Funcs ( the carrier of V,REAL) holds
( b2 is 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 RLVECT_2:def 6 :
for V being non empty addLoopStr
for L being Element of Funcs ( the carrier of V,REAL) holds Carrier L = { v where v is Element of V : L . v <> 0 } ;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
:: deftheorem Def7 defines ZeroLC RLVECT_2:def 7 :
for V being non empty addLoopStr
for b2 being Linear_Combination of V holds
( b2 = ZeroLC V iff Carrier b2 = {} );
theorem
canceled;
theorem Th30:
:: deftheorem Def8 defines Linear_Combination RLVECT_2:def 8 :
for V being non empty addLoopStr
for A being Subset of V
for b3 being Linear_Combination of V holds
( b3 is Linear_Combination of A iff Carrier b3 c= A );
theorem
canceled;
theorem
canceled;
theorem
theorem Th34:
theorem Th35:
:: deftheorem Def9 defines (#) RLVECT_2:def 9 :
for V being RealLinearSpace
for F being FinSequence of V
for f being Function of the carrier of V,REAL
for b4 being FinSequence of the carrier of V holds
( b4 = f (#) F iff ( len b4 = len F & ( for i being Element of NAT st i in dom b4 holds
b4 . i = (f . (F /. i)) * (F /. i) ) ) );
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem Th40:
theorem
theorem Th42:
theorem Th43:
theorem
:: deftheorem Def10 defines Sum RLVECT_2:def 10 :
for V being RealLinearSpace
for L being Linear_Combination of V
for b3 being Element of V holds
( b3 = Sum L iff ex F being FinSequence of V st
( F is one-to-one & rng F = Carrier L & b3 = Sum (L (#) F) ) );
Lm2:
for V being RealLinearSpace holds Sum (ZeroLC V) = 0. V
theorem
canceled;
theorem
canceled;
theorem
theorem
theorem
theorem Th50:
theorem Th51:
theorem
theorem
theorem
:: deftheorem defines = RLVECT_2:def 11 :
for V being non empty addLoopStr
for L1, L2 being Linear_Combination of V holds
( L1 = L2 iff for v being Element of V holds L1 . v = L2 . v );
:: deftheorem Def12 defines + RLVECT_2:def 12 :
for V being non empty addLoopStr
for L1, L2, b4 being Linear_Combination of V holds
( b4 = L1 + L2 iff for v being Element of V holds b4 . v = (L1 . v) + (L2 . v) );
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem Th58:
theorem Th59:
theorem
theorem Th61:
theorem Th62:
:: deftheorem Def13 defines * RLVECT_2:def 13 :
for V being RealLinearSpace
for a being Real
for L, b4 being Linear_Combination of V holds
( b4 = a * L iff for v being VECTOR of V holds b4 . v = a * (L . v) );
theorem
canceled;
theorem
canceled;
theorem Th65:
theorem Th66:
theorem Th67:
theorem Th68:
theorem Th69:
theorem Th70:
theorem Th71:
:: deftheorem defines - RLVECT_2:def 14 :
for V being RealLinearSpace
for L being Linear_Combination of V holds - L = (- 1) * L;
theorem
canceled;
theorem Th73:
theorem
theorem
theorem
theorem
:: deftheorem defines - RLVECT_2:def 15 :
for V being RealLinearSpace
for L1, L2 being Linear_Combination of V holds L1 - L2 = L1 + (- L2);
theorem
canceled;
theorem Th79:
theorem
theorem
theorem Th82:
:: deftheorem Def16 defines LinComb RLVECT_2:def 16 :
for V being RealLinearSpace
for b2 being set holds
( b2 = LinComb V iff for x being set holds
( x in b2 iff x is Linear_Combination of V ) );
:: deftheorem defines @ RLVECT_2:def 17 :
for V being RealLinearSpace
for e being Element of LinComb V holds @ e = e;
:: deftheorem defines @ RLVECT_2:def 18 :
for V being RealLinearSpace
for L being Linear_Combination of V holds @ L = L;
:: deftheorem Def19 defines LCAdd RLVECT_2:def 19 :
for V being RealLinearSpace
for b2 being BinOp of (LinComb V) holds
( b2 = LCAdd V iff for e1, e2 being Element of LinComb V holds b2 . (e1,e2) = (@ e1) + (@ e2) );
definition
let V be
RealLinearSpace;
func LCMult V -> Function of
[:REAL,(LinComb V):],
(LinComb V) means :
Def20:
for
a being
Real for
e being
Element of
LinComb V holds
it . [a,e] = a * (@ e);
existence
ex b1 being Function of [:REAL,(LinComb V):],(LinComb V) st
for a being Real
for e being Element of LinComb V holds b1 . [a,e] = a * (@ e)
uniqueness
for b1, b2 being Function of [:REAL,(LinComb V):],(LinComb V) st ( for a being Real
for e being Element of LinComb V holds b1 . [a,e] = a * (@ e) ) & ( for a being Real
for e being Element of LinComb V holds b2 . [a,e] = a * (@ e) ) holds
b1 = b2
end;
:: deftheorem Def20 defines LCMult RLVECT_2:def 20 :
for V being RealLinearSpace
for b2 being Function of [:REAL,(LinComb V):],(LinComb V) holds
( b2 = LCMult V iff for a being Real
for e being Element of LinComb V holds b2 . [a,e] = a * (@ e) );
:: deftheorem defines LC_RLSpace RLVECT_2:def 21 :
for V being RealLinearSpace holds LC_RLSpace V = RLSStruct(# (LinComb V),(@ (ZeroLC V)),(LCAdd V),(LCMult V) #);
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
theorem
theorem
theorem
theorem Th96:
theorem Th97:
theorem Th98:
theorem
:: deftheorem defines LC_RLSpace RLVECT_2:def 22 :
for V being RealLinearSpace
for A being Subset of V
for b3 being strict Subspace of LC_RLSpace V holds
( b3 = LC_RLSpace A iff the carrier of b3 = { l where l is Linear_Combination of A : verum } );
theorem Th100:
theorem
theorem
theorem
theorem
theorem