begin
:: deftheorem defines 0* MATRIX14:def 1 :
for K being non empty ZeroStr
for n being Element of NAT holds 0* (K,n) = n |-> (0. K);
theorem Th1:
theorem Th2:
theorem Th3:
theorem Th4:
theorem Th5:
theorem Th6:
theorem Th7:
theorem Th8:
theorem
theorem Th10:
theorem Th11:
theorem Th12:
theorem Th13:
begin
theorem Th14:
theorem Th15:
theorem
theorem Th17:
theorem Th18:
theorem Th19:
theorem Th20:
theorem Th21:
theorem Th22:
:: deftheorem defines Base_FinSeq MATRIX14:def 2 :
for K being Field
for n, i being Nat holds Base_FinSeq (K,n,i) = Replace ((n |-> (0. K)),i,(1. K));
theorem Th23:
theorem Th24:
theorem Th25:
theorem Th26:
theorem Th27:
theorem
theorem Th29:
begin
theorem Th30:
theorem Th31:
theorem Th32:
theorem
theorem Th34:
theorem Th35:
theorem Th36:
theorem Th37:
for
n being
Element of
NAT for
K being
Field for
a being
Element of
K for
P,
Q being
Matrix of
n,
K st
n > 0 &
a <> 0. K &
P * (1,1)
= a " & ( for
i being
Element of
NAT st 1
< i &
i <= n holds
P . i = Base_FinSeq (
K,
n,
i) ) &
Q * (1,1)
= a & ( for
j being
Element of
NAT st 1
< j &
j <= n holds
Q * (1,
j)
= - (a * (P * (1,j))) ) & ( for
i being
Element of
NAT st 1
< i &
i <= n holds
Q . i = Base_FinSeq (
K,
n,
i) ) holds
(
P is
invertible &
Q = P ~ )
theorem Th38:
theorem Th39:
theorem Th40:
theorem Th41:
begin
theorem Th42:
definition
let K be
Field;
let n be
Element of
NAT ;
let i0 be
Nat;
func SwapDiagonal (
K,
n,
i0)
-> Matrix of
n,
K equals
Swap (
(1. (K,n)),1,
i0);
correctness
coherence
Swap ((1. (K,n)),1,i0) is Matrix of n,K;
end;
:: deftheorem defines SwapDiagonal MATRIX14:def 3 :
for K being Field
for n being Element of NAT
for i0 being Nat holds SwapDiagonal (K,n,i0) = Swap ((1. (K,n)),1,i0);
theorem Th43:
for
K being
Field for
n being
Element of
NAT for
i0 being
Nat for
A being
Matrix of
n,
K st 1
<= i0 &
i0 <= n &
A = SwapDiagonal (
K,
n,
i0) holds
for
i,
j being
Nat st 1
<= i &
i <= n & 1
<= j &
j <= n &
i0 <> 1 holds
( (
i = 1 &
j = i0 implies
A * (
i,
j)
= 1. K ) & (
i = i0 &
j = 1 implies
A * (
i,
j)
= 1. K ) & (
i = 1 &
j = 1 implies
A * (
i,
j)
= 0. K ) & (
i = i0 &
j = i0 implies
A * (
i,
j)
= 0. K ) & ( ( ( not
i = 1 & not
i = i0 ) or ( not
j = 1 & not
j = i0 ) ) implies ( (
i = j implies
A * (
i,
j)
= 1. K ) & (
i <> j implies
A * (
i,
j)
= 0. K ) ) ) )
theorem Th44:
theorem Th45:
theorem Th46:
theorem Th47:
for
K being
Field for
n,
i0 being
Element of
NAT for
A being
Matrix of
n,
K st 1
<= i0 &
i0 <= n &
i0 <> 1 & ( for
i,
j being
Nat st 1
<= i &
i <= n & 1
<= j &
j <= n holds
( (
i = 1 &
j = i0 implies
A * (
i,
j)
= 1. K ) & (
i = i0 &
j = 1 implies
A * (
i,
j)
= 1. K ) & (
i = 1 &
j = 1 implies
A * (
i,
j)
= 0. K ) & (
i = i0 &
j = i0 implies
A * (
i,
j)
= 0. K ) & ( ( ( not
i = 1 & not
i = i0 ) or ( not
j = 1 & not
j = i0 ) ) implies ( (
i = j implies
A * (
i,
j)
= 1. K ) & (
i <> j implies
A * (
i,
j)
= 0. K ) ) ) ) ) holds
A = SwapDiagonal (
K,
n,
i0)
theorem Th48:
for
n being
Element of
NAT for
K being
Field for
A being
Matrix of
n,
K for
i0 being
Element of
NAT st 1
<= i0 &
i0 <= n holds
( ( for
j being
Element of
NAT st 1
<= j &
j <= n holds
(
((SwapDiagonal (K,n,i0)) * A) * (
i0,
j)
= A * (1,
j) &
((SwapDiagonal (K,n,i0)) * A) * (1,
j)
= A * (
i0,
j) ) ) & ( for
i,
j being
Element of
NAT st
i <> 1 &
i <> i0 & 1
<= i &
i <= n & 1
<= j &
j <= n holds
((SwapDiagonal (K,n,i0)) * A) * (
i,
j)
= A * (
i,
j) ) )
theorem Th49:
theorem Th50:
theorem Th51:
for
n being
Element of
NAT for
K being
Field for
A being
Matrix of
n,
K for
j0 being
Element of
NAT st 1
<= j0 &
j0 <= n holds
( ( for
i being
Element of
NAT st 1
<= i &
i <= n holds
(
(A * (SwapDiagonal (K,n,j0))) * (
i,
j0)
= A * (
i,1) &
(A * (SwapDiagonal (K,n,j0))) * (
i,1)
= A * (
i,
j0) ) ) & ( for
i,
j being
Element of
NAT st
j <> 1 &
j <> j0 & 1
<= i &
i <= n & 1
<= j &
j <= n holds
(A * (SwapDiagonal (K,n,j0))) * (
i,
j)
= A * (
i,
j) ) )
theorem Th52:
begin
theorem Th53:
theorem Th54:
theorem Th55:
theorem