begin
Lm1:
for K being non empty addLoopStr
for p1, p2 being FinSequence of K holds dom (p1 + p2) = (dom p1) /\ (dom p2)
theorem Th1:
theorem Th2:
theorem Th3:
theorem Th4:
theorem
:: deftheorem Def1 defines min MATRIXJ1:def 1 :
for f being FinSequence of NAT
for i being Nat st i in Seg (Sum f) holds
for b3 being Element of NAT holds
( b3 = min (f,i) iff ( i <= Sum (f | b3) & b3 in dom f & ( for j being Nat st i <= Sum (f | j) holds
b3 <= j ) ) );
theorem
theorem Th7:
theorem Th8:
theorem Th9:
Lm2:
for i being Nat
for f being FinSequence of NAT st i in dom f holds
Sum (f | i) = (Sum (f | (i -' 1))) + (f . i)
theorem Th10:
theorem
begin
:: deftheorem Def2 defines Matrix-yielding MATRIXJ1:def 2 :
for D being non empty set
for F being FinSequence of (D *) * holds
( F is Matrix-yielding iff for i being Nat st i in dom F holds
F . i is Matrix of D );
theorem Th12:
Lm3:
for D being non empty set
for M being Matrix of D holds <*M*> is Matrix-yielding
begin
:: deftheorem Def3 defines Len MATRIXJ1:def 3 :
for D being non empty set
for F being FinSequence_of_Matrix of D
for b3 being FinSequence of NAT holds
( b3 = Len F iff ( dom b3 = dom F & ( for i being Nat st i in dom b3 holds
b3 . i = len (F . i) ) ) );
:: deftheorem Def4 defines Width MATRIXJ1:def 4 :
for D being non empty set
for F being FinSequence_of_Matrix of D
for b3 being FinSequence of NAT holds
( b3 = Width F iff ( dom b3 = dom F & ( for i being Nat st i in dom b3 holds
b3 . i = width (F . i) ) ) );
theorem Th13:
theorem Th14:
theorem Th15:
Lm4:
for D being non empty set
for M being Matrix of D holds Sum (Len <*M*>) = len M
theorem Th16:
theorem Th17:
theorem Th18:
theorem Th19:
Lm5:
for D being non empty set
for M being Matrix of D holds Sum (Width <*M*>) = width M
theorem Th20:
theorem Th21:
begin
definition
let D be non
empty set ;
let d be
Element of
D;
let F be
FinSequence_of_Matrix of
D;
func block_diagonal (
F,
d)
-> Matrix of
D means :
Def5:
(
len it = Sum (Len F) &
width it = Sum (Width F) & ( for
i,
j being
Nat st
[i,j] in Indices it holds
( ( (
j <= Sum ((Width F) | ((min ((Len F),i)) -' 1)) or
j > Sum ((Width F) | (min ((Len F),i))) ) implies
it * (
i,
j)
= d ) & (
Sum ((Width F) | ((min ((Len F),i)) -' 1)) < j &
j <= Sum ((Width F) | (min ((Len F),i))) implies
it * (
i,
j)
= (F . (min ((Len F),i))) * (
(i -' (Sum ((Len F) | ((min ((Len F),i)) -' 1)))),
(j -' (Sum ((Width F) | ((min ((Len F),i)) -' 1))))) ) ) ) );
existence
ex b1 being Matrix of D st
( len b1 = Sum (Len F) & width b1 = Sum (Width F) & ( for i, j being Nat st [i,j] in Indices b1 holds
( ( ( j <= Sum ((Width F) | ((min ((Len F),i)) -' 1)) or j > Sum ((Width F) | (min ((Len F),i))) ) implies b1 * (i,j) = d ) & ( Sum ((Width F) | ((min ((Len F),i)) -' 1)) < j & j <= Sum ((Width F) | (min ((Len F),i))) implies b1 * (i,j) = (F . (min ((Len F),i))) * ((i -' (Sum ((Len F) | ((min ((Len F),i)) -' 1)))),(j -' (Sum ((Width F) | ((min ((Len F),i)) -' 1))))) ) ) ) )
uniqueness
for b1, b2 being Matrix of D st len b1 = Sum (Len F) & width b1 = Sum (Width F) & ( for i, j being Nat st [i,j] in Indices b1 holds
( ( ( j <= Sum ((Width F) | ((min ((Len F),i)) -' 1)) or j > Sum ((Width F) | (min ((Len F),i))) ) implies b1 * (i,j) = d ) & ( Sum ((Width F) | ((min ((Len F),i)) -' 1)) < j & j <= Sum ((Width F) | (min ((Len F),i))) implies b1 * (i,j) = (F . (min ((Len F),i))) * ((i -' (Sum ((Len F) | ((min ((Len F),i)) -' 1)))),(j -' (Sum ((Width F) | ((min ((Len F),i)) -' 1))))) ) ) ) & len b2 = Sum (Len F) & width b2 = Sum (Width F) & ( for i, j being Nat st [i,j] in Indices b2 holds
( ( ( j <= Sum ((Width F) | ((min ((Len F),i)) -' 1)) or j > Sum ((Width F) | (min ((Len F),i))) ) implies b2 * (i,j) = d ) & ( Sum ((Width F) | ((min ((Len F),i)) -' 1)) < j & j <= Sum ((Width F) | (min ((Len F),i))) implies b2 * (i,j) = (F . (min ((Len F),i))) * ((i -' (Sum ((Len F) | ((min ((Len F),i)) -' 1)))),(j -' (Sum ((Width F) | ((min ((Len F),i)) -' 1))))) ) ) ) holds
b1 = b2
end;
:: deftheorem Def5 defines block_diagonal MATRIXJ1:def 5 :
for D being non empty set
for d being Element of D
for F being FinSequence_of_Matrix of D
for b4 being Matrix of D holds
( b4 = block_diagonal (F,d) iff ( len b4 = Sum (Len F) & width b4 = Sum (Width F) & ( for i, j being Nat st [i,j] in Indices b4 holds
( ( ( j <= Sum ((Width F) | ((min ((Len F),i)) -' 1)) or j > Sum ((Width F) | (min ((Len F),i))) ) implies b4 * (i,j) = d ) & ( Sum ((Width F) | ((min ((Len F),i)) -' 1)) < j & j <= Sum ((Width F) | (min ((Len F),i))) implies b4 * (i,j) = (F . (min ((Len F),i))) * ((i -' (Sum ((Len F) | ((min ((Len F),i)) -' 1)))),(j -' (Sum ((Width F) | ((min ((Len F),i)) -' 1))))) ) ) ) ) );
theorem
theorem Th23:
for
D being non
empty set for
d being
Element of
D for
M1,
M2 being
Matrix of
D for
M being
Matrix of
Sum (Len <*M1,M2*>),
Sum (Width <*M1,M2*>),
D holds
(
M = block_diagonal (
<*M1,M2*>,
d) iff for
i being
Nat holds
( (
i in dom M1 implies
Line (
M,
i)
= (Line (M1,i)) ^ ((width M2) |-> d) ) & (
i in dom M2 implies
Line (
M,
(i + (len M1)))
= ((width M1) |-> d) ^ (Line (M2,i)) ) ) )
theorem Th24:
for
D being non
empty set for
d being
Element of
D for
M1,
M2 being
Matrix of
D for
M being
Matrix of
Sum (Len <*M1,M2*>),
Sum (Width <*M1,M2*>),
D holds
(
M = block_diagonal (
<*M1,M2*>,
d) iff for
i being
Nat holds
( (
i in Seg (width M1) implies
Col (
M,
i)
= (Col (M1,i)) ^ ((len M2) |-> d) ) & (
i in Seg (width M2) implies
Col (
M,
(i + (width M1)))
= ((len M1) |-> d) ^ (Col (M2,i)) ) ) )
theorem Th25:
theorem Th26:
theorem Th27:
theorem Th28:
theorem Th29:
theorem Th30:
theorem Th31:
theorem Th32:
theorem Th33:
theorem Th34:
theorem Th35:
theorem Th36:
theorem
theorem
theorem
theorem Th40:
theorem Th41:
theorem Th42:
theorem Th43:
theorem Th44:
:: deftheorem Def6 defines Square-Matrix-yielding MATRIXJ1:def 6 :
for D being non empty set
for F being FinSequence of (D *) * holds
( F is Square-Matrix-yielding iff for i being Nat st i in dom F holds
ex n being Nat st F . i is Matrix of n,D );
theorem
Lm6:
for n being Nat
for D being non empty set
for M being Matrix of n,D holds <*M*> is FinSequence_of_Square-Matrix of D
theorem Th46:
theorem Th47:
theorem
:: deftheorem Def7 defines Det MATRIXJ1:def 7 :
for K being Field
for R being FinSequence_of_Square-Matrix of K
for b3 being FinSequence of K holds
( b3 = Det R iff ( dom b3 = dom R & ( for i being Nat st i in dom b3 holds
b3 . i = Det (R . i) ) ) );
theorem Th49:
theorem Th50:
theorem
theorem Th52:
theorem
theorem Th54:
theorem
begin
:: deftheorem Def8 defines 1. MATRIXJ1:def 8 :
for K being Field
for f being FinSequence of NAT
for b3 being FinSequence_of_Square-Matrix of K holds
( b3 = 1. (K,f) iff ( dom b3 = dom f & ( for i being Nat st i in dom b3 holds
b3 . i = 1. (K,(f . i)) ) ) );
theorem
theorem Th57:
theorem Th58:
theorem
theorem Th60:
theorem
begin
:: deftheorem Def9 defines mlt MATRIXJ1:def 9 :
for K being Field
for G being FinSequence_of_Matrix of K
for p being FinSequence of K
for b4 being FinSequence_of_Matrix of K holds
( b4 = mlt (p,G) iff ( dom b4 = dom G & ( for i being Nat st i in dom b4 holds
b4 . i = (p /. i) * (G . i) ) ) );
theorem
theorem Th63:
theorem Th64:
Lm7:
for K being Field
for a, a1 being Element of K
for A1, A2 being Matrix of K holds a * (block_diagonal (<*A1,A2*>,a1)) = block_diagonal (<*(a * A1),(a * A2)*>,(a * a1))
theorem
:: deftheorem Def10 defines (+) MATRIXJ1:def 10 :
for K being Field
for G1, G2, b4 being FinSequence_of_Matrix of K holds
( b4 = G1 (+) G2 iff ( dom b4 = dom G1 & ( for i being Nat st i in dom b4 holds
b4 . i = (G1 . i) + (G2 . i) ) ) );
theorem Th66:
theorem Th67:
theorem Th68:
theorem Th69:
theorem Th70:
Lm8:
for i being Nat
for K being Field
for A1, A2 being Matrix of K st width A1 = width A2 & i in dom A1 holds
Line ((A1 + A2),i) = (Line (A1,i)) + (Line (A2,i))
theorem Th71:
for
K being
Field for
a1,
a2 being
Element of
K for
A1,
B1,
A2,
B2 being
Matrix of
K st
len A1 = len B1 &
len A2 = len B2 &
width A1 = width B1 &
width A2 = width B2 holds
(block_diagonal (<*A1,A2*>,a1)) + (block_diagonal (<*B1,B2*>,a2)) = block_diagonal (
(<*A1,A2*> (+) <*B1,B2*>),
(a1 + a2))
theorem
:: deftheorem Def11 defines (#) MATRIXJ1:def 11 :
for K being Field
for G1, G2, b4 being FinSequence_of_Matrix of K holds
( b4 = G1 (#) G2 iff ( dom b4 = dom G1 & ( for i being Nat st i in dom b4 holds
b4 . i = (G1 . i) * (G2 . i) ) ) );
theorem Th73:
theorem Th74:
theorem Th75:
theorem Th76:
theorem Th77:
Lm9:
for i, j being Nat
for K being Field
for R1, S1 being Element of i -tuples_on the carrier of K
for R2, S2 being Element of j -tuples_on the carrier of K holds Sum (mlt ((R1 ^ R2),(S1 ^ S2))) = (Sum (mlt (R1,S1))) + (Sum (mlt (R2,S2)))
theorem Th78:
for
K being
Field for
A1,
B1,
A2,
B2 being
Matrix of
K st
width A1 = len B1 &
width A2 = len B2 holds
(block_diagonal (<*A1,A2*>,(0. K))) * (block_diagonal (<*B1,B2*>,(0. K))) = block_diagonal (
(<*A1,A2*> (#) <*B1,B2*>),
(0. K))
theorem