begin
:: deftheorem Def1 defines tabular MATRIX_1:def 1 :
for f being FinSequence holds
( f is tabular iff ex n being Nat st
for x being set st x in rng f holds
ex s being FinSequence st
( s = x & len s = n ) );
theorem Th1:
theorem Th2:
theorem Th3:
theorem Th4:
theorem Th5:
theorem
theorem
theorem
theorem Th9:
:: deftheorem MATRIX_1:def 2 :
canceled;
:: deftheorem Def3 defines Matrix MATRIX_1:def 3 :
for D being non empty set
for m, n being Nat
for b4 being Matrix of D holds
( b4 is Matrix of m,n,D iff ( len b4 = m & ( for p being FinSequence of D st p in rng b4 holds
len p = n ) ) );
theorem Th10:
theorem Th11:
theorem Th12:
theorem Th13:
theorem
theorem
theorem
theorem
theorem
theorem
:: deftheorem Def4 defines width MATRIX_1:def 4 :
for M being tabular FinSequence
for b2 being Element of NAT holds
( ( len M > 0 implies ( b2 = width M iff ex s being FinSequence st
( s in rng M & len s = b2 ) ) ) & ( not len M > 0 implies ( b2 = width M iff b2 = 0 ) ) );
theorem Th20:
:: deftheorem defines Indices MATRIX_1:def 5 :
for M being tabular FinSequence holds Indices M = [:(dom M),(Seg (width M)):];
:: deftheorem Def6 defines * MATRIX_1:def 6 :
for D being set
for M being Matrix of D
for i, j being Nat st [i,j] in Indices M holds
for b5 being Element of D holds
( b5 = M * (i,j) iff ex p being FinSequence of D st
( p = M . i & b5 = p . j ) );
theorem Th21:
scheme
MatrixEx{
F1()
-> non
empty set ,
F2()
-> Element of
NAT ,
F3()
-> Element of
NAT ,
P1[
set ,
set ,
set ] } :
provided
for
i,
j being
Nat st
[i,j] in [:(Seg F2()),(Seg F3()):] holds
for
x1,
x2 being
Element of
F1() st
P1[
i,
j,
x1] &
P1[
i,
j,
x2] holds
x1 = x2
and A1:
for
i,
j being
Nat st
[i,j] in [:(Seg F2()),(Seg F3()):] holds
ex
x being
Element of
F1() st
P1[
i,
j,
x]
theorem
canceled;
theorem
theorem Th24:
theorem Th25:
theorem Th26:
theorem Th27:
theorem Th28:
theorem
definition
let D be non
empty set ;
let M be
Matrix of
D;
func M @ -> Matrix of
D means :
Def7:
(
len it = width M & ( for
i,
j being
Nat holds
(
[i,j] in Indices it iff
[j,i] in Indices M ) ) & ( for
i,
j being
Nat st
[j,i] in Indices M holds
it * (
i,
j)
= M * (
j,
i) ) );
existence
ex b1 being Matrix of D st
( len b1 = width M & ( for i, j being Nat holds
( [i,j] in Indices b1 iff [j,i] in Indices M ) ) & ( for i, j being Nat st [j,i] in Indices M holds
b1 * (i,j) = M * (j,i) ) )
uniqueness
for b1, b2 being Matrix of D st len b1 = width M & ( for i, j being Nat holds
( [i,j] in Indices b1 iff [j,i] in Indices M ) ) & ( for i, j being Nat st [j,i] in Indices M holds
b1 * (i,j) = M * (j,i) ) & len b2 = width M & ( for i, j being Nat holds
( [i,j] in Indices b2 iff [j,i] in Indices M ) ) & ( for i, j being Nat st [j,i] in Indices M holds
b2 * (i,j) = M * (j,i) ) holds
b1 = b2
end;
:: deftheorem Def7 defines @ MATRIX_1:def 7 :
for D being non empty set
for M, b3 being Matrix of D holds
( b3 = M @ iff ( len b3 = width M & ( for i, j being Nat holds
( [i,j] in Indices b3 iff [j,i] in Indices M ) ) & ( for i, j being Nat st [j,i] in Indices M holds
b3 * (i,j) = M * (j,i) ) ) );
Lm1:
for D being non empty set
for M being Matrix of D st width M > 0 holds
( len (M @) = width M & width (M @) = len M )
Lm2:
for n being Element of NAT
for K being non empty set
for A being Matrix of n,K holds A @ is Matrix of n,K
definition
let D be non
empty set ;
let M be
Matrix of
D;
let i be
Nat;
func Line (
M,
i)
-> FinSequence of
D means :
Def8:
(
len it = width M & ( for
j being
Nat st
j in Seg (width M) holds
it . j = M * (
i,
j) ) );
existence
ex b1 being FinSequence of D st
( len b1 = width M & ( for j being Nat st j in Seg (width M) holds
b1 . j = M * (i,j) ) )
uniqueness
for b1, b2 being FinSequence of D st len b1 = width M & ( for j being Nat st j in Seg (width M) holds
b1 . j = M * (i,j) ) & len b2 = width M & ( for j being Nat st j in Seg (width M) holds
b2 . j = M * (i,j) ) holds
b1 = b2
func Col (
M,
i)
-> FinSequence of
D means :
Def9:
(
len it = len M & ( for
j being
Nat st
j in dom M holds
it . j = M * (
j,
i) ) );
existence
ex b1 being FinSequence of D st
( len b1 = len M & ( for j being Nat st j in dom M holds
b1 . j = M * (j,i) ) )
uniqueness
for b1, b2 being FinSequence of D st len b1 = len M & ( for j being Nat st j in dom M holds
b1 . j = M * (j,i) ) & len b2 = len M & ( for j being Nat st j in dom M holds
b2 . j = M * (j,i) ) holds
b1 = b2
end;
:: deftheorem Def8 defines Line MATRIX_1:def 8 :
for D being non empty set
for M being Matrix of D
for i being Nat
for b4 being FinSequence of D holds
( b4 = Line (M,i) iff ( len b4 = width M & ( for j being Nat st j in Seg (width M) holds
b4 . j = M * (i,j) ) ) );
:: deftheorem Def9 defines Col MATRIX_1:def 9 :
for D being non empty set
for M being Matrix of D
for i being Nat
for b4 being FinSequence of D holds
( b4 = Col (M,i) iff ( len b4 = len M & ( for j being Nat st j in dom M holds
b4 . j = M * (j,i) ) ) );
definition
let K be non
empty doubleLoopStr ;
let n be
Nat;
func n -Matrices_over K -> set equals
n -tuples_on (n -tuples_on the carrier of K);
coherence
n -tuples_on (n -tuples_on the carrier of K) is set
;
func 0. (
K,
n)
-> Matrix of
n,
K equals
n |-> (n |-> (0. K));
coherence
n |-> (n |-> (0. K)) is Matrix of n,K
by Th10;
func 1. (
K,
n)
-> Matrix of
n,
K means :
Def12:
( ( for
i being
Nat st
[i,i] in Indices it holds
it * (
i,
i)
= 1. K ) & ( for
i,
j being
Nat st
[i,j] in Indices it &
i <> j holds
it * (
i,
j)
= 0. K ) );
existence
ex b1 being Matrix of n,K st
( ( for i being Nat st [i,i] in Indices b1 holds
b1 * (i,i) = 1. K ) & ( for i, j being Nat st [i,j] in Indices b1 & i <> j holds
b1 * (i,j) = 0. K ) )
uniqueness
for b1, b2 being Matrix of n,K st ( for i being Nat st [i,i] in Indices b1 holds
b1 * (i,i) = 1. K ) & ( for i, j being Nat st [i,j] in Indices b1 & i <> j holds
b1 * (i,j) = 0. K ) & ( for i being Nat st [i,i] in Indices b2 holds
b2 * (i,i) = 1. K ) & ( for i, j being Nat st [i,j] in Indices b2 & i <> j holds
b2 * (i,j) = 0. K ) holds
b1 = b2
let A be
Matrix of
n,
K;
func - A -> Matrix of
n,
K means :
Def13:
for
i,
j being
Nat st
[i,j] in Indices A holds
it * (
i,
j)
= - (A * (i,j));
existence
ex b1 being Matrix of n,K st
for i, j being Nat st [i,j] in Indices A holds
b1 * (i,j) = - (A * (i,j))
uniqueness
for b1, b2 being Matrix of n,K st ( for i, j being Nat st [i,j] in Indices A holds
b1 * (i,j) = - (A * (i,j)) ) & ( for i, j being Nat st [i,j] in Indices A holds
b2 * (i,j) = - (A * (i,j)) ) holds
b1 = b2
let B be
Matrix of
n,
K;
func A + B -> Matrix of
n,
K means :
Def14:
for
i,
j being
Nat st
[i,j] in Indices A holds
it * (
i,
j)
= (A * (i,j)) + (B * (i,j));
existence
ex b1 being Matrix of n,K st
for i, j being Nat st [i,j] in Indices A holds
b1 * (i,j) = (A * (i,j)) + (B * (i,j))
uniqueness
for b1, b2 being Matrix of n,K st ( for i, j being Nat st [i,j] in Indices A holds
b1 * (i,j) = (A * (i,j)) + (B * (i,j)) ) & ( for i, j being Nat st [i,j] in Indices A holds
b2 * (i,j) = (A * (i,j)) + (B * (i,j)) ) holds
b1 = b2
end;
:: deftheorem defines -Matrices_over MATRIX_1:def 10 :
for K being non empty doubleLoopStr
for n being Nat holds n -Matrices_over K = n -tuples_on (n -tuples_on the carrier of K);
:: deftheorem defines 0. MATRIX_1:def 11 :
for K being non empty doubleLoopStr
for n being Nat holds 0. (K,n) = n |-> (n |-> (0. K));
:: deftheorem Def12 defines 1. MATRIX_1:def 12 :
for K being non empty doubleLoopStr
for n being Nat
for b3 being Matrix of n,K holds
( b3 = 1. (K,n) iff ( ( for i being Nat st [i,i] in Indices b3 holds
b3 * (i,i) = 1. K ) & ( for i, j being Nat st [i,j] in Indices b3 & i <> j holds
b3 * (i,j) = 0. K ) ) );
:: deftheorem Def13 defines - MATRIX_1:def 13 :
for K being non empty doubleLoopStr
for n being Nat
for A, b4 being Matrix of n,K holds
( b4 = - A iff for i, j being Nat st [i,j] in Indices A holds
b4 * (i,j) = - (A * (i,j)) );
:: deftheorem Def14 defines + MATRIX_1:def 14 :
for K being non empty doubleLoopStr
for n being Nat
for A, B, b5 being Matrix of n,K holds
( b5 = A + B iff for i, j being Nat st [i,j] in Indices A holds
b5 * (i,j) = (A * (i,j)) + (B * (i,j)) );
theorem Th30:
theorem Th31:
:: deftheorem defines diagonal MATRIX_1:def 15 :
for K being non empty ZeroStr
for M being Matrix of K holds
( M is diagonal iff for i, j being Nat st [i,j] in Indices M & M * (i,j) <> 0. K holds
i = j );
theorem Th32:
theorem Th33:
theorem Th34:
theorem Th35:
definition
let F be non
empty right_complementable Abelian add-associative right_zeroed doubleLoopStr ;
let n be
Nat;
func n -G_Matrix_over F -> strict AbGroup means
( the
carrier of
it = n -Matrices_over F & ( for
A,
B being
Matrix of
n,
F holds the
addF of
it . (
A,
B)
= A + B ) &
0. it = 0. (
F,
n) );
existence
ex b1 being strict AbGroup st
( the carrier of b1 = n -Matrices_over F & ( for A, B being Matrix of n,F holds the addF of b1 . (A,B) = A + B ) & 0. b1 = 0. (F,n) )
uniqueness
for b1, b2 being strict AbGroup st the carrier of b1 = n -Matrices_over F & ( for A, B being Matrix of n,F holds the addF of b1 . (A,B) = A + B ) & 0. b1 = 0. (F,n) & the carrier of b2 = n -Matrices_over F & ( for A, B being Matrix of n,F holds the addF of b2 . (A,B) = A + B ) & 0. b2 = 0. (F,n) holds
b1 = b2
end;
:: deftheorem defines -G_Matrix_over MATRIX_1:def 16 :
for F being non empty right_complementable Abelian add-associative right_zeroed doubleLoopStr
for n being Nat
for b3 being strict AbGroup holds
( b3 = n -G_Matrix_over F iff ( the carrier of b3 = n -Matrices_over F & ( for A, B being Matrix of n,F holds the addF of b3 . (A,B) = A + B ) & 0. b3 = 0. (F,n) ) );
theorem
begin
theorem Th37:
theorem
theorem Th39:
theorem