:: MATRIX_3 semantic presentation
Lemma1:
for b1, b2 being Nat holds
( b1 = b1 + b2 implies b2 = 0 )
by XCMPLX_1:3;
:: deftheorem Def1 defines 0. MATRIX_3:def 1 :
definition
let c
1 be
Field;
let c
2 be
Matrix of c
1;
func - c
2 -> Matrix of a
1 means :
Def2:
:: MATRIX_3:def 2
(
len a
3 = len a
2 &
width a
3 = width a
2 & ( for b
1, b
2 being
Nat holds
(
[b1,b2] in Indices a
2 implies a
3 * b
1,b
2 = - (a2 * b1,b2) ) ) );
existence
ex b1 being Matrix of c1 st
( len b1 = len c2 & width b1 = width c2 & ( for b2, b3 being Nat holds
( [b2,b3] in Indices c2 implies b1 * b2,b3 = - (c2 * b2,b3) ) ) )
uniqueness
for b1, b2 being Matrix of c1 holds
( len b1 = len c2 & width b1 = width c2 & ( for b3, b4 being Nat holds
( [b3,b4] in Indices c2 implies b1 * b3,b4 = - (c2 * b3,b4) ) ) & len b2 = len c2 & width b2 = width c2 & ( for b3, b4 being Nat holds
( [b3,b4] in Indices c2 implies b2 * b3,b4 = - (c2 * b3,b4) ) ) implies b1 = b2 )
end;
:: deftheorem Def2 defines - MATRIX_3:def 2 :
definition
let c
1 be
Field;
let c
2, c
3 be
Matrix of c
1;
func c
2 + c
3 -> Matrix of a
1 means :
Def3:
:: MATRIX_3:def 3
(
len a
4 = len a
2 &
width a
4 = width a
2 & ( for b
1, b
2 being
Nat holds
(
[b1,b2] in Indices a
2 implies a
4 * b
1,b
2 = (a2 * b1,b2) + (a3 * b1,b2) ) ) );
existence
ex b1 being Matrix of c1 st
( len b1 = len c2 & width b1 = width c2 & ( for b2, b3 being Nat holds
( [b2,b3] in Indices c2 implies b1 * b2,b3 = (c2 * b2,b3) + (c3 * b2,b3) ) ) )
uniqueness
for b1, b2 being Matrix of c1 holds
( len b1 = len c2 & width b1 = width c2 & ( for b3, b4 being Nat holds
( [b3,b4] in Indices c2 implies b1 * b3,b4 = (c2 * b3,b4) + (c3 * b3,b4) ) ) & len b2 = len c2 & width b2 = width c2 & ( for b3, b4 being Nat holds
( [b3,b4] in Indices c2 implies b2 * b3,b4 = (c2 * b3,b4) + (c3 * b3,b4) ) ) implies b1 = b2 )
end;
:: deftheorem Def3 defines + MATRIX_3:def 3 :
theorem Th1: :: MATRIX_3:1
canceled;
theorem Th2: :: MATRIX_3:2
canceled;
theorem Th3: :: MATRIX_3:3
theorem Th4: :: MATRIX_3:4
theorem Th5: :: MATRIX_3:5
theorem Th6: :: MATRIX_3:6
for b
1, b
2 being
Natfor b
3 being
Fieldfor b
4 being
Matrix of b
1,b
2,b
3 holds b
4 + (0. b3,b1,b2) = b
4
theorem Th7: :: MATRIX_3:7
for b
1, b
2 being
Natfor b
3 being
Fieldfor b
4 being
Matrix of b
1,b
2,b
3 holds b
4 + (- b4) = 0. b
3,b
1,b
2
definition
let c
1 be
Field;
let c
2, c
3 be
Matrix of c
1;
assume E5:
width c
2 = len c
3
;
func c
2 * c
3 -> Matrix of a
1 means :
Def4:
:: MATRIX_3:def 4
(
len a
4 = len a
2 &
width a
4 = width a
3 & ( for b
1, b
2 being
Nat holds
(
[b1,b2] in Indices a
4 implies a
4 * b
1,b
2 = (Line a2,b1) "*" (Col a3,b2) ) ) );
existence
ex b1 being Matrix of c1 st
( len b1 = len c2 & width b1 = width c3 & ( for b2, b3 being Nat holds
( [b2,b3] in Indices b1 implies b1 * b2,b3 = (Line c2,b2) "*" (Col c3,b3) ) ) )
uniqueness
for b1, b2 being Matrix of c1 holds
( len b1 = len c2 & width b1 = width c3 & ( for b3, b4 being Nat holds
( [b3,b4] in Indices b1 implies b1 * b3,b4 = (Line c2,b3) "*" (Col c3,b4) ) ) & len b2 = len c2 & width b2 = width c3 & ( for b3, b4 being Nat holds
( [b3,b4] in Indices b2 implies b2 * b3,b4 = (Line c2,b3) "*" (Col c3,b4) ) ) implies b1 = b2 )
end;
:: deftheorem Def4 defines * MATRIX_3:def 4 :
definition
let c
1, c
2, c
3 be
Nat;
let c
4 be
Field;
let c
5 be
Matrix of c
1,c
2,c
4;
let c
6 be
Matrix of
width c
5,c
3,c
4;
redefine func * as c
5 * c
6 -> Matrix of
len a
5,
width a
6,a
4;
coherence
c5 * c6 is Matrix of len c5, width c6,c4
end;
definition
let c
1 be
Field;
let c
2 be
Matrix of c
1;
let c
3 be
Element of c
1;
func c
3 * c
2 -> Matrix of a
1 means :: MATRIX_3:def 5
(
len a
4 = len a
2 &
width a
4 = width a
2 & ( for b
1, b
2 being
Nat holds
(
[b1,b2] in Indices a
2 implies a
4 * b
1,b
2 = a
3 * (a2 * b1,b2) ) ) );
existence
ex b1 being Matrix of c1 st
( len b1 = len c2 & width b1 = width c2 & ( for b2, b3 being Nat holds
( [b2,b3] in Indices c2 implies b1 * b2,b3 = c3 * (c2 * b2,b3) ) ) )
uniqueness
for b1, b2 being Matrix of c1 holds
( len b1 = len c2 & width b1 = width c2 & ( for b3, b4 being Nat holds
( [b3,b4] in Indices c2 implies b1 * b3,b4 = c3 * (c2 * b3,b4) ) ) & len b2 = len c2 & width b2 = width c2 & ( for b3, b4 being Nat holds
( [b3,b4] in Indices c2 implies b2 * b3,b4 = c3 * (c2 * b3,b4) ) ) implies b1 = b2 )
end;
:: deftheorem Def5 defines * MATRIX_3:def 5 :
:: deftheorem Def6 defines * MATRIX_3:def 6 :
theorem Th8: :: MATRIX_3:8
theorem Th9: :: MATRIX_3:9
theorem Th10: :: MATRIX_3:10
theorem Th11: :: MATRIX_3:11
theorem Th12: :: MATRIX_3:12
Lemma6:
for b1 being non empty add-associative right_zeroed right_complementable LoopStr
for b2 being FinSequence of b1 holds
( ( for b3 being Nat holds
( b3 in dom b2 implies b2 . b3 = 0. b1 ) ) implies Sum b2 = 0. b1 )
theorem Th13: :: MATRIX_3:13
theorem Th14: :: MATRIX_3:14
theorem Th15: :: MATRIX_3:15
theorem Th16: :: MATRIX_3:16
for b
1 being
Fieldfor b
2, b
3 being
FinSequence of b
1for b
4 being
Nat holds
( b
4 in dom b
2 & b
2 . b
4 = 1. b
1 & ( for b
5 being
Nat holds
( b
5 in dom b
2 & b
5 <> b
4 implies b
2 . b
5 = 0. b
1 ) ) implies for b
5 being
Nat holds
( b
5 in dom (mlt b2,b3) implies ( ( b
4 = b
5 implies
(mlt b2,b3) . b
5 = b
3 . b
4 ) & ( b
4 <> b
5 implies
(mlt b2,b3) . b
5 = 0. b
1 ) ) ) )
theorem Th17: :: MATRIX_3:17
theorem Th18: :: MATRIX_3:18
theorem Th19: :: MATRIX_3:19
theorem Th20: :: MATRIX_3:20
theorem Th21: :: MATRIX_3:21
theorem Th22: :: MATRIX_3:22
theorem Th23: :: MATRIX_3:23
for b
1 being
Fieldfor b
2, b
3, b
4, b
5, b
6, b
7, b
8, b
9 being
Element of b
1 holds
(b2,b4 ][ b6,b8) * (b3,b5 ][ b7,b9) = ((b2 * b3) + (b4 * b7)),
((b2 * b5) + (b4 * b9)) ][ ((b6 * b3) + (b8 * b7)),
((b6 * b5) + (b8 * b9))
theorem Th24: :: MATRIX_3:24
definition
let c
1, c
2, c
3 be non
empty set ;
let c
4 be
BinOp of c
3;
let c
5 be
Function of c
1,c
3;
let c
6 be
Function of c
2,c
3;
redefine func * as c
4 * c
5,c
6 -> Function of
[:a1,a2:],a
3;
coherence
c4 * c5,c6 is Function of [:c1,c2:],c3
by FINSEQOP:84;
end;
theorem Th25: :: MATRIX_3:25
for b
1, b
2, b
3 being non
empty set for b
4, b
5 being
BinOp of b
3for b
6 being
Function of b
1,b
3for b
7 being
Function of b
2,b
3for b
8 being
Element of
Fin b
1for b
9 being
Element of
Fin b
2 holds
( b
4 is
commutative & b
4 is
associative & not ( not
[:b9,b8:] <> {} & not b
4 has_a_unity ) & b
5 is
commutative implies b
4 $$ [:b8,b9:],
(b5 * b6,b7) = b
4 $$ [:b9,b8:],
(b5 * b7,b6) )
theorem Th26: :: MATRIX_3:26
for b
1, b
2, b
3 being non
empty set for b
4, b
5 being
BinOp of b
1for b
6 being
Function of b
2,b
1for b
7 being
Function of b
3,b
1 holds
( b
4 is
commutative & b
4 is
associative & b
4 has_a_unity implies for b
8 being
Element of b
2for b
9 being
Element of b
3 holds b
4 $$ [:{.b8.},{.b9.}:],
(b5 * b6,b7) = b
4 $$ {.b8.},
(b5 [:] b6,(b4 $$ {.b9.},b7)) )
theorem Th27: :: MATRIX_3:27
for b
1, b
2, b
3 being non
empty set for b
4, b
5 being
BinOp of b
1for b
6 being
Function of b
2,b
1for b
7 being
Function of b
3,b
1for b
8 being
Element of
Fin b
2for b
9 being
Element of
Fin b
3 holds
( b
4 is
commutative & b
4 is
associative & b
4 has_a_unity & b
4 has_an_inverseOp & b
5 is_distributive_wrt b
4 implies for b
10 being
Element of b
2 holds b
4 $$ [:{.b10.},b9:],
(b5 * b6,b7) = b
4 $$ {.b10.},
(b5 [:] b6,(b4 $$ b9,b7)) )
theorem Th28: :: MATRIX_3:28
for b
1, b
2, b
3 being non
empty set for b
4, b
5 being
BinOp of b
1for b
6 being
Function of b
2,b
1for b
7 being
Function of b
3,b
1for b
8 being
Element of
Fin b
2for b
9 being
Element of
Fin b
3 holds
( b
4 has_a_unity & b
4 is
commutative & b
4 is
associative & b
4 has_an_inverseOp & b
5 is_distributive_wrt b
4 implies b
4 $$ [:b8,b9:],
(b5 * b6,b7) = b
4 $$ b
8,
(b5 [:] b6,(b4 $$ b9,b7)) )
theorem Th29: :: MATRIX_3:29
for b
1, b
2, b
3 being non
empty set for b
4, b
5 being
BinOp of b
1for b
6 being
Function of b
2,b
1for b
7 being
Function of b
3,b
1 holds
( b
4 is
commutative & b
4 is
associative & b
4 has_a_unity & b
5 is
commutative implies for b
8 being
Element of b
2for b
9 being
Element of b
3 holds b
4 $$ [:{.b8.},{.b9.}:],
(b5 * b6,b7) = b
4 $$ {.b9.},
(b5 [;] (b4 $$ {.b8.},b6),b7) )
theorem Th30: :: MATRIX_3:30
for b
1, b
2, b
3 being non
empty set for b
4, b
5 being
BinOp of b
1for b
6 being
Function of b
2,b
1for b
7 being
Function of b
3,b
1for b
8 being
Element of
Fin b
2for b
9 being
Element of
Fin b
3 holds
( b
4 has_a_unity & b
4 is
commutative & b
4 is
associative & b
4 has_an_inverseOp & b
5 is_distributive_wrt b
4 & b
5 is
commutative implies b
4 $$ [:b8,b9:],
(b5 * b6,b7) = b
4 $$ b
9,
(b5 [;] (b4 $$ b8,b6),b7) )
theorem Th31: :: MATRIX_3:31
theorem Th32: :: MATRIX_3:32
theorem Th33: :: MATRIX_3:33
theorem Th34: :: MATRIX_3:34
theorem Th35: :: MATRIX_3:35
definition
let c
1 be
Nat;
let c
2 be
Field;
let c
3 be
Matrix of c
1,c
2;
let c
4 be
Element of
Permutations c
1;
func Path_matrix c
4,c
3 -> FinSequence of a
2 means :
Def7:
:: MATRIX_3:def 7
(
len a
5 = a
1 & ( for b
1, b
2 being
Nat holds
( b
1 in dom a
5 & b
2 = a
4 . b
1 implies a
5 . b
1 = a
3 * b
1,b
2 ) ) );
existence
ex b1 being FinSequence of c2 st
( len b1 = c1 & ( for b2, b3 being Nat holds
( b2 in dom b1 & b3 = c4 . b2 implies b1 . b2 = c3 * b2,b3 ) ) )
uniqueness
for b1, b2 being FinSequence of c2 holds
( len b1 = c1 & ( for b3, b4 being Nat holds
( b3 in dom b1 & b4 = c4 . b3 implies b1 . b3 = c3 * b3,b4 ) ) & len b2 = c1 & ( for b3, b4 being Nat holds
( b3 in dom b2 & b4 = c4 . b3 implies b2 . b3 = c3 * b3,b4 ) ) implies b1 = b2 )
end;
:: deftheorem Def7 defines Path_matrix MATRIX_3:def 7 :
definition
let c
1 be
Nat;
let c
2 be
Field;
let c
3 be
Matrix of c
1,c
2;
func Path_product c
3 -> Function of
Permutations a
1,the
carrier of a
2 means :
Def8:
:: MATRIX_3:def 8
for b
1 being
Element of
Permutations a
1 holds a
4 . b
1 = - (the mult of a2 $$ (Path_matrix b1,a3)),b
1;
existence
ex b1 being Function of Permutations c1,the carrier of c2 st
for b2 being Element of Permutations c1 holds b1 . b2 = - (the mult of c2 $$ (Path_matrix b2,c3)),b2
uniqueness
for b1, b2 being Function of Permutations c1,the carrier of c2 holds
( ( for b3 being Element of Permutations c1 holds b1 . b3 = - (the mult of c2 $$ (Path_matrix b3,c3)),b3 ) & ( for b3 being Element of Permutations c1 holds b2 . b3 = - (the mult of c2 $$ (Path_matrix b3,c3)),b3 ) implies b1 = b2 )
end;
:: deftheorem Def8 defines Path_product MATRIX_3:def 8 :
:: deftheorem Def9 defines Det MATRIX_3:def 9 :
theorem Th36: :: MATRIX_3:36
:: deftheorem Def10 defines diagonal_of_Matrix MATRIX_3:def 10 :