begin
theorem
canceled;
theorem
canceled;
theorem Th3:
theorem Th4:
theorem Th5:
theorem Th6:
begin
theorem
canceled;
theorem
canceled;
begin
theorem Th9:
theorem Th10:
theorem Th11:
theorem Th12:
:: deftheorem MATRLIN:def 1 :
canceled;
:: deftheorem MATRLIN:def 2 :
canceled;
:: deftheorem Def3 defines finite-dimensional MATRLIN:def 3 :
for K being Field
for V being VectSp of K holds
( V is finite-dimensional iff ex A being finite Subset of V st A is Basis of V );
:: deftheorem Def4 defines OrdBasis MATRLIN:def 4 :
for K being Field
for V being finite-dimensional VectSp of K
for b3 being FinSequence of V holds
( b3 is OrdBasis of V iff ( b3 is one-to-one & rng b3 is Basis of V ) );
:: deftheorem Def5 defines + MATRLIN:def 5 :
for K being non empty doubleLoopStr
for V1, V2 being non empty VectSpStr of K
for f1, f2, b6 being Function of V1,V2 holds
( b6 = f1 + f2 iff for v being Element of V1 holds b6 . v = (f1 . v) + (f2 . v) );
:: deftheorem Def6 defines * MATRLIN:def 6 :
for K being non empty doubleLoopStr
for V1, V2 being non empty VectSpStr of K
for f being Function of V1,V2
for a being Element of K
for b6 being Function of V1,V2 holds
( b6 = a * f iff for v being Element of V1 holds b6 . v = a * (f . v) );
theorem Th13:
theorem Th14:
theorem Th15:
:: deftheorem defines lmlt MATRLIN:def 7 :
for K being Field
for V1 being finite-dimensional VectSp of K
for p1 being FinSequence of K
for p2 being FinSequence of V1 holds lmlt (p1,p2) = the lmult of V1 .: (p1,p2);
theorem Th16:
:: deftheorem Def8 defines Sum MATRLIN:def 8 :
for V1 being non empty addLoopStr
for M being FinSequence of the carrier of V1 *
for b3 being FinSequence of V1 holds
( b3 = Sum M iff ( len b3 = len M & ( for k being Nat st k in dom b3 holds
b3 /. k = Sum (M /. k) ) ) );
theorem Th17:
theorem Th18:
theorem Th19:
theorem Th20:
theorem Th21:
theorem Th22:
theorem Th23:
theorem Th24:
theorem Th25:
theorem Th26:
definition
let D be non
empty set ;
let n,
m,
k be
Nat;
let M1 be
Matrix of
n,
k,
D;
let M2 be
Matrix of
m,
k,
D;
^redefine func M1 ^ M2 -> Matrix of
n + m,
k,
D;
coherence
M1 ^ M2 is Matrix of n + m,k,D
end;
theorem
theorem Th28:
theorem
for
t,
k,
m,
n,
i being
Nat for
D being non
empty set for
M1 being
Matrix of
t,
k,
D for
M2 being
Matrix of
m,
k,
D st
n in dom M2 &
i = (len M1) + n holds
Line (
(M1 ^ M2),
i)
= Line (
M2,
n)
theorem Th30:
theorem Th31:
theorem Th32:
theorem Th33:
theorem Th34:
theorem Th35:
theorem
canceled;
theorem Th37:
theorem Th38:
for
n,
m being
Nat for
K being
Field for
V1 being
finite-dimensional VectSp of
K for
M being
Matrix of
n,
m, the
carrier of
K st
n > 0 &
m > 0 holds
for
p,
d being
FinSequence of
K st
len p = n &
len d = m & ( for
j being
Nat st
j in dom d holds
d /. j = Sum (mlt (p,(Col (M,j)))) ) holds
for
b,
c being
FinSequence of
V1 st
len b = m &
len c = n & ( for
i being
Nat st
i in dom c holds
c /. i = Sum (lmlt ((Line (M,i)),b)) ) holds
Sum (lmlt (p,c)) = Sum (lmlt (d,b))
begin
:: deftheorem Def9 defines |-- MATRLIN:def 9 :
for K being Field
for V being finite-dimensional VectSp of K
for b1 being OrdBasis of V
for W being Element of V
for b5 being FinSequence of K holds
( b5 = W |-- b1 iff ( len b5 = len b1 & ex KL being Linear_Combination of V st
( W = Sum KL & Carrier KL c= rng b1 & ( for k being Nat st 1 <= k & k <= len b5 holds
b5 /. k = KL . (b1 /. k) ) ) ) );
Lm1:
for K being Field
for V being finite-dimensional VectSp of K
for b being OrdBasis of V
for W being Element of V holds dom (W |-- b) = dom b
theorem Th39:
theorem Th40:
theorem Th41:
Lm2:
for p being FinSequence
for k being set st k in dom p holds
len p > 0
theorem Th42:
begin
:: deftheorem Def10 defines AutMt MATRLIN:def 10 :
for K being Field
for V1, V2 being finite-dimensional VectSp of K
for f being Function of V1,V2
for b1 being FinSequence of V1
for b2 being OrdBasis of V2
for b7 being Matrix of K holds
( b7 = AutMt (f,b1,b2) iff ( len b7 = len b1 & ( for k being Nat st k in dom b1 holds
b7 /. k = (f . (b1 /. k)) |-- b2 ) ) );
Lm3:
for K being Field
for V1, V2 being finite-dimensional VectSp of K
for f being Function of V1,V2
for b1 being OrdBasis of V1
for b2 being OrdBasis of V2 holds dom (AutMt (f,b1,b2)) = dom b1
theorem Th43:
theorem Th44:
theorem
theorem
for
K being
Field for
V1,
V2,
V3 being
finite-dimensional VectSp of
K for
f being
Function of
V1,
V2 for
g being
Function of
V2,
V3 for
b1 being
OrdBasis of
V1 for
b2 being
OrdBasis of
V2 for
b3 being
OrdBasis of
V3 st
g is
additive &
g is
homogeneous &
len b1 > 0 &
len b2 > 0 holds
AutMt (
(g * f),
b1,
b3)
= (AutMt (f,b1,b2)) * (AutMt (g,b2,b3))
theorem
theorem