Theory Diagonal_form

Up to index of Isabelle/HOL/HOL-Algebra/Diagonal_form

theory Diagonal_form
imports Matrix IntRing

theory Diagonal_form
imports Main
"~~/src/HOL/Matrix/Matrix"
"~~/src/HOL/Divides"
"~~/src/HOL-Algebra/Ring"
"~~/src/HOL-Algebra/IntRing"
begin


(* Authors: Jesus Maria Aransay
Jose Divason
*)


text{*This function interchanges the columns n and m of an infinite matrix*}
definition interchange_columns_infmatrix :: "'a infmatrix => nat => nat => 'a infmatrix" where
"interchange_columns_infmatrix A n m == (λi j. if j=n then A i m else if j = m then A i n else A i j)"


text{*This function interchanges the columns n and m of a finite matrix*}
definition interchange_columns_matrix :: "int matrix => nat => nat => int matrix"
where "interchange_columns_matrix A n m == Abs_matrix (interchange_columns_infmatrix (Rep_matrix A) n m)"



lemma Rep_Abs_aux_interchange':
shows "Rep_matrix (Abs_matrix (λi j. if j = n then Rep_matrix A i m else if j = m then Rep_matrix A i n else Rep_matrix A i j)) =
(λi j. if j = n then Rep_matrix A i m else if j = m then Rep_matrix A i n else Rep_matrix A i j)"

proof (rule RepAbs_matrix)
obtain a where a:"∀j s. a≤j --> Rep_matrix A s j = (0::'a)" by (metis ncols)
obtain b where b:"∀i s. b≤i --> Rep_matrix A i s = (0::'a)" by (metis nrows)
show "∃ma. ∀j i. ma ≤ j --> (if i = n then Rep_matrix A j m else if i = m then Rep_matrix A j n else Rep_matrix A j i) = (0::'a)"
using b by auto
show "∃na. ∀j i. na ≤ i --> (if i = n then Rep_matrix A j m else if i = m then Rep_matrix A j n else Rep_matrix A j i) = (0::'a)"
proof (rule exI[of _ "max (n+1) (max (m+1) a)"], auto+)
fix i j
assume "Suc n ≤ i" and "Suc m ≤ i" and a_le_i:"a≤i"
show "Rep_matrix A j i = (0::'a)" using a a_le_i by auto
qed
qed

lemma Rep_Abs_aux_interchange2':
shows "Rep_matrix (Abs_matrix (λi j. if i = n then Rep_matrix A m j else if i = m then Rep_matrix A n j else Rep_matrix A i j)) =
(λi j. if i = n then Rep_matrix A m j else if i = m then Rep_matrix A n j else Rep_matrix A i j)"

proof (rule RepAbs_matrix)
obtain a where a:"∀j s. a≤j --> Rep_matrix A s j = (0::'a)" by (metis ncols)
obtain b where b:"∀i s. b≤i --> Rep_matrix A i s = (0::'a)" by (metis nrows)
show "∃ma. ∀j i. ma ≤ j --> (if j = n then Rep_matrix A m i else if j = m then Rep_matrix A n i else Rep_matrix A j i) = (0::'a)"
proof (rule exI[of _ "max (n+1) (max (m+1) b)"], auto+)
fix i j
assume "Suc n ≤ j" and "Suc m ≤ j" and b_le_j:"b≤j"
show "Rep_matrix A j i = (0::'a)" using b b_le_j by blast
qed
show "∃na. ∀j i. na ≤ i --> (if j = n then Rep_matrix A m i else if j = m then Rep_matrix A n i else Rep_matrix A j i) = (0::'a)"
using a by auto
qed

text{*If we interchange two columns twice, we will obtain the original matrix.*}
lemma interchange_columns_matrix_id:
shows "interchange_columns_matrix (interchange_columns_matrix A n m) n m = A"

proof (unfold interchange_columns_matrix_def, unfold interchange_columns_infmatrix_def)
have "A=Abs_matrix(Rep_matrix A)" using Rep_matrix_inverse [of A, symmetric] .
also have "...= Abs_matrix
(λi j. if j = n then Rep_matrix (Abs_matrix (λi j. if j = n then Rep_matrix A i m else if j = m then Rep_matrix A i n else Rep_matrix A i j)) i m
else if j = m then Rep_matrix (Abs_matrix (λi j. if j = n then Rep_matrix A i m else if j = m then Rep_matrix A i n else Rep_matrix A i j)) i n
else Rep_matrix (Abs_matrix (λi j. if j = n then Rep_matrix A i m else if j = m then Rep_matrix A i n else Rep_matrix A i j)) i j)"

proof (rule cong[of _ Abs_matrix], simp, (rule ext)+)
fix i j
show "Rep_matrix A i j =
(if j = n then Rep_matrix (Abs_matrix (λi j. if j = n then Rep_matrix A i m else if j = m then Rep_matrix A i n else Rep_matrix A i j)) i m
else if j = m then Rep_matrix (Abs_matrix (λi j. if j = n then Rep_matrix A i m else if j = m then Rep_matrix A i n else Rep_matrix A i j)) i n
else Rep_matrix (Abs_matrix (λi j. if j = n then Rep_matrix A i m else if j = m then Rep_matrix A i n else Rep_matrix A i j)) i j)"

using Rep_Abs_aux_interchange'[of n A m] by auto
qed
finally show "Abs_matrix
(λi j. if j = n then Rep_matrix (Abs_matrix (λi j. if j = n then Rep_matrix A i m else if j = m then Rep_matrix A i n else Rep_matrix A i j)) i m
else if j = m then Rep_matrix (Abs_matrix (λi j. if j = n then Rep_matrix A i m else if j = m then Rep_matrix A i n else Rep_matrix A i j)) i n
else Rep_matrix (Abs_matrix (λi j. if j = n then Rep_matrix A i m else if j = m then Rep_matrix A i n else Rep_matrix A i j)) i j) = A"

by auto
qed

text{*If we interchange two columns a and b, the elements situated in column a now there will be in column b*}
lemma interchange_columns:
shows "∀i. Rep_matrix(interchange_columns_matrix A a b) i a = Rep_matrix A i b"

unfolding interchange_columns_matrix_def interchange_columns_infmatrix_def
proof
fix i
show "Rep_matrix (Abs_matrix (λi j. if j = a then Rep_matrix A i b else if j = b then Rep_matrix A i a else Rep_matrix A i j)) i a = Rep_matrix A i b"
by (metis (no_types) Rep_Abs_aux_interchange')
qed

lemma interchange_columns':
shows "∀i. Rep_matrix(interchange_columns_matrix A a b) i b = Rep_matrix A i a"

unfolding interchange_columns_matrix_def interchange_columns_infmatrix_def
proof
fix i
show "Rep_matrix (Abs_matrix (λi j. if j = a then Rep_matrix A i b else if j = b then Rep_matrix A i a else Rep_matrix A i j)) i b = Rep_matrix A i a"
using Rep_Abs_aux_interchange'[of a A b] by simp
qed


lemma interchange_columns_eq:
shows "∀i s. (s≠n ∧ s≠m) --> Rep_matrix (interchange_columns_matrix (A::int matrix) n m) i s = (Rep_matrix A) i s"

unfolding interchange_columns_matrix_def interchange_columns_infmatrix_def
using Rep_Abs_aux_interchange'[of n A m] by simp

text{*This function interchanges two rows of an infinite matrix*}
definition interchange_rows_infmatrix :: "'a infmatrix => nat => nat => 'a infmatrix"
where "interchange_rows_infmatrix A n m == (λi j. if i=n then A m j else if i = m then A n j else A i j)"


text{*This function interchanges two rows of a finite matrix*}
definition interchange_rows_matrix :: "int matrix => nat => nat => int matrix"
where "interchange_rows_matrix A n m== Abs_matrix (interchange_rows_infmatrix (Rep_matrix A) n m)"


lemma interchange_rows_matrix_id:
shows "interchange_rows_matrix (interchange_rows_matrix A n m) n m = A"

proof (unfold interchange_rows_matrix_def, unfold interchange_rows_infmatrix_def)
have "A=Abs_matrix(Rep_matrix A)" using Rep_matrix_inverse [of A, symmetric] .
also have "...=Abs_matrix
(λi j. if i = n then Rep_matrix (Abs_matrix (λi j. if i = n then Rep_matrix A m j else if i = m then Rep_matrix A n j else Rep_matrix A i j)) m j
else if i = m then Rep_matrix (Abs_matrix (λi j. if i = n then Rep_matrix A m j else if i = m then Rep_matrix A n j else Rep_matrix A i j)) n j
else Rep_matrix (Abs_matrix (λi j. if i = n then Rep_matrix A m j else if i = m then Rep_matrix A n j else Rep_matrix A i j)) i j)"

proof (rule cong[of _ Abs_matrix], simp, (rule ext)+)
fix i j
show "Rep_matrix A i j =
(if i = n then Rep_matrix (Abs_matrix (λi j. if i = n then Rep_matrix A m j else if i = m then Rep_matrix A n j else Rep_matrix A i j)) m j
else if i = m then Rep_matrix (Abs_matrix (λi j. if i = n then Rep_matrix A m j else if i = m then Rep_matrix A n j else Rep_matrix A i j)) n j
else Rep_matrix (Abs_matrix (λi j. if i = n then Rep_matrix A m j else if i = m then Rep_matrix A n j else Rep_matrix A i j)) i j)"

using Rep_Abs_aux_interchange2'[of n A m] by auto
qed
finally show "Abs_matrix
(λi j. if i = n then Rep_matrix (Abs_matrix (λi j. if i = n then Rep_matrix A m j else if i = m then Rep_matrix A n j else Rep_matrix A i j)) m j
else if i = m then Rep_matrix (Abs_matrix (λi j. if i = n then Rep_matrix A m j else if i = m then Rep_matrix A n j else Rep_matrix A i j)) n j
else Rep_matrix (Abs_matrix (λi j. if i = n then Rep_matrix A m j else if i = m then Rep_matrix A n j else Rep_matrix A i j)) i j) = A"

by simp
qed


lemma interchange_rows:
shows "∀j. Rep_matrix(interchange_rows_matrix A a b) a j = Rep_matrix A b j"

unfolding interchange_rows_matrix_def interchange_rows_infmatrix_def
proof
fix j
show "Rep_matrix (Abs_matrix (λi j. if i = a then Rep_matrix A b j else if i = b then Rep_matrix A a j else Rep_matrix A i j)) a j = Rep_matrix A b j"
by (metis (no_types) Rep_Abs_aux_interchange2')
qed


lemma interchange_rows':
shows "∀j. Rep_matrix(interchange_rows_matrix A a b) b j = Rep_matrix A a j"

unfolding interchange_rows_matrix_def interchange_rows_infmatrix_def
proof
fix j
show "Rep_matrix (Abs_matrix (λi j. if i = a then Rep_matrix A b j else if i = b then Rep_matrix A a j else Rep_matrix A i j)) b j = Rep_matrix A a j"
using Rep_Abs_aux_interchange2'[of a A b] by simp
qed

lemma interchange_rows_matrix_eq:
shows "∀i s. (s≠n ∧ s≠m) --> Rep_matrix (interchange_rows_matrix A n m) s i = (Rep_matrix A) s i"

unfolding interchange_rows_matrix_def interchange_rows_infmatrix_def
using Rep_Abs_aux_interchange2'[of n A m] by simp

lemma interchange_columns_interchange_rows:
shows "Rep_matrix (interchange_columns_matrix (interchange_rows_matrix A a b) n m) a n = (Rep_matrix A) b m"

by (metis interchange_columns interchange_rows)


text{*For a given matrix, this function moves the element (i,j) to the position (k, l)*}
definition move_element :: "int matrix => nat => nat => nat => nat => int matrix"
where "move_element A i j k l = (interchange_columns_matrix (interchange_rows_matrix A i k) j l)"


lemma move_element:
shows "Rep_matrix (move_element A i j k l) k l = (Rep_matrix A) i j"

by (metis interchange_columns' interchange_rows' move_element_def)


text{*This function adds to row i the row j multiplied by q*}
definition row_add_infmatrix :: "('a::{plus,times}) infmatrix => nat => nat => 'a => 'a infmatrix"
where "row_add_infmatrix A i j q = (λa b. if a=i then ((A i b) + q*(A j b)) else A a b)"


definition row_add_matrix :: "int matrix => nat => nat => int => int matrix"
where "row_add_matrix A i j q = Abs_matrix (row_add_infmatrix (Rep_matrix A) i j q)"


lemma Rep_aux_row_add:
shows "Rep_matrix (Abs_matrix (λa b. if a = i then Rep_matrix (A::('a::ring) matrix) i b + q * Rep_matrix A j b else Rep_matrix A a b))
=(λa b. if a = i then Rep_matrix A i b + q * Rep_matrix A j b else Rep_matrix A a b)"

proof (rule RepAbs_matrix)
obtain a where a:"∀j s. a≤j --> Rep_matrix A s j = (0::'a)" by (metis ncols)
obtain b where b:"∀i s. b≤i --> Rep_matrix A i s = (0::'a)" by (metis nrows)
show "∃n. ∀ja ia. n ≤ ia --> (if ja = i then Rep_matrix A i ia + q * Rep_matrix A j ia else Rep_matrix A ja ia) = (0::'a)"
proof (rule exI[of _ "max (i+1) (max (j+1) a)"], auto+)
fix ia assume "a ≤ ia"
thus "Rep_matrix A i ia + q * Rep_matrix A j ia = (0::'a)" using a by auto
next
fix ja ia assume "a ≤ ia" thus "Rep_matrix A ja ia = (0::'a)" using a by simp
qed
show "∃m. ∀ja ia. m ≤ ja --> (if ja = i then Rep_matrix A i ia + q * Rep_matrix A j ia else Rep_matrix A ja ia) = (0::'a)"
proof (rule exI[of _ "max (i+1) (max (j+1) b)"], auto+)
fix ia ja assume "b≤ja" thus "Rep_matrix A ja ia = (0::'a)" using b by simp
qed
qed


lemma row_add_matrix_eq:
shows "∀s. Rep_matrix (row_add_matrix A i j q) i s = (Rep_matrix A) i s + q*(Rep_matrix A) j s"

proof
fix s
show "Rep_matrix (row_add_matrix A i j q) i s = Rep_matrix A i s + q * Rep_matrix A j s"
unfolding row_add_matrix_def row_add_infmatrix_def
proof -
have "Rep_matrix (Abs_matrix (λa b. if a = i then Rep_matrix A i b + q * Rep_matrix A j b else Rep_matrix A a b))
=(λa b. if a = i then Rep_matrix A i b + q * Rep_matrix A j b else Rep_matrix A a b)"

using Rep_aux_row_add .
thus "Rep_matrix (Abs_matrix (λa b. if a = i then Rep_matrix A i b + q * Rep_matrix A j b else Rep_matrix A a b)) i s
= Rep_matrix A i s + q * Rep_matrix A j s"

by auto
qed
qed

lemma Rep_matrix_row_add_matrix_preserve:
shows "∀s t. s≠i --> Rep_matrix (row_add_matrix A i j q) s t = (Rep_matrix A) s t"

unfolding row_add_matrix_def row_add_infmatrix_def using Rep_aux_row_add[of i A q] by auto


text{*This function adds to column i the column j multiplied by q*}
definition column_add_infmatrix :: "('a::{plus,times}) infmatrix => nat => nat => 'a => 'a infmatrix"
where "column_add_infmatrix A i j q = (λa b. if b=i then ((A a i) + q*(A a j)) else A a b)"


definition column_add_matrix :: "int matrix => nat => nat => int => int matrix"
where "column_add_matrix A i j q = Abs_matrix (column_add_infmatrix (Rep_matrix A) i j q)"


lemma Rep_aux_column_add:
shows "Rep_matrix (Abs_matrix (λa b. if b = i then Rep_matrix (A::('a::ring) matrix) a i + q * Rep_matrix A a j else Rep_matrix A a b))
=(λa b. if b = i then Rep_matrix A a i + q * Rep_matrix A a j else Rep_matrix A a b)"

proof (rule RepAbs_matrix)
obtain a where a:"∀j s. a≤j --> Rep_matrix A s j = (0::'a)" by (metis ncols)
obtain b where b:"∀i s. b≤i --> Rep_matrix A i s = (0::'a)" by (metis nrows)
show "∃m. ∀ja ia. m ≤ ja --> (if ia = i then Rep_matrix A ja i + q * Rep_matrix A ja j else Rep_matrix A ja ia) = (0::'a)"
proof (rule exI[of _ "max (i+1) (max (j+1) b)"], auto+)
fix ja assume "b ≤ ja" thus "Rep_matrix A ja i + q * Rep_matrix A ja j = (0::'a)" using b by simp
next
fix ja ia assume "b ≤ ja" thus "Rep_matrix A ja ia = (0::'a)" using b by simp
qed
show "∃n. ∀ja ia. n ≤ ia --> (if ia = i then Rep_matrix A ja i + q * Rep_matrix A ja j else Rep_matrix A ja ia) = (0::'a)"
proof (rule exI[of _ "max (i+1) (max (j+1) a)"], auto+)
fix ja ia assume "a ≤ ia" thus "Rep_matrix A ja ia = (0::'a)"using a by simp
qed
qed


lemma column_add_matrix_eq:
shows "∀s. Rep_matrix (column_add_matrix A i j q) s i = (Rep_matrix A) s i + q*(Rep_matrix A) s j "

proof
fix s
show "Rep_matrix (column_add_matrix A i j q) s i = (Rep_matrix A) s i + q*(Rep_matrix A) s j"
unfolding column_add_matrix_def column_add_infmatrix_def
proof -
have "Rep_matrix (Abs_matrix (λa b. if b = i then Rep_matrix A a i + q * Rep_matrix A a j else Rep_matrix A a b))
=(λa b. if b = i then Rep_matrix A a i + q * Rep_matrix A a j else Rep_matrix A a b)"

using Rep_aux_column_add .
thus "Rep_matrix (Abs_matrix (λa b. if b = i then Rep_matrix A a i + q * Rep_matrix A a j else Rep_matrix A a b)) s i
= Rep_matrix A s i + q * Rep_matrix A s j"

by simp
qed
qed


lemma Rep_matrix_column_add_matrix_preserve:
shows "∀s t. t≠i --> Rep_matrix (column_add_matrix A i j q) s t = (Rep_matrix A) s t "

unfolding column_add_matrix_def column_add_infmatrix_def using Rep_aux_column_add[of i A q j] by simp


text{*Function that multiply a column by a number*}
definition column_multiply_infmatrix :: "('a::{plus,times}) infmatrix => nat => 'a => 'a infmatrix"
where "column_multiply_infmatrix A i q = (λa b. if b=i then q*(A a i) else A a b)"


definition column_multiply_matrix :: "int matrix => nat => int => int matrix"
where "column_multiply_matrix A i q = Abs_matrix (column_multiply_infmatrix (Rep_matrix A) i q)"


lemma Rep_abs_aux_column_multiply:
shows "Rep_matrix (Abs_matrix (λa b. if b = i then q * Rep_matrix (A::('a::ring) matrix) a i else Rep_matrix A a b))
=(λa b. if b = i then q * Rep_matrix A a i else Rep_matrix A a b)"

proof (rule RepAbs_matrix)
obtain a where a:"∀j s. a≤j --> Rep_matrix A s j = (0::'a)" by (metis ncols)
obtain b where b:"∀i s. b≤i --> Rep_matrix A i s = (0::'a)" by (metis nrows)
show "∃m. ∀j ia. m ≤ j --> (if ia = i then q * Rep_matrix A j i else Rep_matrix A j ia) = (0::'a)"
proof (rule exI[of _ "max (i+1) (max (j+1) b)"], auto+)
fix ja assume "b ≤ ja" thus "q * Rep_matrix A ja i = (0::'a)" using b by simp
next
fix ja ia assume "b ≤ ja" thus "Rep_matrix A ja ia = (0::'a)" using b by simp
qed
show "∃n. ∀j ia. n ≤ ia --> (if ia = i then q * Rep_matrix A j i else Rep_matrix A j ia) = (0::'a)"
proof (rule exI[of _ "max (i+1) (max (j+1) a)"], auto+)
fix ja ia assume " a ≤ ia" thus "Rep_matrix A ja ia = (0::'a)"using a by simp
qed
qed


lemma column_multiply_matrix_eq:
shows "∀s. Rep_matrix (column_multiply_matrix A i q) s i = q*((Rep_matrix A) s i)"

proof
fix s
show " Rep_matrix (column_multiply_matrix A i q) s i = q * Rep_matrix A s i"
unfolding column_multiply_matrix_def column_multiply_infmatrix_def
proof -
have "Rep_matrix (Abs_matrix (λa b. if b = i then q * Rep_matrix A a i else Rep_matrix A a b))
=(λa b. if b = i then q * Rep_matrix A a i else Rep_matrix A a b)"

using Rep_abs_aux_column_multiply .
thus "Rep_matrix (Abs_matrix (λa b. if b = i then q * Rep_matrix A a i else Rep_matrix A a b)) s i = q * Rep_matrix A s i"
by simp
qed
qed


lemma column_multiply_matrix_eq2:
shows "∀s t. t≠i --> Rep_matrix (column_multiply_matrix A i q) s t = (Rep_matrix A) s t"

find_theorems "Rep_matrix ?A"
unfolding column_multiply_matrix_def column_multiply_infmatrix_def
using Rep_abs_aux_column_multiply[of i q A] by simp


text{*This function multiplies a row i by a number q.*}
definition row_multiply_infmatrix :: "('a::{plus,times}) infmatrix => nat => 'a => 'a infmatrix"
where "row_multiply_infmatrix A i q = (λa b. if a=i then q*(A i b) else A a b)"


definition row_multiply_matrix :: "int matrix => nat => int => int matrix"
where "row_multiply_matrix A i q = Abs_matrix (row_multiply_infmatrix (Rep_matrix A) i q)"


lemma Rep_abs_aux_row_multiply:
shows "Rep_matrix (Abs_matrix (λa b. if a = i then q * Rep_matrix (A::('a::ring) matrix) i b else Rep_matrix A a b))
=(λa b. if a = i then q * Rep_matrix A i b else Rep_matrix A a b)"

proof (rule RepAbs_matrix)
obtain a where a:"∀j s. a≤j --> Rep_matrix A s j = (0::'a)" by (metis ncols)
obtain b where b:"∀i s. b≤i --> Rep_matrix A i s = (0::'a)" by (metis nrows)
show "∃m. ∀j ia. m ≤ j --> (if j = i then q * Rep_matrix A i ia else Rep_matrix A j ia) = (0::'a)"
proof (rule exI[of _ "max (i+1) (max (j+1) b)"], auto+)
fix ja ia assume "b ≤ ja" thus "Rep_matrix A ja ia = (0::'a)" using b by simp
qed
show "∃n. ∀j ia. n ≤ ia --> (if j = i then q * Rep_matrix A i ia else Rep_matrix A j ia) = (0::'a)"
using exI[of _ "max (i+1) (max (j+1) a)"] a by auto
qed


lemma row_multiply_matrix_eq:
shows "∀s. Rep_matrix (row_multiply_matrix A i q) i s = q*((Rep_matrix A) i s)"

proof
fix s
show "Rep_matrix (row_multiply_matrix A i q) i s = q * Rep_matrix A i s"
unfolding row_multiply_matrix_def row_multiply_infmatrix_def
proof -
have "Rep_matrix (Abs_matrix (λa b. if a = i then q * Rep_matrix A i b else Rep_matrix A a b))
=(λa b. if a = i then q * Rep_matrix A i b else Rep_matrix A a b)"

using Rep_abs_aux_row_multiply .
thus "Rep_matrix (Abs_matrix (λa b. if a = i then q * Rep_matrix A i b else Rep_matrix A a b)) i s = q * Rep_matrix A i s"
by simp
qed
qed


lemma row_multiply_matrix_eq2:
shows "∀s t. s≠i --> Rep_matrix (row_multiply_matrix A i q) s t = (Rep_matrix A) s t"

unfolding row_multiply_matrix_def row_multiply_infmatrix_def
using Rep_abs_aux_row_multiply[of i q A] by simp


definition partRowReduce_infmatrix :: "('a::ring_div) infmatrix => nat => nat => 'a infmatrix"
where "partRowReduce_infmatrix A k l = (λi j. if i>k then (row_add_infmatrix A i k (-((A i l) div (A k l)))) i j else A i j)"


definition partRowReduce_matrix :: "int matrix => nat => nat => int matrix"
where "partRowReduce_matrix A k l = Abs_matrix (partRowReduce_infmatrix (Rep_matrix A) k l)"



lemma aux_partRowReduce:
shows "Rep_matrix (Abs_matrix (λi j. if k < i then row_add_infmatrix
(Rep_matrix (A::('a::ring_div) matrix)) i k (- (Rep_matrix A i l div Rep_matrix A k l)) i j else Rep_matrix A i j))
= (λi j. if k < i then row_add_infmatrix (Rep_matrix A) i k (- (Rep_matrix A i l div Rep_matrix A k l)) i j else Rep_matrix A i j)"

proof (rule RepAbs_matrix)
obtain a where a:"∀j s. a≤j --> Rep_matrix A s j = (0::'a)" by (metis ncols)
obtain b where b:"∀i s. b≤i --> Rep_matrix A i s = (0::'a)" by (metis nrows)
show "∃m. ∀j i. m ≤ j --> (if k < j then row_add_infmatrix (Rep_matrix A) j k (- (Rep_matrix A j l div Rep_matrix A k l)) j i else Rep_matrix A j i) = (0::'a)"
proof (rule exI[of _ "max (k+1) (max (l+1) b)"], auto+)
fix j i assume b_le_ja: "b ≤ j"
show "row_add_infmatrix (Rep_matrix A) j k (- (Rep_matrix A j l div Rep_matrix A k l)) j i = (0::'a)"
unfolding row_add_infmatrix_def
proof (auto)
have 1: "Rep_matrix A j i = 0" using b_le_ja b by simp
have "Rep_matrix A j l = 0" using b_le_ja b by simp
hence "(Rep_matrix A j l div Rep_matrix A k l)=0" using div_0[of "Rep_matrix A k l"] by simp
thus "Rep_matrix A j i + - (Rep_matrix A j l div Rep_matrix A k l * Rep_matrix A k i) = (0::'a)" using 1 by auto
qed
qed
show "∃n. ∀j i. n ≤ i --> (if k < j then row_add_infmatrix (Rep_matrix A) j k (- (Rep_matrix A j l div Rep_matrix A k l)) j i else Rep_matrix A j i) = (0::'a)"
proof (rule exI[of _ "max (k+1) (max (l+1) a)"], auto+)
fix j i assume a_le_i: "a ≤ i" thus "Rep_matrix A j i = (0::'a)" using a by simp
show "row_add_infmatrix (Rep_matrix A) j k (- (Rep_matrix A j l div Rep_matrix A k l)) j i = (0::'a)"
proof (unfold row_add_infmatrix_def, auto)
have 1: "Rep_matrix A j i = 0" using a_le_i a by simp
have "Rep_matrix A k i = 0" using a_le_i a by simp
thus "Rep_matrix A j i + - (Rep_matrix A j l div Rep_matrix A k l * Rep_matrix A k i) = (0::'a)" using 1 by auto
qed
qed
qed


lemma partRowReduce':
shows "∀s t. s≤k --> Rep_matrix (partRowReduce_matrix A k l) s t = (Rep_matrix A) s t"

unfolding partRowReduce_matrix_def partRowReduce_infmatrix_def
using aux_partRowReduce[of k A l] by simp

lemma partRowReduce:
shows "∀i j. i>k --> Rep_matrix (partRowReduce_matrix A k l) i j = row_add_infmatrix (Rep_matrix A) i k (- (Rep_matrix A i l div Rep_matrix A k l)) i j"

unfolding partRowReduce_matrix_def partRowReduce_infmatrix_def
using aux_partRowReduce[of k A l] by simp


lemma mod_abs:
assumes b_not_zero: "b≠0
"
shows "abs ((a::int) mod b) < abs (b)"

proof (cases "b>0")
case True thus ?thesis using mod_le_divisor by simp
next
case False
hence "b<0" using b_not_zero by auto
hence mod_le_zero: "a mod b ≤ 0" and b_l_mod: "b < a mod b"
using neg_mod_conj by simp+
thus ?thesis by auto
qed


lemma partRowReduce_abs:
assumes i_g_k: "i>k"
and Akk_not_zero: "Rep_matrix (A::int matrix) k k ≠ 0"
shows "abs (Rep_matrix (partRowReduce_matrix A k k) i k) < abs ((Rep_matrix A) k k)"

proof -
have "abs (Rep_matrix (partRowReduce_matrix A k k) i k) = abs (row_add_infmatrix (Rep_matrix A) i k (- (Rep_matrix A i k div Rep_matrix A k k)) i k)"
using partRowReduce[of k A k]
using i_g_k
unfolding partRowReduce_matrix_def partRowReduce_infmatrix_def by simp
also have "...=abs (if i = i then Rep_matrix A i k + - (Rep_matrix A i k div Rep_matrix A k k) * Rep_matrix A k k else Rep_matrix A i k)"
unfolding row_add_infmatrix_def ..
also have "...=abs (Rep_matrix A i k + - (Rep_matrix A i k div Rep_matrix A k k) * Rep_matrix A k k)" by simp
also have "... = abs (Rep_matrix A i k mod Rep_matrix A k k)"
by (metis diff_int_def minus_mult_right zmod_zdiv_equality' zmult_commute)
also have "... < abs (Rep_matrix A k k)"
using Akk_not_zero mod_abs[of "Rep_matrix A k k" "Rep_matrix A i k"] by simp
finally show ?thesis .
qed

definition partColumnReduce_infmatrix :: "('a::ring_div) infmatrix => nat => nat => 'a infmatrix"
where "partColumnReduce_infmatrix A k l = (λi j. if j>l then (column_add_infmatrix A j l (-((A k j) div (A k l)))) i j else A i j)"


definition partColumnReduce_matrix :: "int matrix => nat => nat => int matrix"
where "partColumnReduce_matrix A k l = Abs_matrix (partColumnReduce_infmatrix (Rep_matrix A) k l)"


lemma aux_partColumnReduce:
shows "Rep_matrix (Abs_matrix (λi j. if l < j then column_add_infmatrix (Rep_matrix A) j l (- (Rep_matrix A k j div Rep_matrix A k l)) i j else Rep_matrix A i j))
=(λi j. if l < j then column_add_infmatrix (Rep_matrix (A::('a::ring_div) matrix)) j l (- (Rep_matrix A k j div Rep_matrix A k l)) i j else Rep_matrix A i j)"

proof (rule RepAbs_matrix)
obtain a where a:"∀j s. a≤j --> Rep_matrix A s j = (0::'a)" by (metis ncols)
obtain b where b:"∀i s. b≤i --> Rep_matrix A i s = (0::'a)" by (metis nrows)
show "∃m. ∀j i. m ≤ j --> (if l < i then column_add_infmatrix (Rep_matrix A) i l (- (Rep_matrix A k i div Rep_matrix A k l)) j i else Rep_matrix A j i) = (0::'a)"
proof (rule exI[of _ "max (k+1) (max (l+1) b)"], auto+)
fix i j assume b_le_j:"b ≤ j" thus "Rep_matrix A j i = (0::'a)" using b by simp
show "column_add_infmatrix (Rep_matrix A) i l (- (Rep_matrix A k i div Rep_matrix A k l)) j i = (0::'a)"
unfolding column_add_infmatrix_def using b_le_j b by auto
qed
show "∃n. ∀j i. n ≤ i --> (if l < i then column_add_infmatrix (Rep_matrix A) i l (- (Rep_matrix A k i div Rep_matrix A k l)) j i else Rep_matrix A j i) = (0::'a)"
using exI[of _ "max (k+1) (max (l+1) a)"] using a unfolding column_add_infmatrix_def by auto
qed


lemma partColumnReduce':
shows "∀s t. t≤l --> Rep_matrix (partColumnReduce_matrix A k l) s t = (Rep_matrix A) s t"

unfolding partColumnReduce_matrix_def partColumnReduce_infmatrix_def
using aux_partColumnReduce[of l A k] by simp

lemma partColumnReduce:
shows "∀i j. j>l --> Rep_matrix (partColumnReduce_matrix A k l) i j = (column_add_infmatrix (Rep_matrix A) j l (-((Rep_matrix A k j) div (Rep_matrix A k l)))) i j"

unfolding partColumnReduce_matrix_def partColumnReduce_infmatrix_def
using aux_partColumnReduce[of l A k] by simp

lemma partColumnReduce_abs:
assumes j_g_k: "j>k"
and Akk_not_zero: "Rep_matrix (A::int matrix) k k ≠ 0"
shows "abs (Rep_matrix (partColumnReduce_matrix A k k) k j) < abs ((Rep_matrix A) k k)"

proof -
have "abs (Rep_matrix (partColumnReduce_matrix A k k) k j) = abs (column_add_infmatrix (Rep_matrix A) j k (- (Rep_matrix A k j div Rep_matrix A k k)) k j)"
using partColumnReduce [of k A k]
unfolding partColumnReduce_matrix_def partColumnReduce_infmatrix_def
using j_g_k by simp
also have "...=abs (if j = j then Rep_matrix A k j + - (Rep_matrix A k j div Rep_matrix A k k) * Rep_matrix A k k else Rep_matrix A k j)"
unfolding column_add_infmatrix_def ..
also have "...=abs (Rep_matrix A k j + - (Rep_matrix A k j div Rep_matrix A k k) * Rep_matrix A k k)" by simp
also have "...=abs (Rep_matrix A k j mod Rep_matrix A k k)"
by (metis diff_int_def minus_mult_right zmod_zdiv_equality' zmult_commute)
also have "... < abs (Rep_matrix A k k)" using Akk_not_zero mod_abs[of "Rep_matrix A k k" "Rep_matrix A k j"] by simp
finally show ?thesis .
qed

definition fst_bis:: "('a × 'b × 'c)=>'a"
where "fst_bis A = fst A"


lemma fst_bis[simp]:"fst_bis(a,b,c)=a" unfolding fst_bis_def by simp

definition snd_bis :: "('a × 'b × 'c)=>'b"
where "snd_bis A = fst (snd A)"


lemma snd_bis[simp]:"snd_bis(a,b,c)=b" unfolding snd_bis_def by simp

definition trd_bis :: "('a × 'b × 'c)=>'c"
where "trd_bis A = snd (snd A)"


lemma trd_bis[simp]:"trd_bis (a,b,c)=c" unfolding trd_bis_def by simp


instantiation prod :: (ord, ord) ord
begin


definition less_prod :: "'a × 'b => 'a × 'b => bool"
where "less_prod P Q = (fst P < fst Q ∨ ((fst P = fst Q) ∧ (snd P < snd Q)))"


(*BE CAREFUL! We have to write less (or symbol <) and not less_prod*)
definition less_eq_prod :: "'a × 'b => 'a × 'b => bool"
where "less_eq_prod P Q == ((less P Q) ∨ ((fst P = fst Q) ∧ (snd P = snd Q)))"


instance
proof
qed

end

text{*For a given matrix A and a natural number k, this function returns us (False, 0, 0) if the submatrix from position A k k (included) are null. In other case,
this function returns (True, pos1, pos2), where (pos1, pos2) is the position of the element which has the least notzero absolute value*}

definition minNonzero_nc :: "(int matrix) => nat => (bool × nat × nat)"
where "minNonzero_nc A k = (if (∀ i j. i∈{k..<nrows A} ∧ j∈{k..<ncols A} --> Rep_matrix A i j = 0) then (False,0::nat,0::nat)
else (True,LEAST pos. (fst pos)≥k ∧ (snd pos)≥k ∧ Rep_matrix A (fst pos) (snd pos) ≠ 0
∧ (∀m n. m≥k ∧ n≥k ∧ Rep_matrix A m n ≠ 0 --> (abs(Rep_matrix A m n)) ≥ (abs(Rep_matrix A (fst pos) (snd pos))))))"


text{*If the matrix has non zero elements from position A k k, then the boolean that the function minNonzero\_nc A k returns will be True.*}
lemma minNonzero_true:
assumes "¬ (∀(i::nat) j::nat. i ∈ {k..<nrows A} ∧ j ∈ {k..<ncols A} --> Rep_matrix A i j = 0)"
shows "fst_bis (minNonzero_nc A k) = True"

unfolding minNonzero_nc_def fst_bis_def using assms by simp

text{*Analogously, if function (minNonzero\_nc A k) returns True, then not every element are zero in the submatrix from element A k k.*}
lemma minNonzero_true_inv:
assumes "fst_bis (minNonzero_nc A k) = True"
shows "¬ (∀(i::nat) j::nat. i ∈ {k..<nrows A} ∧ j ∈ {k..<ncols A} --> Rep_matrix A i j = 0)"

by (metis assms fst_bis minNonzero_nc_def)


text{*This function moves the element of the submatrix (k:nrows,k:ncols) with the least nonzero absolute value to the position (k,k)*}
definition move_minNonzero_nc :: "(int matrix) => nat => (int matrix)"
where "move_minNonzero_nc A k = (move_element A (snd_bis (minNonzero_nc A k)) (trd_bis (minNonzero_nc A k)) k k)"


partial_function (tailrec) iterate_partRowReduce :: "(int matrix) => nat => (int matrix)"
where "iterate_partRowReduce A k = (if (∀m. m>k --> Rep_matrix A m k = 0) then A else
iterate_partRowReduce (partRowReduce_matrix (move_minNonzero_nc A k) k k) k)"


declare iterate_partRowReduce.simps [simp del]


lemma partRowReduce_abs2:
assumes Akk_not_zero: "Rep_matrix (A::int matrix) k k ≠ 0"
and eq: "abs(Rep_matrix A k k) ≤ abs(Rep_matrix A i k)"
and i:"i>k"
shows "abs (Rep_matrix (partRowReduce_matrix A k k) i k) < abs (Rep_matrix A i k)"

proof -
have "abs (Rep_matrix (partRowReduce_matrix A k k) i k) = abs (row_add_infmatrix (Rep_matrix A) i k (- (Rep_matrix A i k div Rep_matrix A k k)) i k)"
using partRowReduce[of k A k]
using i
unfolding partRowReduce_matrix_def partRowReduce_infmatrix_def by simp
also have "...=abs (if i = i then Rep_matrix A i k + - (Rep_matrix A i k div Rep_matrix A k k) * Rep_matrix A k k else Rep_matrix A i k)"
unfolding row_add_infmatrix_def ..
also have "...=abs (Rep_matrix A i k + - (Rep_matrix A i k div Rep_matrix A k k) * Rep_matrix A k k)" by simp
also have "... = abs (Rep_matrix A i k mod Rep_matrix A k k)"
by (metis diff_int_def minus_mult_right zmod_zdiv_equality' zmult_commute)
also have "... < abs (Rep_matrix A k k)"
using Akk_not_zero mod_abs[of "Rep_matrix A k k" "Rep_matrix A i k"] by simp
also have "...≤abs(Rep_matrix A i k)" using eq .
finally show ?thesis .
qed


instantiation prod :: (order, order) order
begin


instance
proof
fix x :: "('a::order × 'b::order)"
and y :: "('a::order × 'b::order)"
and z :: "('a::order × 'b::order)"

show "(x < y) = (x ≤ y ∧ ¬ y ≤ x)"
by (metis less_eq_prod_def less_prod_def xt1(9))
show "x ≤ x" by (metis less_eq_prod_def)
assume x: "x ≤ y" and y: "y ≤ z"
thus "x ≤ z"
unfolding less_prod_def less_eq_prod_def apply auto by (metis xt1(10))+
assume "x ≤ y" and "y ≤ x"
thus "x = y" by (metis `(x < y) = (x ≤ y ∧ ¬ y ≤ x)` less_eq_prod_def prod_eqI)
qed

end

instantiation prod :: (linorder, linorder) linorder
begin


instance
proof
fix x :: "('a::linorder × 'b::linorder)"
and y :: "('a::linorder × 'b::linorder)"

show "x ≤ y ∨ y ≤ x"
unfolding less_eq_prod_def less_prod_def
by (metis linorder_neq_iff)
qed

end

text{*In the file Wellfounded.thy is has been proved that well orders with
the lexicographic order are themselves well orders, so we just have to rewrite
our definitions of @{term "less"} and @{term "less_eq"} adequetely and then use
the result @{thm "wf_lex_prod"}.*}


instantiation prod :: (wellorder, wellorder) wellorder
begin

find_theorems "less_nat_induct"
instance
proof (rule wf_wellorderI)
have wf1: "wf {(x,y). x < (y::'a)}" and wf2: "wf {(x,y). x < (y::'b)}"
using wf by simp_all
have wf_lex_prod: "wf ({(x,y). x < (y::'a)} <*lex*> {(x,y). x < (y::'b)})"
using wf_lex_prod [OF wf1 wf2] .
show "wf {(x, y). x < (y::'a × 'b)}"
proof -
have rw: "{(x,y). x < (y::'a)} <*lex*> {(x,y). x < (y::'b)} = {(x, y). x < y}"
unfolding lex_prod_def
apply auto
by (metis fst_conv snd_conv less_prod_def)+
show ?thesis using rw using wf_lex_prod by simp
qed
show "OFCLASS('a × 'b, linorder_class)"
by (intro_classes)
qed

end


(*Definitions obtained from http://afp.sourceforge.net/browser_info/devel/HOL/Completeness/Base.html*)
definition
preImage :: "('a => 'b) => 'b set => 'a set" where
"preImage f A = {x . f x ∈ A}"


definition
pre :: "('a => 'b) => 'b => 'a set" where
"pre f a = {x . f x = a}"


definition
equalOn :: "['a set,'a => 'b,'a => 'b] => bool" where
"equalOn A f g = (!x:A. f x = g x)"


lemma preImage_insert: "preImage f (insert a A) = pre f a Un preImage f A"
by(auto simp add: preImage_def pre_def)

lemma preImageI: "f x : A ==> x : preImage f A"
by(simp add: preImage_def)

lemma preImageE: "x : preImage f A ==> f x : A"
by(simp add: preImage_def)

lemma equalOn_Un: "equalOn (A ∪ B) f g = (equalOn A f g ∧ equalOn B f g)"
by(auto simp add: equalOn_def)

lemma equalOnD: "equalOn A f g ==> (∀x ∈ A . f x = g x)"
by(simp add: equalOn_def)

lemma equalOnI:"(∀x ∈ A . f x = g x) ==> equalOn A f g"
by(simp add: equalOn_def)

lemma equalOn_UnD: "equalOn (A ∪ B) f g ==> equalOn A f g ∧ equalOn B f g"
by(auto simp: equalOn_def)

-- "FIXME move following elsewhere?"

lemma inj_inv_singleton[simp]: "[| inj f; f z = y |] ==> {x. f x = y} = {z}"
apply rule
apply(auto simp: inj_on_def) done

lemma finite_pre[simp]: "inj f ==> finite (pre f x)"
apply(simp add: pre_def)
apply (cases "∃y. f y = x", auto) done

lemma finite_preImage[simp]: "[| finite A; inj f |] ==> finite (preImage f A)"
apply(induct A rule: finite_induct)
apply(simp add: preImage_def)
apply(simp add: preImage_insert) done



lemma pre:
assumes l: "∃x∈A. ∀y∈A. P x y"
shows "{x∈A. ∀y∈A. P x y} ≠ {}"

using l by auto

definition Rep_matrix' :: "('a::zero matrix)=>(nat×nat)=>'a"
where "Rep_matrix' A x = Rep_matrix A (fst x) (snd x)"


text{*Next function returns us the absolute values of the nonzero elements of a matrix A*}
definition abs_set :: "(int matrix) => nat set"
where "abs_set A = (natoabso(Rep_matrix' A))` (nonzero_positions (Rep_matrix A))"


text{*If there exist nonzero elements in A, then the set returned by function (abs\_set A) won't be empty.*}
lemma
assumes "nonzero_positions (Rep_matrix A)≠{}"
shows "abs_set A≠{}"

using abs_set_def[of A] assms empty_is_image[of "natoabso(Rep_matrix' A)" "nonzero_positions (Rep_matrix A)"]
by metis

text{*If a element is belongs to (abs\_set A), then it is not a zero.*}
lemma abs_notzero:
assumes x: "x ∈ abs_set A"
shows "x≠0"

using x unfolding abs_set_def nonzero_positions_def apply auto unfolding Rep_matrix'_def by simp

lemma
assumes f:"finite A"
and ne:"(A::('a::{wellorder} set))≠{}"
shows "Least A ∈ A"

proof -
thm mem_def
from ne obtain x where x:"x∈A" by auto
thus ?thesis using LeastI_ex mem_def by metis
qed

lemma nat_times_nat_least:
assumes ne:"(A::(nat × nat) set)≠{}"
shows "Least A ∈ A"

proof -
thm mem_def
from ne obtain x where x:"x∈A" by auto
thus ?thesis using LeastI_ex mem_def by metis
qed

text{*If the set returned by (abs\_set A) is not empty, then the LEAST element belongs to this set. It is indispensable that the set has a well-order.*}
lemma Least_in:
assumes "abs_set A ≠ {}"
shows "Least (abs_set A) ∈ (abs_set A)"

by (metis Collect_def LeastI_ex assms empty_Collect_eq mem_def)

text{*Then, each element belonging to (abs\_set A) is greater than or equal than the LEAST.*}
lemma Least_le_abs_set:
assumes y:"y ∈ abs_set A" --"Con esto ya se que no es vacio"
shows "Least (abs_set A) ≤ y"

using Least_le mem_def y by metis

text{*Now we want to obtain the position of the LEAST element of (abs\_set A). We use the definition of preimage.
We have to take into account that the LEAST could be in several positions, that is, we can have several preimages (we obtain a set of positions).*}

lemma pre_abs_set:
assumes "abs_set A≠{}"
shows "pre (natoabs o Rep_matrix' A) (Least (abs_set A)) ≠ {}"

using pre_def[of "(natoabs o Rep_matrix' A)" "(Least (abs_set A))"]
using image_iff using assms Collect_empty_eq[of "pre (natoabs o Rep_matrix' A) (Least (abs_set A))"] Least_in[of A] abs_set_def[of A] by force

lemma Least_in_pre:
assumes "abs_set A≠{}"
shows "Least (pre (natoabs o Rep_matrix' A) (Least (abs_set A))) ∈ pre (natoabs o Rep_matrix' A) (Least (abs_set A))"

by (metis assms nat_times_nat_least pre_abs_set)


text{*In positions of preimage set, there will be the elements with the least absolute value.*}
lemma pre_is_the_least_abs:
assumes "abs_set A≠{}"
and i:"Rep_matrix A i j ≠ 0"
and x:"x ∈ pre (natoabs o Rep_matrix' A) (Least (abs_set A))"
shows "nat(abs(Rep_matrix' A x)) ≤ nat(abs(Rep_matrix' A (i,j)))"

using x unfolding pre_def apply auto
proof (rule Least_le_abs_set[of " nat ¦Rep_matrix' A (i, j)¦" A])
have "(i,j) ∈ nonzero_positions (Rep_matrix A)" using i unfolding nonzero_positions_def by auto
thus "nat ¦Rep_matrix' A (i, j)¦ ∈ abs_set A"
by (metis abs_set_def imageI image_compose)
qed


lemma
assumes a:"abs_set A≠{}"
and x:"x ∈ pre (natoabs o Rep_matrix' A) (Least (abs_set A))"
shows "x ∈ nonzero_positions (Rep_matrix A)"

using x unfolding pre_def abs_set_def image_def apply auto
proof -
assume p:"nat ¦Rep_matrix' A x¦ = Least {y. ∃x∈nonzero_positions (Rep_matrix A). y = nat ¦Rep_matrix' A x¦}"
have "Least {y. ∃x∈nonzero_positions (Rep_matrix A). y = nat ¦Rep_matrix' A x¦}≠0"
using x unfolding pre_def using p unfolding o_def nonzero_positions_def unfolding mem_def using abs_notzero
using Least_in a unfolding Collect_def by force
thus ?thesis using p unfolding nonzero_positions_def Rep_matrix'_def using abs_zero mem_def nat_0 by auto
qed




text{*Now we have to do the same but for a submatrix obtained from a position (k,k):*}
definition abs_set_k :: "(int matrix) => nat => nat set"
where "abs_set_k A k = (natoabso(Rep_matrix' A))` (nonzero_positions (Rep_matrix A)∩{(i,j). i≥k ∧ j≥k})"


lemma nonzero_implies_abs_k_ne:
assumes "nonzero_positions (Rep_matrix A)∩{(i,j). i≥k ∧ j≥k}≠{}"
shows "abs_set_k A k ≠ {}"

using abs_set_k_def[of A] assms empty_is_image[of "natoabso(Rep_matrix' A)" "nonzero_positions (Rep_matrix A)∩{(i,j). i≥k ∧ j≥k}"]
by metis

lemma abs_k_notzero:
assumes x: "x ∈ abs_set_k A k"
shows "x≠0"

using x unfolding abs_set_k_def nonzero_positions_def apply auto unfolding Rep_matrix'_def by simp


lemma Least_in_abs_k:
assumes "abs_set_k A k ≠ {}"
shows "Least (abs_set_k A k) ∈ (abs_set_k A k)"

by (metis Collect_def LeastI_ex assms empty_Collect_eq mem_def)


lemma Least_le_abs_set_k:
assumes y:"y ∈ abs_set_k A k" --"Con esto ya se que no es vacio"
shows "Least (abs_set_k A k) ≤ y"

using Least_le mem_def y by metis


lemma pre_abs_set_k:
assumes "abs_set_k A k≠{}"
shows "pre (natoabs o Rep_matrix' A) (Least (abs_set_k A k)) ≠ {}"

using pre_def[of "(natoabs o Rep_matrix' A)" "(Least (abs_set_k A k))"]
using image_iff using assms Collect_empty_eq[of "pre (natoabs o Rep_matrix' A) (Least (abs_set_k A k))"]
Least_in_abs_k[of A k] abs_set_k_def[of A k]
by force

lemma pre_abs_set_k':
assumes "abs_set_k A k≠{}"
shows "pre (natoabs o Rep_matrix' A) (Least (abs_set_k A k)) ∩ {(i,j). i≥k ∧ j≥k} ≠ {}"

using pre_def[of "(natoabs o Rep_matrix' A)" "(Least (abs_set_k A k))"]
using image_iff using assms Collect_empty_eq[of "pre (natoabs o Rep_matrix' A) (Least (abs_set_k A k))∩ {(i,j). i≥k ∧ j≥k}"]
using Least_in_abs_k[of A k] abs_set_k_def[of A k]
by force

lemma Least_in_pre_abs_k:
assumes "abs_set_k A k≠{}"
shows "Least (pre (natoabs o Rep_matrix' A) (Least (abs_set_k A k))) ∈ pre (natoabs o Rep_matrix' A) (Least (abs_set_k A k))"

by (metis assms nat_times_nat_least pre_abs_set_k)

lemma Least_in_pre_abs_k':
assumes "abs_set_k A k≠{}"
shows "Least (pre (natoabs o Rep_matrix' A) (Least (abs_set_k A k)) ∩ {(i,j). i≥k ∧ j≥k})
∈ (pre (natoabs o Rep_matrix' A) (Least (abs_set_k A k))) ∩ {(i,j). i≥k ∧ j≥k}"

by (metis assms nat_times_nat_least pre_abs_set_k')

lemma pre_is_the_least_abs_k:
assumes "abs_set_k A k≠{}"
and i:"Rep_matrix A i j ≠ 0"
and i2:"i≥k"
and i3: "j≥k"
and x:"x ∈ pre (natoabs o Rep_matrix' A) (Least (abs_set_k A k))"
shows "nat(abs(Rep_matrix' A x)) ≤ nat(abs(Rep_matrix' A (i,j)))"

using x unfolding pre_def apply auto proof (rule Least_le_abs_set_k[of " nat ¦Rep_matrix' A (i, j)¦" A k])
have "(i,j) ∈ nonzero_positions (Rep_matrix A)∩{(a,b). a≥k ∧ b≥k}" using i i2 i3 unfolding nonzero_positions_def by auto
thus "nat ¦Rep_matrix' A (i, j)¦ ∈ abs_set_k A k" unfolding abs_set_k_def by auto
qed

lemma pre_is_the_least_abs_k':
assumes abs:"abs_set_k A k≠{}"
and i:"Rep_matrix A i j ≠ 0"
and i2:"i≥k"
and i3: "j≥k"
and x:"x ∈ pre (natoabs o Rep_matrix' A) (Least (abs_set_k A k))∩{(i,j).i≥k ∧ j≥k}"
shows "nat(abs(Rep_matrix' A x)) ≤ nat(abs(Rep_matrix' A (i,j)))"

proof -
have "x ∈ pre (natoabs o Rep_matrix' A) (Least (abs_set_k A k))" using x by simp
thus ?thesis using pre_is_the_least_abs_k[OF abs i i2 i3 _] by auto
qed

lemma aux_nonzero:
assumes "y∈{y. ∃x∈nonzero_positions (Rep_matrix A). y = nat ¦Rep_matrix' A x¦}"
shows "y≠0"

using assms unfolding nonzero_positions_def Rep_matrix'_def by auto

lemma aux_least: "(Least A) = (LEAST x. x∈A)"
by (metis Collect_def Collect_disj_eq Un_absorb Un_def)

lemma pre_in_nonzero_k:
assumes a:"abs_set_k A k≠{}"
and x:"x ∈ pre (natoabs o Rep_matrix' A) (Least (abs_set_k A k)) ∩ {(i,j). i≥k ∧ j≥k}"
shows "x ∈ nonzero_positions (Rep_matrix A)∩{(i,j). i≥k ∧ j ≥ k}"

proof -
have x1: "x ∈ {(i,j). i≥k ∧ j ≥ k}" using x by simp
have x2: "x ∈ pre (natoabs o Rep_matrix' A) (Least (abs_set_k A k))" using x by simp
have "x ∈ nonzero_positions (Rep_matrix A)"
using x2 unfolding pre_def abs_set_k_def image_def apply auto
proof -
assume p: "nat ¦Rep_matrix' A x¦ = Least {y. ∃x∈nonzero_positions (Rep_matrix A) ∩ Collect (op ≤ k) × Collect (op ≤ k). y = nat ¦Rep_matrix' A x¦}"
have "nonzero_positions (Rep_matrix A) ∩ Collect (op ≤ k) × Collect (op ≤ k) ≠ {}" using a unfolding abs_set_k_def image_def by auto
hence ne: "{y. ∃x∈nonzero_positions (Rep_matrix A) ∩ Collect (op ≤ k) × Collect (op ≤ k). y = nat ¦Rep_matrix' A x¦}≠{}" by auto
from this obtain x where x3: "x∈{y. ∃x∈nonzero_positions (Rep_matrix A) ∩ Collect (op ≤ k) × Collect (op ≤ k). y = nat ¦Rep_matrix' A x¦}" by auto
have "{y. ∃x∈nonzero_positions (Rep_matrix A) ∩ Collect (op ≤ k) × Collect (op ≤ k). y = nat ¦Rep_matrix' A x¦} ⊆
{y. ∃x∈nonzero_positions (Rep_matrix A). y = nat ¦Rep_matrix' A x¦}"
by auto
hence "!!y. y∈{y. ∃x∈nonzero_positions (Rep_matrix A) ∩ Collect (op ≤ k) × Collect (op ≤ k). y = nat ¦Rep_matrix' A x¦}
==> y ∈ {y. ∃x∈nonzero_positions (Rep_matrix A). y = nat ¦Rep_matrix' A x¦}"
by auto
hence p2:"!!y. y∈{y. ∃x∈nonzero_positions (Rep_matrix A) ∩ Collect (op ≤ k) × Collect (op ≤ k). y = nat ¦Rep_matrix' A x¦} ==> y≠0"
using aux_nonzero by auto
have "Least{y. ∃x∈nonzero_positions (Rep_matrix A) ∩ Collect (op ≤ k) × Collect (op ≤ k). y = nat ¦Rep_matrix' A x¦}
∈ {y. ∃x∈nonzero_positions (Rep_matrix A) ∩ Collect (op ≤ k) × Collect (op ≤ k). y = nat ¦Rep_matrix' A x¦}"

using LeastI[of "{y. ∃x∈nonzero_positions (Rep_matrix A) ∩ Collect (op ≤ k) × Collect (op ≤ k). y = nat ¦Rep_matrix' A x¦}" x]
using x3
unfolding aux_least[of "{y. ∃x∈nonzero_positions (Rep_matrix A) ∩ Collect (op ≤ k) × Collect (op ≤ k). y = nat ¦Rep_matrix' A x¦}"]
unfolding mem_def[of x "{y. ∃x∈nonzero_positions (Rep_matrix A) ∩ Collect (op ≤ k) × Collect (op ≤ k). y = nat ¦Rep_matrix' A x¦}"]
unfolding mem_def[of "(LEAST x. x ∈ {y. ∃x∈nonzero_positions (Rep_matrix A) ∩ Collect (op ≤ k) × Collect (op ≤ k). y = nat ¦Rep_matrix' A x¦})" ]
by simp
hence "Least {y. ∃x∈nonzero_positions (Rep_matrix A) ∩ Collect (op ≤ k) × Collect (op ≤ k). y = nat ¦Rep_matrix' A x¦} ≠ 0" using p2 by auto
thus ?thesis using x2 unfolding pre_def o_def image_def abs_set_k_def unfolding nonzero_positions_def Rep_matrix'_def using abs_zero mem_def nat_0 by auto
qed
thus ?thesis using x1 by simp
qed

lemma nat_abs_int:
assumes "(nat(abs(a::int))) ≤ nat(abs(b::int))"
shows "abs(a)≤ abs(b)"

by (metis abs_ge_zero assms nat_le_eq_zle)

lemma nat_abs_int':
assumes "abs(a)≤ abs(b)"
shows "(nat(abs(a::int))) ≤ nat(abs(b::int))"

by (metis assms nat_mono)

lemma minNonzero_least:
assumes P_def: "P=(λpos. k ≤ fst pos ∧ k ≤ snd pos ∧ Rep_matrix A (fst pos) (snd pos)≠0 ∧
(∀m n. k ≤ m ∧ k ≤ n ∧ Rep_matrix A m n ≠ 0 --> ¦Rep_matrix A (fst pos) (snd pos)¦ ≤ ¦Rep_matrix A m n¦))"

and ex_minNonzero: "fst_bis (minNonzero_nc (A::int matrix) k) = True"
shows "P (LEAST k. P k)"

proof -
have notall: "¬ (∀i j. i ∈ {k..<nrows A} ∧ j ∈ {k..<ncols A} --> Rep_matrix A i j = 0)"
using minNonzero_true_inv[OF ex_minNonzero] .
from this obtain i j where i:"i ∈ {k..<nrows A}" and j:"j ∈ {k..<ncols A}" and nz:"Rep_matrix A i j ≠ 0" by auto
have ex:"∃x. P x" proof (rule exI[of _ "Least (pre (natoabs o Rep_matrix' A) (Least (abs_set_k A k)) ∩ {(i,j). i≥k ∧ j≥k})"], unfold P_def, auto)
have abs:"abs_set_k A k≠{}"
proof (rule nonzero_implies_abs_k_ne)
have 1:"(i,j) ∈ nonzero_positions (Rep_matrix A)" using nz unfolding nonzero_positions_def by auto
have 2:"(i,j)∈{(a, b). k ≤ a ∧ k ≤ b}" using i j by auto
show "nonzero_positions (Rep_matrix A) ∩ {(a, b). k ≤ a ∧ k ≤ b} ≠ {}" using 1 and 2 by auto
qed
show k1:"k ≤ fst (Least (pre (nat o abs o Rep_matrix' A) (Least (abs_set_k A k)) ∩ {i. k ≤ i} × {j. k ≤ j}))"
using pre_in_nonzero_k[OF abs Least_in_pre_abs_k'[OF abs]] unfolding nonzero_positions_def by auto
show k2: "k ≤ snd (Least (pre (nat o abs o Rep_matrix' A) (Least (abs_set_k A k)) ∩ {i. k ≤ i} × {j. k ≤ j}))"
using pre_in_nonzero_k[OF abs Least_in_pre_abs_k'[OF abs]] unfolding nonzero_positions_def
by auto
show "Rep_matrix A (fst (Least (pre (nat o abs o Rep_matrix' A) (Least (abs_set_k A k)) ∩ {i. k ≤ i} × {j. k ≤ j})))
(snd (Least (pre (nat o abs o Rep_matrix' A) (Least (abs_set_k A k)) ∩ {i. k ≤ i} × {j. k ≤ j}))) =
0 ==> False"
using pre_in_nonzero_k[OF abs Least_in_pre_abs_k'[OF abs]] unfolding nonzero_positions_def by auto
have "!!m n. [|k ≤ m; k ≤ n; Rep_matrix A m n ≠ 0|]
==> nat¦Rep_matrix A (fst (Least (pre (nat o abs o Rep_matrix' A) (Least (abs_set_k A k)) ∩ {i. k ≤ i} × {j. k ≤ j})))
(snd (Least (pre (nat o abs o Rep_matrix' A) (Least (abs_set_k A k)) ∩ {i. k ≤ i} × {j. k ≤ j})))¦
≤ nat¦Rep_matrix A m n¦"

using pre_is_the_least_abs_k'[OF abs _ _ _ Least_in_pre_abs_k'[OF abs]] unfolding Rep_matrix'_def by auto
thus "!!m n. [|k ≤ m; k ≤ n; Rep_matrix A m n ≠ 0|]
==> ¦Rep_matrix A (fst (Least (pre (nat o abs o Rep_matrix' A) (Least (abs_set_k A k)) ∩ {i. k ≤ i} × {j. k ≤ j})))
(snd (Least (pre (nat o abs o Rep_matrix' A) (Least (abs_set_k A k)) ∩ {i. k ≤ i} × {j. k ≤ j})))¦
≤ ¦Rep_matrix A m n¦"
using nat_abs_int by auto
qed
thus ?thesis using LeastI_ex[of "P"] by fast
qed

text{*We move the element with the least nonzero absolute value to position (k,k). Then, we will have there a non zero element.*}
lemma move_minNonzero_not_zero:
assumes ex_minNonzero: "fst_bis (minNonzero_nc A k) = True"
shows "Rep_matrix (move_minNonzero_nc A k) k k ≠ 0"

proof -
have notall: "¬ (∀(i::nat) j::nat. i ∈ {k..<nrows A} ∧ j ∈ {k..<ncols A} --> Rep_matrix A i j = 0)"
using minNonzero_true_inv[OF ex_minNonzero] .
def P=="(λpos. k ≤ fst pos ∧ k ≤ snd pos ∧ Rep_matrix A (fst pos) (snd pos) ≠ 0 ∧
(∀m n. k ≤ m ∧ k ≤ n ∧ Rep_matrix A m n ≠ 0 --> ¦Rep_matrix A (fst pos) (snd pos)¦ ≤ ¦Rep_matrix A m n¦))"

def pos=="LEAST k. P k"
have "P (pos)" using minNonzero_least[of P k A] ex_minNonzero unfolding P_def pos_def by auto
hence not_zero: "Rep_matrix A (fst pos) (snd pos) ≠ 0" unfolding P_def by simp --"NO SE PUEDE HACER CALCULATION CON EL SIMBOLO (NOT EQUAL) porque lo trata como bool."
have "Rep_matrix (move_minNonzero_nc A k) k k = Rep_matrix (move_element A (snd_bis (minNonzero_nc A k)) (trd_bis (minNonzero_nc A k)) k k) k k"
unfolding move_minNonzero_nc_def ..
also have "...=Rep_matrix A (snd_bis (minNonzero_nc A k)) (trd_bis (minNonzero_nc A k))" using move_element .
also have "...=Rep_matrix A (fst pos) (snd pos)" using notall unfolding minNonzero_nc_def snd_bis_def trd_bis_def pos_def P_def by auto
finally have "Rep_matrix (move_minNonzero_nc A k) k k=Rep_matrix A (fst pos) (snd pos)" .
thus ?thesis using not_zero by simp
qed

text{*That element will be less than any other nonzero element (in absolute value)*}
lemma minNonzero:
assumes i:"i≥k" and j:"j≥k"
and ex_minNonzero: "fst_bis (minNonzero_nc (A::int matrix) k) = True" --"No se puede quitar y hacer por casos, solo si k = 0"
and ij_not_zero: "Rep_matrix A i j ≠ 0"
shows "abs(Rep_matrix A (snd_bis(minNonzero_nc A k)) (trd_bis(minNonzero_nc A k))) ≤ abs(Rep_matrix A i j)"

proof -
have notall: "¬ (∀(i::nat) j::nat. i ∈ {k..<nrows A} ∧ j ∈ {k..<ncols A} --> Rep_matrix A i j = 0)"
using minNonzero_true_inv[OF ex_minNonzero] .
def P=="(λpos. k ≤ fst pos ∧ k ≤ snd pos ∧ Rep_matrix A (fst pos) (snd pos) ≠ 0 ∧
(∀m n. k ≤ m ∧ k ≤ n ∧ Rep_matrix A m n ≠ 0 --> ¦Rep_matrix A (fst pos) (snd pos)¦ ≤ ¦Rep_matrix A m n¦))"

def pos=="LEAST k. P k"
have P_pos: "P (pos)" using minNonzero_least[OF _ ex_minNonzero] P_def pos_def by auto
thus ?thesis using P_pos unfolding P_def pos_def snd_bis_def trd_bis_def minNonzero_nc_def
using notall i j ij_not_zero by auto
qed


lemma move_minNonzero:
assumes i:"i≥k" and j:"j≥k"
and Aik_not_zero: "Rep_matrix (A::int matrix) i j ≠ 0"
shows "abs(Rep_matrix (move_minNonzero_nc A k) k k) ≤ abs(Rep_matrix A i j)"

proof -
have i':"i<nrows A" using Aik_not_zero nrows_notzero by auto
have j':"j<ncols A" using Aik_not_zero ncols_notzero by auto
have notall: "¬ (∀(i::nat) j::nat. i ∈ {k..<nrows A} ∧ j ∈ {k..<ncols A} --> Rep_matrix A i j = 0)"
using i j i' j' Aik_not_zero by force
have "abs(Rep_matrix A (snd_bis(minNonzero_nc A k)) (trd_bis(minNonzero_nc A k))) ≤ abs(Rep_matrix A i j)"
proof (rule minNonzero)
show "k ≤ i" using i .
show "k ≤ j" using j .
show "fst_bis (minNonzero_nc A k) = True" using minNonzero_true[OF notall] .
show "Rep_matrix A i j ≠ 0" using Aik_not_zero .
qed
thus ?thesis
by (metis move_element move_minNonzero_nc_def)
qed

lemma abs_part_min_strict_decreasing':
assumes Aik_not_zero: "Rep_matrix (A::int matrix) i k ≠ 0"
and i:"i>k"
shows "abs (Rep_matrix (partRowReduce_matrix (move_minNonzero_nc A k) k k) i k) < abs (Rep_matrix A i k)"

proof -
have "abs (Rep_matrix (partRowReduce_matrix (move_minNonzero_nc A k) k k) i k) < abs (Rep_matrix (move_minNonzero_nc A k) k k)"
proof (rule partRowReduce_abs)
have i_l_nrows:"i<nrows A" using Aik_not_zero nrows_notzero by auto
have k_l_ncols: "k<ncols A" using Aik_not_zero ncols_notzero by auto
show "k < i" using i .
show "Rep_matrix (move_minNonzero_nc A k) k k ≠ 0"
using move_minNonzero_not_zero[of A k]
using minNonzero_true[of k A] Aik_not_zero i_l_nrows k_l_ncols i by fastforce
qed
also have "...≤ abs(Rep_matrix A i k)"
using move_minNonzero i Aik_not_zero by simp
finally show ?thesis .
qed


lemma abs_part_min_strict_decreasing_ik_kk:
assumes Akk_not_zero: "Rep_matrix (A::int matrix) k k ≠ 0"
and i:"i>k"
shows "abs (Rep_matrix (partRowReduce_matrix (move_minNonzero_nc A k) k k) i k) < abs (Rep_matrix A k k)"

proof -
have "abs (Rep_matrix (partRowReduce_matrix (move_minNonzero_nc A k) k k) i k) < abs (Rep_matrix (move_minNonzero_nc A k) k k)"
proof (rule partRowReduce_abs)
have k_l_nrows:"k<nrows A" using Akk_not_zero nrows_notzero by auto
have k_l_ncols: "k<ncols A" using Akk_not_zero ncols_notzero by auto
show "k < i" using i .
show "Rep_matrix (move_minNonzero_nc A k) k k ≠ 0"
using move_minNonzero_not_zero[of A k]
using minNonzero_true[of k A] Akk_not_zero k_l_nrows k_l_ncols i by fastforce
qed
also have "...≤abs(Rep_matrix A k k)"
using move_minNonzero Akk_not_zero by simp
finally show ?thesis .
qed

lemma Akk_strict_decreasing_move_minNonzero:
assumes i: "i>k"
and menor: "abs(Rep_matrix (A::int matrix) i k) < abs(Rep_matrix A k k)"
and Aik_not_zero: "Rep_matrix A i k ≠ 0"
shows "abs (Rep_matrix (move_minNonzero_nc A k) k k) < abs (Rep_matrix A k k)"

proof -
have "abs (Rep_matrix (move_minNonzero_nc A k) k k) ≤ abs(Rep_matrix A i k)"
using move_minNonzero i Aik_not_zero by simp
also have "...<abs(Rep_matrix A k k)" using menor .
finally show ?thesis .
qed


lemma Akk_strict_decreasing_move_minNonzero':
assumes i: "i>k"
and j: "j≥k"
and menor: "abs(Rep_matrix (A::int matrix) i j) < abs(Rep_matrix A k k)"
and Aij_not_zero: "Rep_matrix A i j ≠ 0"
shows "abs (Rep_matrix (move_minNonzero_nc A k) k k) < abs (Rep_matrix A k k)"

proof -
have "abs (Rep_matrix (move_minNonzero_nc A k) k k) ≤ abs(Rep_matrix A i j)"
using move_minNonzero i j Aij_not_zero by simp
also have "...<abs(Rep_matrix A k k)" using menor .
finally show ?thesis .
qed


lemma Akk_strict_decreasing_move_minNonzero2:
assumes i: "i≥k"
and j: "j>k"
and menor: "abs(Rep_matrix (A::int matrix) i j) < abs(Rep_matrix A k k)"
and Aij_not_zero: "Rep_matrix A i j ≠ 0"
shows "abs (Rep_matrix (move_minNonzero_nc A k) k k) < abs (Rep_matrix A k k)"

proof -
have "abs (Rep_matrix (move_minNonzero_nc A k) k k) ≤ abs(Rep_matrix A i j)"
using move_minNonzero i j Aij_not_zero by simp
also have "...<abs(Rep_matrix A k k)" using menor .
finally show ?thesis .
qed


lemma abs_part_min_strict_decreasing_kk:
assumes Akk_not_zero: "Rep_matrix (A::int matrix) k k ≠ 0"
and "i>k"
and "Rep_matrix A i k ≠ 0"
and menor: "abs(Rep_matrix (A::int matrix) i k) < abs(Rep_matrix A k k)"
shows "abs (Rep_matrix (partRowReduce_matrix (move_minNonzero_nc A k) k k) k k) < abs (Rep_matrix A k k)"

by (metis Akk_strict_decreasing_move_minNonzero assms(2) assms(3) menor order_refl partRowReduce')


lemma partRowReduce_eq_abs:
"¦Rep_matrix (partRowReduce_matrix A k k) k k¦ = ¦Rep_matrix A k k¦"

by (metis eq_imp_le partRowReduce')

lemma abs_part_move_ik_kk:
assumes i: "i>k"
and t: "fst_bis(minNonzero_nc (A::int matrix) k)=True"
shows "abs (Rep_matrix (partRowReduce_matrix (move_minNonzero_nc A k) k k) i k)
< abs (Rep_matrix (partRowReduce_matrix (move_minNonzero_nc A k) k k) k k)"

proof -
have "abs (Rep_matrix (partRowReduce_matrix (move_minNonzero_nc A k) k k) i k) < abs (Rep_matrix (move_minNonzero_nc A k) k k)"
proof (rule partRowReduce_abs)
show "k < i" using i .
show "Rep_matrix (move_minNonzero_nc A k) k k ≠ 0"
using move_minNonzero_not_zero[OF t] .
qed
also have "...=abs (Rep_matrix (partRowReduce_matrix (move_minNonzero_nc A k) k k) k k)"
by (metis le_refl partRowReduce')
finally show ?thesis .
qed

definition f' :: "(int matrix × nat) => nat"
where "f' A = nat(abs(Rep_matrix (move_minNonzero_nc (fst A) (snd A)) (snd A) (snd A)))"


lemma iterate_partRowReduce:
fixes A:: "int matrix"
assumes Akk_not_zero: "Rep_matrix A k k ≠ 0"
shows "∀i>k. Rep_matrix (iterate_partRowReduce A k) i k = 0"

using Akk_not_zero
proof (induct "(A,k)" arbitrary: A k rule: measure_induct_rule [of f'])
fix A k assume h1: "!!Aa ka. [|f' (Aa, ka) < f' (A, k); Rep_matrix Aa ka ka ≠ 0|] ==> ∀i>ka. Rep_matrix (iterate_partRowReduce Aa ka) i ka = 0"
and h2: "Rep_matrix A k k ≠ 0"

show "∀i>k. Rep_matrix (iterate_partRowReduce A k) i k = 0"
proof (cases "∀i>k. Rep_matrix A i k = 0")
case True show ?thesis apply (subst iterate_partRowReduce.simps) using True by auto
next
case False
have "∀i>k. Rep_matrix (iterate_partRowReduce (partRowReduce_matrix(move_minNonzero_nc A k) k k) k) i k = 0"
proof (cases "∀i>k. Rep_matrix (partRowReduce_matrix(move_minNonzero_nc A k) k k) i k = 0")
case True thus ?thesis using iterate_partRowReduce.simps by auto
next
case False
from this obtain i where i:"i>k" and nz: "Rep_matrix (partRowReduce_matrix(move_minNonzero_nc A k) k k) i k ≠ 0" by blast
show ?thesis
proof (rule h1)
show part_move_nonzero: "Rep_matrix (partRowReduce_matrix (move_minNonzero_nc A k) k k) k k ≠ 0"
proof -
have "Rep_matrix (partRowReduce_matrix (move_minNonzero_nc A k) k k) k k = Rep_matrix (move_minNonzero_nc A k) k k"
by (metis le_refl partRowReduce')
also have "...≠0"
proof (rule move_minNonzero_not_zero, rule minNonzero_true, auto, rule exI[of _ k])
have k1: "k < nrows A" using h2
by (metis nrows_notzero)
have k2:"(∃j≥k. j < ncols A ∧ Rep_matrix A k j ≠ 0)"
proof (rule exI[of _ k])
show "k ≤ k ∧ k < ncols A ∧ Rep_matrix A k k ≠ 0" using h2 ncols_notzero by auto
qed
thus "k ≤ k ∧ k < nrows A ∧ (∃j≥k. j < ncols A ∧ Rep_matrix A k j ≠ 0)" using k1 by auto
qed
finally show ?thesis .
qed
show "f' (partRowReduce_matrix (move_minNonzero_nc A k) k k, k) < f' (A, k)"
proof (unfold f'_def, auto)
show "Rep_matrix (move_minNonzero_nc A k) k k = 0 ==> False" using part_move_nonzero
by (metis le_refl partRowReduce')
show "¦Rep_matrix (move_minNonzero_nc (partRowReduce_matrix (move_minNonzero_nc A k) k k) k) k k¦ < ¦Rep_matrix (move_minNonzero_nc A k) k k¦"
proof -
have "¦Rep_matrix (move_minNonzero_nc (partRowReduce_matrix (move_minNonzero_nc A k) k k) k) k k¦
< ¦Rep_matrix (partRowReduce_matrix (move_minNonzero_nc A k) k k) k k¦"

proof (rule Akk_strict_decreasing_move_minNonzero)
show "k < i" using i .
show "¦Rep_matrix (partRowReduce_matrix (move_minNonzero_nc A k) k k) i k¦ < ¦Rep_matrix (partRowReduce_matrix (move_minNonzero_nc A k) k k) k k¦"
by (metis `Rep_matrix (move_minNonzero_nc A k) k k = 0 ==> False` i partRowReduce_abs partRowReduce_eq_abs)
show "Rep_matrix (partRowReduce_matrix (move_minNonzero_nc A k) k k) i k ≠ 0" using nz .
qed
also have "...= ¦Rep_matrix (move_minNonzero_nc A k) k k¦" by (metis partRowReduce_eq_abs)
finally show ?thesis .
qed
qed
qed
qed
thus ?thesis using False iterate_partRowReduce.simps[of A k] by auto
qed
qed

lemma iterate_decrease_kk:
assumes "¬(∀i>k. Rep_matrix A i k = 0)"
and "¬(∀i>k. Rep_matrix (partRowReduce_matrix(move_minNonzero_nc A k) k k) i k = 0)"
shows "¦Rep_matrix (iterate_partRowReduce A k) k k¦ < ¦Rep_matrix (move_minNonzero_nc A k) k k¦"

using assms
proof (induct "(A,k)" arbitrary: A k rule: measure_induct_rule [of f'])
fix A k
assume h1 : "!!Aa ka. [|f' (Aa, ka) < f' (A, k); ¬ (∀i>ka. Rep_matrix Aa i ka = 0); ¬ (∀i>ka. Rep_matrix (partRowReduce_matrix (move_minNonzero_nc Aa ka) ka ka) i ka = 0)|]
==> ¦Rep_matrix (iterate_partRowReduce Aa ka) ka ka¦ < ¦Rep_matrix (move_minNonzero_nc Aa ka) ka ka¦"

and nz1: "¬ (∀i>k. Rep_matrix A i k = 0)" and nz2:"¬(∀i>k. Rep_matrix (partRowReduce_matrix (move_minNonzero_nc A k) k k) i k = 0)"

show "¦Rep_matrix (iterate_partRowReduce A k) k k¦ < ¦Rep_matrix (move_minNonzero_nc A k) k k¦"
proof (cases "∀i>k. Rep_matrix (partRowReduce_matrix (move_minNonzero_nc (partRowReduce_matrix (move_minNonzero_nc A k) k k) k) k k) i k = 0")
case True
from nz2 obtain i where i: "i>k" and nz3: "Rep_matrix (partRowReduce_matrix (move_minNonzero_nc A k) k k) i k ≠ 0" by blast
have "¦Rep_matrix (iterate_partRowReduce A k) k k¦
= abs(Rep_matrix (partRowReduce_matrix (move_minNonzero_nc (partRowReduce_matrix (move_minNonzero_nc A k) k k) k) k k) k k)"

using True iterate_partRowReduce.simps nz1 nz2 by auto
also have "...= abs(Rep_matrix (move_minNonzero_nc (partRowReduce_matrix (move_minNonzero_nc A k) k k) k) k k)" by (metis partRowReduce_eq_abs)
also have "...< abs(Rep_matrix(partRowReduce_matrix (move_minNonzero_nc A k) k k) k k)"
proof (rule Akk_strict_decreasing_move_minNonzero)
show "k < i" using i .
show "¦Rep_matrix (partRowReduce_matrix (move_minNonzero_nc A k) k k) i k¦ < ¦Rep_matrix (partRowReduce_matrix (move_minNonzero_nc A k) k k) k k¦"
proof (rule abs_part_move_ik_kk)
show "k < i" using i .
show "fst_bis (minNonzero_nc A k) = True"
proof (rule minNonzero_true)
from nz1 obtain j where j: "j>k" and nz4: "Rep_matrix A j k ≠ 0" by blast
have j2: "j<nrows A" using nz4 nrows_notzero by auto
have k2: "k<ncols A"using nz4 ncols_notzero by auto
show "¬ (∀i j. i ∈ {k..<nrows A} ∧ j ∈ {k..<ncols A} --> Rep_matrix A i j = 0)" using nz4 j j2 k2 by force
qed
qed
show "Rep_matrix (partRowReduce_matrix (move_minNonzero_nc A k) k k) i k ≠ 0" using nz3 .
qed
also have "...=abs(Rep_matrix(move_minNonzero_nc A k) k k)" by (metis partRowReduce_eq_abs)
finally show ?thesis .
next
case False
have " ¦Rep_matrix (iterate_partRowReduce (partRowReduce_matrix(move_minNonzero_nc A k) k k) k) k k¦
< ¦Rep_matrix (move_minNonzero_nc (partRowReduce_matrix(move_minNonzero_nc A k) k k) k) k k¦"

proof (rule h1, unfold f'_def, auto)
show "∃i>k. Rep_matrix (partRowReduce_matrix (move_minNonzero_nc (partRowReduce_matrix (move_minNonzero_nc A k) k k) k) k k) i k ≠ 0" using False by auto
show "∃i>k. Rep_matrix (partRowReduce_matrix (move_minNonzero_nc A k) k k) i k ≠ 0" using nz2 by auto
show move_notzero: "Rep_matrix (move_minNonzero_nc A k) k k = 0 ==> False"
proof -
assume eq_0: "Rep_matrix (move_minNonzero_nc A k) k k = 0"
have "Rep_matrix (move_minNonzero_nc A k) k k ≠ 0"
proof (rule move_minNonzero_not_zero, rule minNonzero_true)
from nz1 obtain j where j: "j>k" and nz4: "Rep_matrix A j k ≠ 0" by blast
have j2: "j<nrows A" using nz4 nrows_notzero by auto
have k2: "k<ncols A"using nz4 ncols_notzero by auto
show "¬ (∀i j. i ∈ {k..<nrows A} ∧ j ∈ {k..<ncols A} --> Rep_matrix A i j = 0)" using nz4 j j2 k2 by force
qed
thus ?thesis using eq_0 by contradiction
qed
show "¦Rep_matrix (move_minNonzero_nc (partRowReduce_matrix (move_minNonzero_nc A k) k k) k) k k¦ < ¦Rep_matrix (move_minNonzero_nc A k) k k¦"
proof -
from nz2 obtain i where i:"i>k" and nz3: "Rep_matrix (partRowReduce_matrix (move_minNonzero_nc A k) k k) i k ≠ 0" by blast
have "¦Rep_matrix (move_minNonzero_nc (partRowReduce_matrix (move_minNonzero_nc A k) k k) k) k k¦
< ¦Rep_matrix (partRowReduce_matrix (move_minNonzero_nc A k) k k) k k¦"

proof (rule Akk_strict_decreasing_move_minNonzero)
show "i>k" using i .
show " ¦Rep_matrix (partRowReduce_matrix (move_minNonzero_nc A k) k k) i k¦ < ¦Rep_matrix (partRowReduce_matrix (move_minNonzero_nc A k) k k) k k¦"
by (metis `Rep_matrix (move_minNonzero_nc A k) k k = 0 ==> False` i partRowReduce_abs partRowReduce_eq_abs)
show "Rep_matrix (partRowReduce_matrix (move_minNonzero_nc A k) k k) i k ≠ 0" using nz3 .
qed
also have "...= ¦Rep_matrix (move_minNonzero_nc A k) k k¦" using partRowReduce_eq_abs by auto
finally show ?thesis .
qed
qed
also have "...≤¦Rep_matrix (partRowReduce_matrix (move_minNonzero_nc A k) k k) k k¦"
proof (rule move_minNonzero, auto)
have "Rep_matrix (partRowReduce_matrix (move_minNonzero_nc A k) k k) k k = Rep_matrix (move_minNonzero_nc A k) k k" using partRowReduce' by force
also have "...≠0"
proof (rule move_minNonzero_not_zero, rule minNonzero_true)
from nz1 obtain j where j: "j>k" and nz4: "Rep_matrix A j k ≠ 0" by blast
have j2: "j<nrows A" using nz4 nrows_notzero by auto
have k2: "k<ncols A"using nz4 ncols_notzero by auto
show "¬ (∀i j. i ∈ {k..<nrows A} ∧ j ∈ {k..<ncols A} --> Rep_matrix A i j = 0)" using nz4 j j2 k2 by force
qed
finally show "Rep_matrix (partRowReduce_matrix (move_minNonzero_nc A k) k k) k k = 0 ==> False" by auto
qed
also have "...≤¦Rep_matrix (move_minNonzero_nc A k) k k¦" using partRowReduce'[of k "(move_minNonzero_nc A k)"] by auto
finally show ?thesis using nz1 iterate_partRowReduce.simps[of A k] by auto
qed
qed


lemma iterate_decrease_kk':
assumes "Rep_matrix A k k ≠ 0"
shows "¦Rep_matrix (iterate_partRowReduce A k) k k¦ ≤ ¦Rep_matrix A k k¦"

proof (cases "∀i>k. Rep_matrix A i k = 0")
case True thus ?thesis using iterate_partRowReduce.simps by auto
next
case False note nz_Aik=False
hence i:"(iterate_partRowReduce A k)=(iterate_partRowReduce (partRowReduce_matrix (move_minNonzero_nc A k) k k) k)" using iterate_partRowReduce.simps[of A k] by auto
show ?thesis
proof (cases "∀i>k. Rep_matrix (partRowReduce_matrix (move_minNonzero_nc A k) k k) i k = 0")
case True
have "¦Rep_matrix (iterate_partRowReduce A k) k k¦ = ¦Rep_matrix (iterate_partRowReduce (partRowReduce_matrix (move_minNonzero_nc A k) k k) k) k k¦"
using i by auto
also have "...=¦Rep_matrix (partRowReduce_matrix (move_minNonzero_nc A k) k k) k k¦" using True iterate_partRowReduce.simps by auto
also have "...=¦Rep_matrix (move_minNonzero_nc A k) k k¦" by (metis partRowReduce_eq_abs)
also have "...≤¦Rep_matrix A k k¦" by (metis assms move_minNonzero order_refl)
finally show ?thesis .
next
case False
have "¦Rep_matrix (iterate_partRowReduce A k) k k¦ < ¦Rep_matrix (move_minNonzero_nc A k) k k¦" using iterate_decrease_kk nz_Aik False by auto
also have "...≤¦Rep_matrix A k k¦" by (metis assms move_minNonzero order_refl)
finally show ?thesis by simp
qed
qed



lemma move_minNonzero_not_zero_Akk:
assumes "Rep_matrix A k k ≠ 0"
shows "Rep_matrix (move_minNonzero_nc A k) k k ≠ 0"

proof (rule move_minNonzero_not_zero, rule minNonzero_true)
have "k<nrows A" and "k<ncols A" using assms ncols_notzero[of A k k] nrows_notzero[of A k k] by auto
thus "¬ (∀i j. i ∈ {k..<nrows A} ∧ j ∈ {k..<ncols A} --> Rep_matrix A i j = 0)" using assms by auto
qed


lemma iterate_nonzero_kk:
assumes "Rep_matrix A k k ≠ 0"
shows "¦Rep_matrix (iterate_partRowReduce A k) k k¦ ≠ 0"

using assms
proof (induct "(A,k)" arbitrary: A k rule: measure_induct_rule [of f'])
fix A k assume h1: "!!Aa ka. [|f' (Aa, ka) < f' (A, k); Rep_matrix Aa ka ka ≠ 0|] ==> ¦Rep_matrix (iterate_partRowReduce Aa ka) ka ka¦ ≠ 0"
and h2: "Rep_matrix A k k ≠ 0"

show " ¦Rep_matrix (iterate_partRowReduce A k) k k¦ ≠ 0"
proof (cases "(∀i>k. Rep_matrix A i k = 0)")
case True thus ?thesis using iterate_partRowReduce.simps h2 by auto
next
case False note False1=False
have "¦Rep_matrix (iterate_partRowReduce (partRowReduce_matrix(move_minNonzero_nc A k) k k) k) k k¦ ≠ 0"
proof (cases "∀i>k. Rep_matrix (partRowReduce_matrix(move_minNonzero_nc A k) k k) i k = 0")
case True
have "¦Rep_matrix (iterate_partRowReduce (partRowReduce_matrix(move_minNonzero_nc A k) k k) k) k k¦
= ¦Rep_matrix (partRowReduce_matrix(move_minNonzero_nc A k) k k) k k¦"

using iterate_partRowReduce.simps[of "((partRowReduce_matrix (move_minNonzero_nc A k) k k))" k]
using True by auto
also have "... = ¦Rep_matrix (move_minNonzero_nc A k) k k¦" using partRowReduce_eq_abs by auto
also have "... ≠ 0" using move_minNonzero_not_zero_Akk[OF h2] by simp
finally show ?thesis .
next
case False
show ?thesis
proof (rule h1, unfold f'_def, auto)
have move_notzero: "Rep_matrix (move_minNonzero_nc A k) k k ≠ 0"
proof (rule move_minNonzero_not_zero, rule minNonzero_true)
from False1 obtain i where i: "i>k" and nz: "Rep_matrix A i k ≠ 0" by blast
have "k<ncols A" and "i<nrows A" using nz ncols_notzero[of A i k] nrows_notzero[of A i k] by auto
thus "¬ (∀i j. i ∈ {k..<nrows A} ∧ j ∈ {k..<ncols A} --> Rep_matrix A i j = 0)" using nz i by force
qed
thus "Rep_matrix (move_minNonzero_nc A k) k k = 0 ==> False" and " Rep_matrix (partRowReduce_matrix (move_minNonzero_nc A k) k k) k k = 0 ==> False"
using partRowReduce'[of k "(move_minNonzero_nc A k)"] by auto+
show " ¦Rep_matrix (move_minNonzero_nc (partRowReduce_matrix (move_minNonzero_nc A k) k k) k) k k¦ < ¦Rep_matrix (move_minNonzero_nc A k) k k¦"
proof -
from False obtain j where j: "j>k" and nz2: " Rep_matrix (partRowReduce_matrix (move_minNonzero_nc A k) k k) j k ≠ 0" by blast
have "¦Rep_matrix (move_minNonzero_nc (partRowReduce_matrix (move_minNonzero_nc A k) k k) k) k k¦
< ¦Rep_matrix (partRowReduce_matrix (move_minNonzero_nc A k) k k) k k¦"

proof (rule Akk_strict_decreasing_move_minNonzero)
show "k < j" using j .
show "Rep_matrix (partRowReduce_matrix (move_minNonzero_nc A k) k k) j k ≠ 0" using nz2 .
show "¦Rep_matrix (partRowReduce_matrix (move_minNonzero_nc A k) k k) j k¦ < ¦Rep_matrix (partRowReduce_matrix (move_minNonzero_nc A k) k k) k k¦"
by (metis `Rep_matrix (move_minNonzero_nc A k) k k = 0 ==> False` j partRowReduce_abs partRowReduce_eq_abs)
qed
also have "...=¦Rep_matrix (move_minNonzero_nc A k) k k¦" using partRowReduce'[of k "(move_minNonzero_nc A k)"] by auto
finally show ?thesis .
qed
qed
qed
thus ?thesis using False1 iterate_partRowReduce.simps[of A k] by auto
qed
qed

lemma partColumnReduce_eq_abs:
"¦Rep_matrix (partColumnReduce_matrix A k k) k k¦ = ¦Rep_matrix A k k¦"

by (metis eq_imp_le partColumnReduce')

lemma partColumnReduce_iterate_partRowReduce:
assumes "Rep_matrix A k k ≠ 0"
shows "∀i>k. Rep_matrix (partColumnReduce_matrix (iterate_partRowReduce A k) k k) i k = 0"

proof (rule allI, rule impI)
(*proof (rule+)*)
fix i
assume i: "i>k"
have "Rep_matrix (partColumnReduce_matrix (iterate_partRowReduce A k) k k) i k = Rep_matrix (iterate_partRowReduce A k) i k"
by (metis le_refl partColumnReduce')
also have "...= 0" using iterate_partRowReduce[OF assms] using i by auto
finally show "Rep_matrix (partColumnReduce_matrix (iterate_partRowReduce A k) k k) i k = 0" .
qed

partial_function (tailrec) Diagonalize :: "(int matrix) => nat => (int matrix)"
where "Diagonalize A k = (if ((∀m. m>k --> Rep_matrix A m k = 0) ∧ (∀n. n>k --> Rep_matrix A k n = 0)) then A else
Diagonalize (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) k)"


declare Diagonalize.simps [simp del]

lemma Diagonalize_iterate_zero:
assumes minNonzero: "fst_bis (minNonzero_nc A k) = True"
shows "(∀m. m>k --> Rep_matrix (Diagonalize A k) m k = 0)"

using assms
proof (induct "(A,k)" arbitrary: A k rule: measure_induct_rule [of f'])
fix A k
assume h1: "!!Aa ka. [|f' (Aa, ka) < f' (A, k); fst_bis (minNonzero_nc Aa ka) = True|] ==> ∀m>ka. Rep_matrix (Diagonalize Aa ka) m ka = 0"
and h2: "fst_bis (minNonzero_nc A k) = True"

show "∀m>k. Rep_matrix (Diagonalize A k) m k = 0"
proof (cases "(∀m>k. Rep_matrix A m k = 0) ∧ (∀n. n>k --> Rep_matrix A k n = 0)")
case True thus ?thesis using Diagonalize.simps by auto
next
case False
have "(∀m>k. Rep_matrix (Diagonalize ((partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k)) k) m k = 0)"
proof (cases "(∀n>k. Rep_matrix (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) k n = 0)")
case True
have "(∀n>k. Rep_matrix (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) n k = 0)"
by (metis h2 move_minNonzero_not_zero partColumnReduce_iterate_partRowReduce)
thus ?thesis using True Diagonalize.simps by auto
next
case False
from this obtain n where n: "n>k" and nz_n: "Rep_matrix (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) k n ≠ 0" by blast
show ?thesis
proof (rule h1, unfold f'_def, auto)
have "Rep_matrix (move_minNonzero_nc A k) k k ≠ 0"
by (metis h2 move_minNonzero_not_zero)
thus "Rep_matrix (move_minNonzero_nc A k) k k = 0 ==> False" by contradiction
show " fst_bis (minNonzero_nc (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) k)"
proof -
have n_l_cols: "n<ncols ((partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k))" using nz_n ncols_notzero by auto
have k_l_rows: "k<nrows ((partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k))" using nz_n nrows_notzero by auto
thus ?thesis using minNonzero_true nz_n n_l_cols n by force
qed
show " ¦Rep_matrix (move_minNonzero_nc (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) k) k k¦
< ¦Rep_matrix (move_minNonzero_nc A k) k k¦"

proof -
have " ¦Rep_matrix (move_minNonzero_nc (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) k) k k¦
< ¦Rep_matrix (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) k k¦"

proof (rule Akk_strict_decreasing_move_minNonzero2)
show "k ≤ k" by simp
show "k < n" using n .
show "¦Rep_matrix (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) k n¦
< ¦Rep_matrix (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) k k¦"

by (metis `Rep_matrix (move_minNonzero_nc A k) k k = 0 ==> False` abs_zero iterate_nonzero_kk n partColumnReduce_abs partColumnReduce_eq_abs)
show "Rep_matrix (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) k n ≠ 0" using nz_n .
qed
also have "...=¦Rep_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k¦" by (metis partColumnReduce_eq_abs)
also have "...≤ ¦Rep_matrix (move_minNonzero_nc A k) k k¦" by (metis `Rep_matrix (move_minNonzero_nc A k) k k = 0 ==> False` iterate_decrease_kk')
finally show ?thesis .
qed
qed
qed
thus ?thesis using Diagonalize.simps False by auto
qed
qed

lemma Diagonalize:
assumes minNonzero: "fst_bis (minNonzero_nc A k) = True"
shows "(∀m. m>k --> Rep_matrix (Diagonalize A k) m k = 0) ∧ (∀n. n>k --> Rep_matrix (Diagonalize A k) k n = 0)"

using minNonzero
proof (induct "(A,k)" arbitrary: A k rule: measure_induct_rule [of f'])
fix A k
assume h1: "!!Aa ka. [|f' (Aa, ka) < f' (A, k); fst_bis (minNonzero_nc Aa ka) = True|]
==> (∀m>ka. Rep_matrix (Diagonalize Aa ka) m ka = 0) ∧ (∀n>ka. Rep_matrix (Diagonalize Aa ka) ka n = 0)"

and h2: "fst_bis (minNonzero_nc A k) = True"

show "(∀m>k. Rep_matrix (Diagonalize A k) m k = 0) ∧ (∀n>k. Rep_matrix (Diagonalize A k) k n = 0)"
proof (cases "(∀m>k. Rep_matrix A m k = 0) ∧ (∀n>k. Rep_matrix A k n = 0)")
case True thus ?thesis using Diagonalize.simps by auto
next
case False note False_1=False
have "(∀m>k. Rep_matrix (Diagonalize ((partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k)) k) m k = 0)
∧ (∀n>k. Rep_matrix (Diagonalize (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) k) k n = 0)"

proof (cases "(∀m>k. Rep_matrix (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) m k= 0)
∧ (∀n>k. Rep_matrix (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) k n = 0)"
)

case True thus ?thesis using Diagonalize.simps by auto
next
case False
have not_all_row_zero: "¬ (∀n>k. Rep_matrix (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) k n = 0)"
proof (rule ccontr)
assume "¬ ¬ (∀n>k. Rep_matrix (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) k n = 0)"
hence "¬ (∀m>k. Rep_matrix (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) m k= 0)" using False by auto
from this obtain i where i:"i>k" and nz_i: "Rep_matrix (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) i k ≠ 0"
by blast
have "Rep_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) i k
= Rep_matrix (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) i k"

by (metis order_eq_refl partColumnReduce')
also have "...≠0" using nz_i by simp
finally show False using iterate_partRowReduce[of "(move_minNonzero_nc A k)" k] i by (metis h2 move_minNonzero_not_zero)
qed
from this obtain i where i: "i>k" and nz_i: "Rep_matrix (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) k i ≠ 0"
by blast
show ?thesis
proof (rule h1, unfold f'_def, auto)
have nz_movemin: "Rep_matrix (move_minNonzero_nc A k) k k ≠ 0" by (metis h2 move_minNonzero_not_zero)
thus "Rep_matrix (move_minNonzero_nc A k) k k = 0 ==> False" by contradiction
show "fst_bis (minNonzero_nc (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) k)"
proof -
have "fst_bis (minNonzero_nc (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) k) = True"
proof (rule minNonzero_true)
have k1:"k<nrows (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k)" using nz_i nrows_notzero by auto
have i1: "i<ncols (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k)" using nz_i ncols_notzero by auto
thus "¬ (∀i j. i ∈ {k..<nrows (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k)} ∧
j ∈ {k..<ncols (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k)} -->
Rep_matrix (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) i j = 0)"
using i k1 nz_i by force
qed
thus ?thesis
by simp
qed
show " ¦Rep_matrix (move_minNonzero_nc (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) k) k k¦
< ¦Rep_matrix (move_minNonzero_nc A k) k k¦"

proof -
have "¦Rep_matrix (move_minNonzero_nc (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) k) k k¦
< ¦Rep_matrix (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) k k¦"

proof (rule Akk_strict_decreasing_move_minNonzero2)
show "k < i" using i .
show " k ≤ k" by simp
show "¦Rep_matrix (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) k i¦
< ¦Rep_matrix (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) k k¦"

by (metis `Rep_matrix (move_minNonzero_nc A k) k k = 0 ==> False` abs_zero i iterate_nonzero_kk partColumnReduce_abs partColumnReduce_eq_abs)
show "Rep_matrix (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) k i ≠ 0" using nz_i .
qed
also have "...= ¦Rep_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k¦" by (metis partColumnReduce_eq_abs)
also have "...≤ ¦Rep_matrix (move_minNonzero_nc A k) k k¦"
by (metis `Rep_matrix (move_minNonzero_nc A k) k k = 0 ==> False` iterate_decrease_kk')
finally show ?thesis .
qed
qed
qed
thus ?thesis using False Diagonalize.simps by auto
qed
qed


lemma Diagonalize_result:
shows "(∀m. m>k --> Rep_matrix (Diagonalize A k) m k = 0) ∧ (∀n. n>k --> Rep_matrix (Diagonalize A k) k n = 0)"

proof (cases "fst_bis (minNonzero_nc A k) = True")
case True show ?thesis using Diagonalize[OF True] .
next
case False
have 1: "∀(i::nat) j::nat. i ∈ {k..<nrows A} ∧ j ∈ {k..<ncols A} --> Rep_matrix A i j = 0"
by (metis False minNonzero_true)
have 2: "∀i j. i ≥ (nrows A) ∨ j ≥ (ncols A) --> Rep_matrix A i j = 0"
by (metis ncols nrows)
have "∀i j. i ≥ k ∧ j ≥k --> Rep_matrix A i j = 0"
proof (rule, rule, rule)
fix i j assume i:"k ≤ i ∧ k ≤ j"
show "Rep_matrix A i j = (0::int)"
proof (cases "i<nrows A")
case False thus ?thesis using 2 by auto
next
case True note i_l_nrows=True show ?thesis
proof (cases "j<ncols A")
case False thus ?thesis using 2 by auto
next
case True show ?thesis using 1 True i_l_nrows i by auto
qed
qed
qed
thus ?thesis using Diagonalize.simps by auto
qed

text{*Next function returns True if a matrix A is diagonal until position k (not included)*}
definition is_Diagonal :: "(int matrix => nat => bool)"
where "is_Diagonal A k = (∀ a b. ((0 ≤ a ∧ a<k) ∨ (0 ≤ b ∧ b<k)) ∧ a≠b --> Rep_matrix A a b = 0)"


lemma is_Diagonal_0:
shows "is_Diagonal A 0 = True"

unfolding is_Diagonal_def by simp

lemma Diagonalize_interchange_rows:
assumes es_H:"is_Diagonal A k"
and i:"i≥k"
shows "is_Diagonal (interchange_rows_matrix A i k) k"

proof (unfold interchange_rows_matrix_def, unfold interchange_rows_infmatrix_def,unfold is_Diagonal_def, auto)
have r: "Rep_matrix (Abs_matrix (λz j. if z = i then Rep_matrix A k j else if z = k then Rep_matrix A i j else Rep_matrix A z j))
= (λz j. if z = i then Rep_matrix A k j else if z = k then Rep_matrix A i j else Rep_matrix A z j)"
using Rep_Abs_aux_interchange2'[of i A k] by simp
fix a b
assume a_noteq_b: "a ≠ b" and a_le_k: "a < k"
have a_not_k: "a≠k" using i a_le_k by simp
have a_not_i: "a≠i" using i a_le_k by simp
have "(λz j. if z = i then Rep_matrix A k j else if z = k then Rep_matrix A i j else Rep_matrix A z j) a b = Rep_matrix A a b" using a_not_k a_not_i by auto
also have "...=0" using es_H unfolding is_Diagonal_def using a_le_k a_noteq_b by simp
finally show "Rep_matrix (Abs_matrix (λia j. if ia = i then Rep_matrix A k j else if ia = k then Rep_matrix A i j else Rep_matrix A ia j)) a b = 0" using r by auto
next
fix a b
assume a_noteq_b: "a ≠ b" and b_le_k: "b < k"
show "Rep_matrix (Abs_matrix (λia j. if ia = i then Rep_matrix A k j else if ia = k then Rep_matrix A i j else Rep_matrix A ia j)) a b = 0"
proof (cases "a=i")
case True
hence "(λia j. if ia = i then Rep_matrix A k j else if ia = k then Rep_matrix A i j else Rep_matrix A ia j) a b = Rep_matrix A k b" by auto
also have "...=0" using es_H unfolding is_Diagonal_def using a_noteq_b and b_le_k by auto
finally show ?thesis using Rep_Abs_aux_interchange2'[of i A k] by auto
next
case False note a_not_i=False
show ?thesis
proof (cases "a=k")
case True
have i_not_b: "i≠b" by (metis b_le_k i linorder_not_less)
have "(λia j. if ia = i then Rep_matrix A k j else if ia = k then Rep_matrix A i j else Rep_matrix A ia j) a b = Rep_matrix A i b" using True by simp
also have "...=0"using es_H unfolding is_Diagonal_def using b_le_k i_not_b by auto
finally show ?thesis using Rep_Abs_aux_interchange2'[of i A k] by auto
next
case False
have "(λia j. if ia = i then Rep_matrix A k j else if ia = k then Rep_matrix A i j else Rep_matrix A ia j) a b = Rep_matrix A a b" using a_not_i False by simp
also have "...=0" using a_noteq_b and b_le_k es_H unfolding is_Diagonal_def by auto
finally show ?thesis using Rep_Abs_aux_interchange2'[of i A k] by auto
qed
qed
qed


lemma Diagonalize_interchange_columns:
assumes es_H:"is_Diagonal A k"
and i:"i≥k"
shows "is_Diagonal (interchange_columns_matrix A i k) k"

proof (unfold interchange_columns_matrix_def, unfold interchange_columns_infmatrix_def,unfold is_Diagonal_def, auto)
fix a b
assume a_noteq_b: "a ≠ b" and b_le_k: "b < k"
have b_not_k: "b≠k" using i b_le_k by simp
have b_not_i: "b≠i" using i b_le_k by simp
have "(λia j. if j = i then Rep_matrix A ia k else if j = k then Rep_matrix A ia i else Rep_matrix A ia j) a b = Rep_matrix A a b" using b_not_k b_not_i by simp
also have "...= 0" using es_H unfolding is_Diagonal_def using b_le_k a_noteq_b by simp
finally show "Rep_matrix (Abs_matrix (λia j. if j = i then Rep_matrix A ia k else if j = k then Rep_matrix A ia i else Rep_matrix A ia j)) a b = 0"
using Rep_Abs_aux_interchange'[of i A k] by auto
next
fix a b
assume a_noteq_b: "a ≠ b" and a_le_k: "a < k"
show "Rep_matrix (Abs_matrix (λia j. if j = i then Rep_matrix A ia k else if j = k then Rep_matrix A ia i else Rep_matrix A ia j)) a b = 0"
proof (cases "b=i")
case True
hence "(λia j. if j = i then Rep_matrix A ia k else if j = k then Rep_matrix A ia i else Rep_matrix A ia j) a b = Rep_matrix A a k" by simp
also have "...=0" using es_H unfolding is_Diagonal_def using a_le_k a_noteq_b by simp
finally show ?thesis using Rep_Abs_aux_interchange'[of i A k] by auto
next
case False note b_not_i=False
show ?thesis
proof (cases "b=k")
case True
have i_not_a: "a≠i" by (metis a_le_k i linorder_not_less)
have "(λia j. if j = i then Rep_matrix A ia k else if j = k then Rep_matrix A ia i else Rep_matrix A ia j) a b = Rep_matrix A a i" using True by simp
also have "...=0" using es_H unfolding is_Diagonal_def using a_le_k i_not_a by simp
finally show ?thesis using Rep_Abs_aux_interchange'[of i A k] by auto
next
case False
hence "(λia j. if j = i then Rep_matrix A ia k else if j = k then Rep_matrix A ia i else Rep_matrix A ia j) a b = Rep_matrix A a b" using b_not_i by simp
also have "...=0" using es_H unfolding is_Diagonal_def using a_le_k a_noteq_b by simp
finally show ?thesis using Rep_Abs_aux_interchange'[of i A k] by auto
qed
qed
qed


lemma k_le_snd_minNonzero:
assumes ex_minNonzero: "fst_bis (minNonzero_nc A k) = True"
shows "k ≤ snd_bis (minNonzero_nc A k)"

proof -
have notall: "¬ (∀(i::nat) j::nat. i ∈ {k..<nrows A} ∧ j ∈ {k..<ncols A} --> Rep_matrix A i j = 0)"
using minNonzero_true_inv[OF ex_minNonzero] .
def P=="(λpos. k ≤ fst pos ∧ k ≤ snd pos ∧ Rep_matrix A (fst pos) (snd pos) ≠ 0 ∧
(∀m n. k ≤ m ∧ k ≤ n ∧ Rep_matrix A m n ≠ 0 --> ¦Rep_matrix A (fst pos) (snd pos)¦ ≤ ¦Rep_matrix A m n¦))"

def pos=="LEAST k. P k"
have "P (pos)" using minNonzero_least[of P k A] ex_minNonzero unfolding P_def pos_def by auto
hence "k≤fst(pos)"unfolding pos_def P_def by auto
thus ?thesis unfolding minNonzero_nc_def using notall unfolding pos_def P_def snd_bis_def by auto
qed

lemma k_le_trd_minNonzero:
assumes ex_minNonzero: "fst_bis (minNonzero_nc A k) = True"
shows "k ≤ trd_bis (minNonzero_nc A k)"

proof -
have notall: "¬ (∀(i::nat) j::nat. i ∈ {k..<nrows A} ∧ j ∈ {k..<ncols A} --> Rep_matrix A i j = 0)"
using minNonzero_true_inv[OF ex_minNonzero] .
def P=="(λpos. k ≤ fst pos ∧ k ≤ snd pos ∧ Rep_matrix A (fst pos) (snd pos) ≠ 0 ∧
(∀m n. k ≤ m ∧ k ≤ n ∧ Rep_matrix A m n ≠ 0 --> ¦Rep_matrix A (fst pos) (snd pos)¦ ≤ ¦Rep_matrix A m n¦))"

def pos=="LEAST k. P k"
have "P (pos)" using minNonzero_least[of P k A] ex_minNonzero unfolding P_def pos_def by auto
hence "k≤snd(pos)"unfolding pos_def P_def by auto
thus ?thesis unfolding minNonzero_nc_def using notall unfolding pos_def P_def trd_bis_def by auto
qed


lemma Diagonalize_move_element:
assumes "is_Diagonal A k"
and "i≥k"
and "j≥k"
shows "is_Diagonal (move_element A i j k k) k"

by (metis Diagonalize_interchange_columns Diagonalize_interchange_rows assms(1) assms(2) assms(3) move_element_def)

lemma Diagonalize_move_minNonzero:
assumes es_H: "is_Diagonal A k"
and f: "fst_bis (minNonzero_nc A k) = True"
shows "is_Diagonal (move_minNonzero_nc A k) k"

proof (unfold move_minNonzero_nc_def, rule Diagonalize_move_element)
show "is_Diagonal A k" using es_H .
show "k ≤ snd_bis (minNonzero_nc A k)" using k_le_snd_minNonzero[OF f] .
show "k ≤ trd_bis (minNonzero_nc A k)" using k_le_trd_minNonzero[OF f] .
qed


lemma is_Diagonal_partRowReduce_move_minNonzero:
assumes es_H:"is_Diagonal A k"
and f: "fst_bis (minNonzero_nc A k) = True"
shows "is_Diagonal (partRowReduce_matrix (move_minNonzero_nc A k) k k) k"

proof (unfold is_Diagonal_def, auto)
fix a b
assume a_not_b: "a≠b" and a_le_k: "a<k"
have "Rep_matrix (partRowReduce_matrix (move_minNonzero_nc A k) k k) a b = Rep_matrix (move_minNonzero_nc A k) a b" by (metis a_le_k nat_less_le partRowReduce')
also have "...=0" using Diagonalize_move_minNonzero[OF es_H f] unfolding is_Diagonal_def using a_not_b a_le_k by simp
finally show "Rep_matrix (partRowReduce_matrix (move_minNonzero_nc A k) k k) a b = 0" .
next
fix a b
assume a_not_b: "a≠b" and b_le_k: "b<k"
have "Rep_matrix (partRowReduce_matrix (move_minNonzero_nc A k) k k) a b = Rep_matrix (move_minNonzero_nc A k) a b"
proof (cases "a≤k")
case True
thus "Rep_matrix (partRowReduce_matrix (move_minNonzero_nc A k) k k) a b = Rep_matrix (move_minNonzero_nc A k) a b" by (metis nat_less_le partRowReduce')
next
case False
hence a_noteq_k:"a≠k" by auto
have eq_0: "Rep_matrix (move_minNonzero_nc A k) a b = 0" using Diagonalize_move_minNonzero[OF es_H f] unfolding is_Diagonal_def using a_not_b b_le_k by auto
have eq_0': "Rep_matrix (move_minNonzero_nc A k) k b = 0" using Diagonalize_move_minNonzero[OF es_H f] unfolding is_Diagonal_def using b_le_k by auto
have r:"Rep_matrix
(Abs_matrix
(λi j. if k < i
then row_add_infmatrix (Rep_matrix (move_minNonzero_nc A k)) i k (- (Rep_matrix (move_minNonzero_nc A k) i k div Rep_matrix (move_minNonzero_nc A k) k k)) i j
else Rep_matrix (move_minNonzero_nc A k) i j)) =(λi j. if k < i
then row_add_infmatrix (Rep_matrix (move_minNonzero_nc A k)) i k (- (Rep_matrix (move_minNonzero_nc A k) i k div Rep_matrix (move_minNonzero_nc A k) k k)) i j
else Rep_matrix (move_minNonzero_nc A k) i j)"
using aux_partRowReduce[of k "(move_minNonzero_nc A k)" k] by auto
hence "Rep_matrix (partRowReduce_matrix (move_minNonzero_nc A k) k k) a b
= row_add_infmatrix (Rep_matrix (move_minNonzero_nc A k)) a k (- (Rep_matrix (move_minNonzero_nc A k) a k div Rep_matrix (move_minNonzero_nc A k) k k)) a b"

unfolding partRowReduce_matrix_def partRowReduce_infmatrix_def using False by auto
also have "...= 0" using eq_0 eq_0' unfolding row_add_infmatrix_def by auto
finally show ?thesis using eq_0 by auto
qed
also have "...=0" using Diagonalize_move_minNonzero[OF es_H f] unfolding is_Diagonal_def using a_not_b b_le_k by auto
finally show "Rep_matrix (partRowReduce_matrix (move_minNonzero_nc A k) k k) a b = 0" .
qed

lemma is_Diagonal_iterate_partRowReduce:
assumes "is_Diagonal A k"
shows "is_Diagonal (iterate_partRowReduce A k) k"

using assms
proof (induct "(A,k)" arbitrary: A k rule: measure_induct_rule [of f'])
fix A k
assume h1:"!!Aa ka. [|f' (Aa, ka) < f' (A, k); is_Diagonal Aa ka|] ==> is_Diagonal (iterate_partRowReduce Aa ka) ka" and h2: "is_Diagonal A k"
show "is_Diagonal (iterate_partRowReduce A k) k"
proof (cases "∀i>k. Rep_matrix A i k = 0")
case True thus ?thesis using iterate_partRowReduce.simps h2 by auto
next
case False
from this obtain i where i:"i>k" and not_zero: "Rep_matrix A i k ≠ 0" by blast
have i_l_rows: "i<nrows A"
by (metis not_zero nrows_notzero)
have k_l_cols: "k<ncols A" by (metis not_zero ncols_notzero)
have "is_Diagonal (iterate_partRowReduce (partRowReduce_matrix(move_minNonzero_nc A k) k k) k) k"
proof (cases "∀i>k. Rep_matrix (partRowReduce_matrix(move_minNonzero_nc A k) k k) i k = 0")
case True show ?thesis unfolding iterate_partRowReduce.simps[of "(partRowReduce_matrix (move_minNonzero_nc A k) k k)" k]
using True apply simp using is_Diagonal_partRowReduce_move_minNonzero[OF h2 _] minNonzero_true[of k A] i_l_rows k_l_cols i not_zero by force
next
case False
from this obtain a where a:"a>k" and nz: "Rep_matrix (partRowReduce_matrix(move_minNonzero_nc A k) k k) a k ≠ 0" by blast
show ?thesis
proof (rule h1, unfold f'_def, auto)
show "is_Diagonal (partRowReduce_matrix (move_minNonzero_nc A k) k k) k"
using is_Diagonal_partRowReduce_move_minNonzero[OF h2 _] minNonzero_true[of k A] i_l_rows k_l_cols i not_zero by force
show move_notzero: "Rep_matrix (move_minNonzero_nc A k) k k = 0 ==> False"
using move_minNonzero_not_zero[of A k] minNonzero_true[of k A] i_l_rows k_l_cols i not_zero by force
show "¦Rep_matrix (move_minNonzero_nc (partRowReduce_matrix (move_minNonzero_nc A k) k k) k) k k¦ < ¦Rep_matrix (move_minNonzero_nc A k) k k¦"
proof -
have "¦Rep_matrix (move_minNonzero_nc (partRowReduce_matrix (move_minNonzero_nc A k) k k) k) k k¦
< ¦Rep_matrix (partRowReduce_matrix (move_minNonzero_nc A k) k k) k k¦"

proof (rule Akk_strict_decreasing_move_minNonzero)
show "a>k" using a .
show " ¦Rep_matrix (partRowReduce_matrix (move_minNonzero_nc A k) k k) a k¦ < ¦Rep_matrix (partRowReduce_matrix (move_minNonzero_nc A k) k k) k k¦"
by (metis `Rep_matrix (move_minNonzero_nc A k) k k = 0 ==> False` a partRowReduce_abs partRowReduce_eq_abs)
show "Rep_matrix (partRowReduce_matrix (move_minNonzero_nc A k) k k) a k ≠ 0" using nz .
qed
also have "...= ¦Rep_matrix (move_minNonzero_nc A k) k k¦" using partRowReduce_eq_abs by auto
finally show ?thesis .
qed
qed
qed
thus ?thesis using False iterate_partRowReduce.simps[of A k] by auto
qed
qed


lemma is_Diagonal_partColumnReduce:
assumes es_H: "is_Diagonal A k"
shows "is_Diagonal (partColumnReduce_matrix A k k) k"

proof (unfold is_Diagonal_def, unfold partColumnReduce_matrix_def, unfold partColumnReduce_infmatrix_def, auto)
fix a b
assume a_not_b: "a ≠ b" and b: "b < k"
show "Rep_matrix (Abs_matrix (λi j. if k < j then column_add_infmatrix (Rep_matrix A) j k (- (Rep_matrix A k j div Rep_matrix A k k)) i j else Rep_matrix A i j)) a b = 0"
using es_H unfolding is_Diagonal_def using b a_not_b using aux_partColumnReduce[of k A k] by auto
next
fix a b
assume a_not_b: "a ≠ b" and a: "a < k"
show "Rep_matrix (Abs_matrix (λi j. if k < j then column_add_infmatrix (Rep_matrix A) j k (- (Rep_matrix A k j div Rep_matrix A k k)) i j else Rep_matrix A i j)) a b =0"
proof (cases "b≤k")
case True show ?thesis using es_H unfolding is_Diagonal_def using True a_not_b using aux_partColumnReduce[of k A k] a by auto
next
case False
have eq_0: "Rep_matrix A a b = 0"
by (metis a a_not_b assms is_Diagonal_def less_eq_nat.simps(1) of_nat_less_iff)
have eq_0': "Rep_matrix A a k = 0"
by (metis a assms is_Diagonal_def less_eq_nat.simps(1) of_nat_less_iff order_less_imp_not_eq2)
have "(λi j. if k < j then column_add_infmatrix (Rep_matrix A) j k (- (Rep_matrix A k j div Rep_matrix A k k)) i j else Rep_matrix A i j) a b
= column_add_infmatrix (Rep_matrix A) b k (- (Rep_matrix A k b div Rep_matrix A k k)) a b"
using False by auto
also have "...=Rep_matrix A a b + - (Rep_matrix A k b div Rep_matrix A k k) * Rep_matrix A a k" unfolding column_add_infmatrix_def by simp
also have "...=0" using eq_0 eq_0' by auto
finally show ?thesis using aux_partColumnReduce[of k A k] by auto
qed
qed


lemma is_Diagonal_partColumnReduce_matrix_iterate_partRow:
assumes "is_Diagonal A k"
and "fst_bis (minNonzero_nc A k) = True"
shows "is_Diagonal (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) k"

by (metis Diagonalize_move_minNonzero assms(1) assms(2) is_Diagonal_iterate_partRowReduce is_Diagonal_partColumnReduce)

lemma is_Diagonal_Diagonalize_k:
assumes es_H: "is_Diagonal A k"
shows "is_Diagonal (Diagonalize A k) k"

proof (cases "fst_bis (minNonzero_nc A k) = True")
case False
hence 1: "(∀i j. i ∈ {k..<nrows A} ∧ j ∈ {k..<ncols A} --> Rep_matrix A i j = 0)" by (metis minNonzero_true)
have 2: "∀i j. i ≥ (nrows A) ∨ j ≥ (ncols A) --> Rep_matrix A i j = 0"
by (metis ncols nrows)
have 3: "∀i j. i ≥ k ∧ j ≥k --> Rep_matrix A i j = 0"
proof (rule, rule, rule)
fix i j assume i:"k ≤ i ∧ k ≤ j"
show "Rep_matrix A i j = (0::int)"
proof (cases "i<nrows A")
case False thus ?thesis using 2 by auto
next
case True note i_l_nrows=True show ?thesis
proof (cases "j<ncols A")
case False thus ?thesis using 2 by auto
next
case True show ?thesis using 1 True i_l_nrows i by auto
qed
qed
qed
hence "(∀m>k. Rep_matrix A m k = 0) ∧ (∀n>k. Rep_matrix A k n = 0)" by simp
thus ?thesis using Diagonalize.simps[of A k] using es_H unfolding is_Diagonal_def by simp
next
case True
thus ?thesis using es_H
proof (induct "(A,k)" arbitrary: A k rule: measure_induct_rule [of f'])
fix A k
assume h1: "!!Aa ka. [|f' (Aa, ka) < f' (A, k); fst_bis (minNonzero_nc Aa ka) = True; is_Diagonal Aa ka|] ==> is_Diagonal (Diagonalize Aa ka) ka"
and h2: "fst_bis (minNonzero_nc A k) = True" and h3: "is_Diagonal A k"

show "is_Diagonal (Diagonalize A k) k"
proof (cases "(∀m>k. Rep_matrix A m k = 0) ∧ (∀n>k. Rep_matrix A k n = 0)")
case True thus ?thesis using Diagonalize.simps using h2 h3 unfolding is_Diagonal_def by auto
next
case False
have "is_Diagonal (Diagonalize (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) k) k"
proof (cases "(∀m>k. Rep_matrix (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) m k = 0)
∧ (∀n>k. Rep_matrix (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) k n = 0)"
)

case True thus ?thesis using Diagonalize.simps[of "partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k" k]
using is_Diagonal_partColumnReduce_matrix_iterate_partRow[OF h3 h2] by auto
next
case False
show ?thesis
proof (rule h1, unfold f'_def, auto)
show move_not_zero: "Rep_matrix (move_minNonzero_nc A k) k k = 0 ==> False"
by (metis h2 move_minNonzero_not_zero)
show "is_Diagonal (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) k"
using is_Diagonal_partColumnReduce_matrix_iterate_partRow[OF h3 h2] .
have "(∀m>k. Rep_matrix (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) m k = 0)"
by (metis move_not_zero partColumnReduce_iterate_partRowReduce)
from this obtain n where n: "n>k" and nz_n: "Rep_matrix (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) k n ≠ 0"
using False by blast
show " fst_bis (minNonzero_nc (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) k)"
proof -
have n_l_cols: "n<ncols ((partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k))" using nz_n ncols_notzero by auto
have k_l_rows: "k<nrows ((partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k))" using nz_n nrows_notzero by auto
thus ?thesis using minNonzero_true nz_n n_l_cols n by force
qed
show " ¦Rep_matrix (move_minNonzero_nc (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) k) k k¦
< ¦Rep_matrix (move_minNonzero_nc A k) k k¦"

proof -
have " ¦Rep_matrix (move_minNonzero_nc (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) k) k k¦
< ¦Rep_matrix (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) k k¦"

proof (rule Akk_strict_decreasing_move_minNonzero2)
show "k ≤ k" by simp
show "k < n" using n .
show "¦Rep_matrix (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) k n¦
< ¦Rep_matrix (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) k k¦"

by (metis `Rep_matrix (move_minNonzero_nc A k) k k = 0 ==> False` abs_zero iterate_nonzero_kk n partColumnReduce_abs partColumnReduce_eq_abs)
show "Rep_matrix (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) k n ≠ 0" using nz_n .
qed
also have "...=¦Rep_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k¦" by (metis partColumnReduce_eq_abs)
also have "...≤ ¦Rep_matrix (move_minNonzero_nc A k) k k¦" by (metis `Rep_matrix (move_minNonzero_nc A k) k k = 0 ==> False` iterate_decrease_kk')
finally show ?thesis .
qed
qed
qed
thus ?thesis using Diagonalize.simps False by auto
qed
qed
qed


lemma is_Diagonal_Diagonalize_suc_k:
assumes es_H: "is_Diagonal A k"
shows "is_Diagonal (Diagonalize A k) (k+1)"

proof -
have 1:"is_Diagonal (Diagonalize A k) k" using is_Diagonal_Diagonalize_k[OF es_H] .
have 2: "(∀m>k. Rep_matrix (Diagonalize A k) m k = 0) ∧ (∀n>k. Rep_matrix (Diagonalize A k) k n = 0)" using Diagonalize_result .
show ?thesis using 1 2 unfolding is_Diagonal_def by smt
qed


lemma not_is_Diagonal_greather:
assumes "¬ is_Diagonal A k"
and "k≤j"
shows "¬ is_Diagonal A j"

using assms unfolding is_Diagonal_def by force

lemma is_Diagonal_less:
assumes "is_Diagonal A k"
and "j≤k"
shows "is_Diagonal A j"

using assms unfolding is_Diagonal_def by force

text{*This function diagonalizes a matrix up to position k (included).*}
fun Diagonalize_up_to_k :: "int matrix => nat => int matrix"
where
"Diagonalize_up_to_k A 0 = Diagonalize A 0"
| "Diagonalize_up_to_k A (Suc k) = Diagonalize (Diagonalize_up_to_k A k) (Suc k)"


text{*This functions returns us True if a matrix is diagonal up to position k (included). False, in other case.*}
definition Diagonalize_p :: "(int matrix => nat => bool)"
where "Diagonalize_p A k = is_Diagonal A (k + 1)"



lemma Diagonalize_p_zero:
shows "Diagonalize_p (Diagonalize A 0) 0"

by (metis Diagonalize_p_def is_Diagonal_0 is_Diagonal_Diagonalize_suc_k of_nat_0)


lemma Diagonalize_p_incr:
assumes "Diagonalize_p A k"
shows "Diagonalize_p (Diagonalize A (Suc k)) (Suc k)"

by (metis Diagonalize_p_def Suc_eq_plus1 assms is_Diagonal_Diagonalize_suc_k)


lemma
Diagonalize_p_mon:
assumes "Diagonalize_p A n"
and "m ≤ n"
shows "Diagonalize_p A m"

using is_Diagonal_less[of A "n+1" "m+1"] using assms unfolding Diagonalize_p_def by auto

lemma Diagonalize_p_Diagonalize_less:
assumes m_le_n: "m ≤ n"
shows "Diagonalize_p (Diagonalize_up_to_k A n) m"

using m_le_n
proof (induct n arbitrary: m)
case 0
show ?case
using "0.prems" using Diagonalize_p_zero [of A] by simp
next
case (Suc n)
fix n m
assume hyp: "!!m. m ≤ n ==> Diagonalize_p (Diagonalize_up_to_k A n) m"
and prem: "m ≤ Suc n"

have hn: "Diagonalize_p (Diagonalize_up_to_k A n) n" using hyp by simp
show "Diagonalize_p (Diagonalize_up_to_k A (Suc n)) m"
proof (cases "m = Suc n")
case False
show ?thesis
by (metis Diagonalize_up_to_k.simps(2) Diagonalize_p_incr Diagonalize_p_mon hn linorder_not_le of_nat_less_imp_less prem)(*
apply (rule Diagonalize_p_mon [OF _ prem])
apply (subst Diagonalize_up_to_k.simps)
apply (rule Diagonalize_p_incr)
by (rule hn)*)

next
case True
show "Diagonalize_p (Diagonalize_up_to_k A (Suc n)) m"
unfolding True
apply (subst Diagonalize_up_to_k.simps)
apply (rule Diagonalize_p_incr)
apply (rule hyp) by simp
qed
qed


corollary Diagonalize_p_Diagonalize_n:
shows "Diagonalize_p (Diagonalize_up_to_k A n) n"

using Diagonalize_p_Diagonalize_less by simp

corollary Diagonalize_p_Diagonalize_max:
shows "Diagonalize_p (Diagonalize_up_to_k A (max (nrows A) (ncols A))) (max (nrows A) (ncols A))"

using Diagonalize_p_Diagonalize_n by auto

definition P_interchange_rows :: "int matrix => nat => nat => int matrix"
where "P_interchange_rows A n m == interchange_rows_matrix (one_matrix (nrows A)) n m"


definition P_interchange_rows' :: "int matrix => nat => nat => int matrix"
where "P_interchange_rows' A n m ≡ Abs_matrix(λi j. if (i=n ∧ j=m) then 1 else (if (i=m ∧ j=n) then 1 else (if (i=j ∧ i≠m ∧ i≠n ∧ i < nrows A) then 1 else 0)))"


definition Q_interchange_columns :: "int matrix => nat => nat => int matrix"
where "Q_interchange_columns A n m == interchange_columns_matrix (one_matrix (ncols A)) n m"


definition Q_interchange_columns' :: "int matrix => nat => nat => int matrix"
where "Q_interchange_columns' A n m ≡ Abs_matrix(λi j. if (i=n ∧ j=m) then 1 else (if (i=m ∧ j=n) then 1 else (if (i=j ∧ i≠m ∧ i≠n ∧ j < ncols A) then 1 else 0)))"


definition P_row_add :: "int matrix => nat => nat => int => int matrix"
where "P_row_add A i j q = row_add_matrix (one_matrix (nrows A)) i j q"


definition P_row_add' :: "int matrix => nat => nat => int => int matrix"
where "P_row_add' A n m q = Abs_matrix(λi j. if (i=n ∧ j=m) then q else (if (i=j ∧ i<nrows A) then 1 else 0))"


definition Q_column_add :: "int matrix => nat => nat => int => int matrix"
where "Q_column_add A i j q = column_add_matrix (one_matrix (ncols A)) i j q"


definition Q_column_add' :: "int matrix => nat => nat => int => int matrix"
where "Q_column_add' A n m q = Abs_matrix(λi j. if (i=m ∧ j=n) then q else (if (i=j ∧ j<ncols A) then 1 else 0))"


lemma P_row_add_eq:
assumes "n<nrows A"
and "m<nrows A"
and "n≠m"
shows "P_row_add' (A::int matrix) n m q = P_row_add A n m q"

proof (unfold P_row_add'_def P_row_add_def row_add_matrix_def row_add_infmatrix_def one_matrix_def)
have aux1: "Rep_matrix (Abs_matrix (λj i. if j = i ∧ j < nrows A then 1 else 0)) =(λj i. if j = i ∧ j < nrows A then 1 else 0::int)"
by (metis (no_types) Rep_one_matrix one_matrix_def)
have "(λi j. if i = n ∧ j = m then q else if i = j ∧ i < nrows A then 1::int else (0::int))=(λa b. if a = n
then (λj i. if j = i ∧ j < nrows A then 1::int else (0::int)) n b +
q * (λj i. if j = i ∧ j < nrows A then 1::int else (0::int)) m b
else (λj i. if j = i ∧ j < nrows A then 1::int else (0::int)) a b)"

apply (rule) using assms by auto
hence " (λi j. if i = n ∧ j = m then q else if i = j ∧ i < nrows A then 1 else 0) =
(λa b. if a = n
then Rep_matrix (Abs_matrix (λj i. if j = i ∧ j < nrows A then 1 else 0)) n b + q * Rep_matrix (Abs_matrix (λj i. if j = i ∧ j < nrows A then 1 else 0)) m b
else Rep_matrix (Abs_matrix (λj i. if j = i ∧ j < nrows A then 1 else 0)) a b)"

using aux1 by force
thus "Abs_matrix (λi j. if i = n ∧ j = m then q else if i = j ∧ i < nrows A then 1 else 0) =
Abs_matrix
(λa b. if a = n
then Rep_matrix (Abs_matrix (λj i. if j = i ∧ j < nrows A then 1 else 0)) n b + q * Rep_matrix (Abs_matrix (λj i. if j = i ∧ j < nrows A then 1 else 0)) m b
else Rep_matrix (Abs_matrix (λj i. if j = i ∧ j < nrows A then 1 else 0)) a b)"
by simp
qed


lemma P_interchange_row_eq:
assumes "n<nrows A"
and "m<nrows A"
shows "P_interchange_rows' (A::int matrix) n m= P_interchange_rows A n m"

proof (unfold P_interchange_rows'_def P_interchange_rows_def interchange_rows_matrix_def interchange_rows_infmatrix_def one_matrix_def)
have aux1: "Rep_matrix (Abs_matrix (λj i. if j = i ∧ j < nrows A then 1 else 0)) =(λj i. if j = i ∧ j < nrows A then 1 else 0::int)"
by (metis (no_types) Rep_one_matrix one_matrix_def)
have "(λi j. if i = n ∧ j = m then 1::int else if i = m ∧ j = n then 1 else if i = j ∧ i ≠ m ∧ i ≠ n ∧ i < nrows A then 1 else 0) =
(λi j. if i = n then (λj i. if j = i ∧ j < nrows A then 1 else 0) m j
else if i = m then (λj i. if j = i ∧ j < nrows A then 1 else 0) n j
else (λj i. if j = i ∧ j < nrows A then 1 else 0) i j)"

apply (rule) using assms by auto
hence " (λi j. if i = n ∧ j = m then 1::int else if i = m ∧ j = n then 1 else if i = j ∧ i ≠ m ∧ i ≠ n ∧ i < nrows A then 1 else 0) =
(λi j. if i = n then Rep_matrix (Abs_matrix (λj i. if j = i ∧ j < nrows A then 1 else 0)) m j
else if i = m then Rep_matrix (Abs_matrix (λj i. if j = i ∧ j < nrows A then 1 else 0)) n j
else Rep_matrix (Abs_matrix (λj i. if j = i ∧ j < nrows A then 1 else 0)) i j)"

using aux1 by force
thus "Abs_matrix (λi j. if i = n ∧ j = m then 1::int else if i = m ∧ j = n then 1 else if i = j ∧ i ≠ m ∧ i ≠ n ∧ i < nrows A then 1 else 0) =
Abs_matrix
(λi j. if i = n then Rep_matrix (Abs_matrix (λj i. if j = i ∧ j < nrows A then 1 else 0)) m j
else if i = m then Rep_matrix (Abs_matrix (λj i. if j = i ∧ j < nrows A then 1 else 0)) n j
else Rep_matrix (Abs_matrix (λj i. if j = i ∧ j < nrows A then 1 else 0)) i j)"
by force
qed


lemma Q_interchange_column_eq:
assumes "n<ncols A"
and "m<ncols A"
shows "Q_interchange_columns' (A::int matrix) n m = Q_interchange_columns A n m"

proof (unfold Q_interchange_columns'_def Q_interchange_columns_def interchange_columns_matrix_def interchange_columns_infmatrix_def one_matrix_def)
have aux1: "Rep_matrix (Abs_matrix (λj i. if j = i ∧ j < ncols A then 1 else 0)) =(λj i. if j = i ∧ j < ncols A then 1 else 0::int)"
by (metis (no_types) Rep_one_matrix one_matrix_def)
have "(λi j. if i = n ∧ j = m then 1::int else if i = m ∧ j = n then 1 else if i = j ∧ i ≠ m ∧ i ≠ n ∧ j < ncols A then 1 else 0) =
(λi j. if j = n then (λj i. if j = i ∧ j < ncols A then 1 else 0) i m
else if j = m then (λj i. if j = i ∧ j < ncols A then 1 else 0) i n
else (λj i. if j = i ∧ j < ncols A then 1 else 0) i j)"
apply rule using assms by auto
hence "(λi j. if i = n ∧ j = m then 1::int else if i = m ∧ j = n then 1 else if i = j ∧ i ≠ m ∧ i ≠ n ∧ j < ncols A then 1 else 0) =
(λi j. if j = n then Rep_matrix (Abs_matrix (λj i. if j = i ∧ j < ncols A then 1 else 0)) i m
else if j = m then Rep_matrix (Abs_matrix (λj i. if j = i ∧ j < ncols A then 1 else 0)) i n
else Rep_matrix (Abs_matrix (λj i. if j = i ∧ j < ncols A then 1 else 0)) i j)"
using aux1 by force
thus "Abs_matrix (λi j. if i = n ∧ j = m then 1::int else if i = m ∧ j = n then 1 else if i = j ∧ i ≠ m ∧ i ≠ n ∧ j < ncols A then 1 else 0) =
Abs_matrix (λi j. if j = n then Rep_matrix (Abs_matrix (λj i. if j = i ∧ j < ncols A then 1 else 0)) i m
else if j = m then Rep_matrix (Abs_matrix (λj i. if j = i ∧ j < ncols A then 1 else 0)) i n
else Rep_matrix (Abs_matrix (λj i. if j = i ∧ j < ncols A then 1 else 0)) i j)"
by auto
qed


lemma Q_column_add_eq:
assumes "n<ncols A"
and "m<ncols A"
and "n≠m"
shows "Q_column_add' (A::int matrix) n m q = Q_column_add A n m q"

proof (unfold Q_column_add'_def Q_column_add_def column_add_matrix_def column_add_infmatrix_def one_matrix_def)
have aux1: "Rep_matrix (Abs_matrix (λj i. if j = i ∧ j < ncols A then 1 else 0)) =(λj i. if j = i ∧ j < ncols A then 1 else 0::int)"
by (metis (no_types) Rep_one_matrix one_matrix_def)
have "(λi j. if i = m ∧ j = n then q else if i = j ∧ j < ncols A then 1 else 0) =
(λa b. if b = n then (λj i. if j = i ∧ j < ncols A then 1 else 0) a n + q * (λj i. if j = i ∧ j < ncols A then 1 else 0) a m
else (λj i. if j = i ∧ j < ncols A then 1 else 0) a b)"
apply rule using assms by auto
hence "(λi j. if i = m ∧ j = n then q else if i = j ∧ j < ncols A then 1 else 0) =
(λa b. if b = n then Rep_matrix (Abs_matrix (λj i. if j = i ∧ j < ncols A then 1 else 0)) a n
+ q * Rep_matrix (Abs_matrix (λj i. if j = i ∧ j < ncols A then 1 else 0)) a m
else Rep_matrix (Abs_matrix (λj i. if j = i ∧ j < ncols A then 1 else 0)) a b)"
using aux1 by force
thus "Abs_matrix (λi j. if i = m ∧ j = n then q else if i = j ∧ j < ncols A then 1 else 0) =
Abs_matrix (λa b. if b = n then Rep_matrix (Abs_matrix (λj i. if j = i ∧ j < ncols A then 1 else 0)) a n
+ q * Rep_matrix (Abs_matrix (λj i. if j = i ∧ j < ncols A then 1 else 0)) a m
else Rep_matrix (Abs_matrix (λj i. if j = i ∧ j < ncols A then 1 else 0)) a b)"
by simp
qed



lemma Rep_abs_aux_P_interchange_rows:
assumes "n<nrows A" and "m<nrows A"
shows "Rep_matrix (Abs_matrix (λi j. if i = n ∧ j = m then 1 else if i = m ∧ j = n then 1 else if i = j ∧ i ≠ m ∧ i ≠ n ∧ i < nrows A then 1 else 0))
=(λi j. if i = n ∧ j = m then 1 else if i = m ∧ j = n then 1 else if i = j ∧ i ≠ m ∧ i ≠ n ∧ i < nrows A then 1 else 0)"

proof (rule RepAbs_matrix)
show "∃ma. ∀j i. ma ≤ j --> (if j = n ∧ i = m then 1::'b else if j = m ∧ i = n then 1::'b else if j = i ∧ j ≠ m ∧ j ≠ n ∧ j < nrows A then 1::'b else (0::'b)) = (0::'b)"
by (metis assms(1) assms(2) less_not_refl order_less_le_trans order_less_trans xt1(7))
thus "∃na. ∀j i. na ≤ i --> (if j = n ∧ i = m then 1::'b else if j = m ∧ i = n then 1::'b else if j = i ∧ j ≠ m ∧ j ≠ n ∧ j < nrows A then 1::'b else (0::'b)) = (0::'b)"
by metis
qed


lemma one_matrix_exists_notnull:
assumes "a<b"
shows "∃i. Rep_matrix ((one_matrix b)::int matrix) a i ≠ 0"

proof (rule exI[of _ a], unfold one_matrix_def)
have "Rep_matrix (Abs_matrix (λj i. if j = i ∧ j < b then 1::int else 0)) = (λj i. if j = i ∧ j < b then 1::int else 0)"
by (metis (no_types) Rep_one_matrix one_matrix_def)
hence "Rep_matrix (Abs_matrix (λj i. if j = i ∧ j < b then 1::int else 0)) a a = (λj i. if j = i ∧ j < b then 1::int else 0) a a" by auto
also have "... = 1" using assms by simp
also have "...≠ 0" by simp
finally show "Rep_matrix (Abs_matrix (λj i. if j = i ∧ j < b then 1::int else 0)) a a ≠ 0" .
qed



lemma interchange_rows_nrows:
assumes n:"n<nrows A"
and m:"m<nrows A"
and i: "∃i. Rep_matrix A (min n m) i ≠ 0"
shows "nrows (interchange_rows_matrix (A::int matrix) n m) = nrows A"

proof (cases "Suc m = nrows A")
case True
have n_le_m: "n≤m" using True n by simp
from this obtain b where not_null: "Rep_matrix A n b ≠ 0" using i by (metis min_max.inf_absorb2 min_max.inf_commute)
have "Rep_matrix (interchange_rows_matrix A n m) m b = Rep_matrix A n b" unfolding interchange_rows_matrix_def interchange_rows_infmatrix_def
unfolding Rep_Abs_aux_interchange2' by simp
also have "...≠ 0" using not_null .
finally have "m < nrows (interchange_rows_matrix (A::int matrix) n m)" using less_nrows[of m "(interchange_rows_matrix (A::int matrix) n m)"] by auto
hence 1:"nrows A ≤ nrows (interchange_rows_matrix (A::int matrix) n m)" using True by simp
have 2:"nrows (interchange_rows_matrix (A::int matrix) n m) ≤ nrows A"
proof -
have "(∀j i. nrows A ≤ j --> Rep_matrix (interchange_rows_matrix A n m) j i = 0)"
unfolding interchange_rows_matrix_def interchange_rows_infmatrix_def Rep_Abs_aux_interchange2'
by (metis Suc_leD True n_le_m not_less_eq_eq nrows)
thus ?thesis
using nrows_le[of "(interchange_rows_matrix (A::int matrix) n m)" "nrows A"] by simp
qed
show ?thesis using 1 and 2 by simp
next
case False note Suc_m_not_nrows_A=False
show ?thesis
proof (cases "Suc n = nrows A")
case True --"Va a ser casi repetir el caso anterior intercambiando n y m"
have m_le_n: "m≤n" using True m by simp
from this obtain b where not_null: "Rep_matrix A m b ≠ 0" using i by (metis min_max.inf_absorb2 min_max.inf_commute)
have "Rep_matrix (interchange_rows_matrix A n m) n b = Rep_matrix A m b" unfolding interchange_rows_matrix_def interchange_rows_infmatrix_def
unfolding Rep_Abs_aux_interchange2' by simp
also have "...≠ 0" using not_null .
finally have "n < nrows (interchange_rows_matrix (A::int matrix) n m)" using less_nrows[of n "(interchange_rows_matrix (A::int matrix) n m)"] by auto
hence 1:"nrows A ≤ nrows (interchange_rows_matrix (A::int matrix) n m)" using True by simp
have 2:"nrows (interchange_rows_matrix (A::int matrix) n m) ≤ nrows A"
proof -
have "(∀j i. nrows A ≤ j --> Rep_matrix (interchange_rows_matrix A n m) j i = 0)"
unfolding interchange_rows_matrix_def interchange_rows_infmatrix_def Rep_Abs_aux_interchange2'
by (metis Suc_leD True m_le_n not_less_eq_eq nrows)
thus ?thesis
using nrows_le[of "(interchange_rows_matrix (A::int matrix) n m)" "nrows A"] by simp
qed
show ?thesis using 1 and 2 by simp
next
case False
obtain a where a:"Suc a = nrows A" using n by (metis gr_implies_not0 nat.exhaust)
obtain c b where b:"Rep_matrix A c b ≠ 0" and "a ≤ c" using a less_nrows[of a A] by auto
have a_eq_c: "a=c"by (metis `a ≤ c` a b le_less_Suc_eq nrows_notzero)
hence b': "Rep_matrix A a b ≠ 0" using b by simp
have "Rep_matrix (interchange_rows_matrix A n m) a b = Rep_matrix A a b"
using False Suc_m_not_nrows_A a
unfolding interchange_rows_matrix_def interchange_rows_infmatrix_def
unfolding Rep_Abs_aux_interchange2' by auto
also have "...≠0" using b' .
finally have "a < nrows (interchange_rows_matrix (A::int matrix) n m)" using less_nrows[of a "(interchange_rows_matrix (A::int matrix) n m)"] by auto
hence 1:"nrows A ≤ nrows (interchange_rows_matrix (A::int matrix) n m)" using a by simp
have 2:"nrows (interchange_rows_matrix (A::int matrix) n m) ≤ nrows A"
proof -
have "(∀j i. nrows A ≤ j --> Rep_matrix (interchange_rows_matrix A n m) j i = 0)"
unfolding interchange_rows_matrix_def interchange_rows_infmatrix_def Rep_Abs_aux_interchange2'
by (metis less_not_refl less_nrows m n order_less_le_trans)
thus ?thesis
using nrows_le[of "(interchange_rows_matrix (A::int matrix) n m)" "nrows A"] by simp
qed
show ?thesis using 1 and 2 by simp
qed
qed



lemma interchange_rows_ncols:
assumes n:"n<nrows A"
and m:"m<nrows A"
shows "ncols (interchange_rows_matrix (A::int matrix) n m) = ncols A"

proof (cases "ncols A = 0")
case True show ?thesis
by (metis True less_nrows less_zeroE m ncols_notzero)
next
case False
from this obtain b where Suc_b:"b+1 = ncols A"
by (metis Suc_eq_plus1 not0_implies_Suc)
obtain a c where not_null: "Rep_matrix A a c ≠ 0" and b_l_c: "b ≤ c"using Suc_b less_ncols[of b A] by auto
hence "c = b"
by (metis Suc_b Suc_eq_plus1 ncols not_less_eq_eq order_eq_iff)
hence not_null_ab:"Rep_matrix A a b ≠ 0" using not_null by simp
have "∃s. Rep_matrix (interchange_rows_matrix (A::int matrix) n m) s b ≠ 0" --"No puedo asegurar que s=a, solo en caso de que a distinto de n y a distinto de m"
proof (cases "a=n")
case True have "Rep_matrix (interchange_rows_matrix (A::int matrix) n m) m b = Rep_matrix A a b"
unfolding interchange_rows_matrix_def interchange_rows_infmatrix_def Rep_Abs_aux_interchange2' True by simp
also have "...≠0" using not_null_ab .
finally show ?thesis by blast
next
case False note a_not_n = False
show ?thesis
proof (cases "a=m")
case True have "Rep_matrix (interchange_rows_matrix (A::int matrix) n m) n b = Rep_matrix A a b"
unfolding interchange_rows_matrix_def interchange_rows_infmatrix_def Rep_Abs_aux_interchange2' True by simp
also have "...≠0" using not_null_ab .
finally show ?thesis by blast
next
case False
have "Rep_matrix (interchange_rows_matrix (A::int matrix) n m) a b = Rep_matrix A a b"
unfolding interchange_rows_matrix_def interchange_rows_infmatrix_def Rep_Abs_aux_interchange2' using False a_not_n by simp
also have "...≠0" using not_null_ab .
finally show ?thesis by blast
qed
qed
from this obtain s where "Rep_matrix (interchange_rows_matrix (A::int matrix) n m) s b ≠ 0" by auto
hence "b < ncols (interchange_rows_matrix (A::int matrix) n m)" using less_ncols[of b "(interchange_rows_matrix (A::int matrix) n m)"] by auto
hence 1:"ncols (interchange_rows_matrix (A::int matrix) n m) ≥ ncols A" using Suc_b by auto
have 2:"ncols (interchange_rows_matrix (A::int matrix) n m) ≤ ncols A"
proof -
have "(∀j i. ncols A ≤ i --> Rep_matrix (interchange_rows_matrix A n m) j i = 0)"
unfolding interchange_rows_matrix_def interchange_rows_infmatrix_def Rep_Abs_aux_interchange2'
by (metis ncols)
thus ?thesis
by (metis ncols_le)
qed
show ?thesis using 1 and 2 by simp
qed

lemma P_interchange_rows_nrows_eq:
assumes "n<nrows A"
and "m<nrows A"
shows "nrows (P_interchange_rows (A::int matrix) n m) = ncols (P_interchange_rows (A::int matrix) n m)"

unfolding P_interchange_rows_def using interchange_rows_nrows interchange_rows_ncols assms nrows_one_matrix ncols_one_matrix
by auto

lemma P_interchange_rows'_nrows_eq:
assumes "n<nrows A"
and "m<nrows A"
shows "nrows (P_interchange_rows' (A::int matrix) n m) = ncols (P_interchange_rows' (A::int matrix) n m)"

by (metis P_interchange_row_eq P_interchange_rows_nrows_eq assms(1) assms(2))

lemma P_interchange_rows_nrows_eq2:
assumes "n<nrows A"
and "m<nrows A"
shows "nrows (P_interchange_rows (A::int matrix) n m) = nrows A"

by (metis (no_types) P_interchange_rows_def P_interchange_rows_nrows_eq assms(1) assms(2) interchange_rows_ncols ncols_one_matrix nrows_one_matrix)

corollary P_interchange_rows'_nrows:
assumes "n<nrows A"
and "m<nrows A"
shows "nrows (P_interchange_rows' (A::int matrix) n m)=nrows (P_interchange_rows (A::int matrix) n m)"

by (metis P_interchange_row_eq assms(1) assms(2))

corollary P_interchange_rows'_nrows_A:
assumes "n<nrows A"
and "m<nrows A"
shows "nrows (P_interchange_rows' (A::int matrix) n m)=nrows A"

by (metis P_interchange_row_eq P_interchange_rows_nrows_eq2 assms(1) assms(2))

corollary P_interchange_rows'_nrows_A_2:
assumes "n<nrows A"
and "m<nrows A"
shows "ncols (P_interchange_rows' (A::int matrix) n m)=nrows A"

by (metis P_interchange_row_eq P_interchange_rows'_nrows_A P_interchange_rows_nrows_eq assms(1) assms(2))



lemma P_interchange_rows_inverse:
assumes n:"n<nrows A"
and m:"m<nrows A"
shows "inverse_matrix (P_interchange_rows' (A::int matrix) n m) (P_interchange_rows' A n m)"

unfolding inverse_matrix_def apply (rule conjI) unfolding right_inverse_matrix_def left_inverse_matrix_def
proof (auto)
show "P_interchange_rows' A n m * P_interchange_rows' A n m = one_matrix (max (nrows (P_interchange_rows' A n m)) (ncols (P_interchange_rows' A n m)))"
proof -
have "Rep_matrix (P_interchange_rows' A n m * P_interchange_rows' A n m)
= Rep_matrix (one_matrix (max (nrows (P_interchange_rows' A n m)) (ncols (P_interchange_rows' A n m))))"

proof (rule,rule)
fix a b
show "Rep_matrix (P_interchange_rows' A n m * P_interchange_rows' A n m) a b =
Rep_matrix (one_matrix (max (nrows (P_interchange_rows' A n m)) (ncols (P_interchange_rows' A n m)))) a b"

proof -
have 1: "Rep_matrix (Abs_matrix (λi j. if i = n ∧ j = m then 1::int else if i = m ∧ j = n then 1 else if i = j ∧ i ≠ m ∧ i ≠ n ∧ i < nrows A then 1 else 0))
= (λi j. if i = n ∧ j = m then 1 else if i = m ∧ j = n then 1 else if i = j ∧ i ≠ m ∧ i ≠ n ∧ i < nrows A then 1 else 0)"

using Rep_abs_aux_P_interchange_rows[OF n m] by simp
have 2: "nrows (Abs_matrix (λi j. if i = n ∧ j = m then 1::int else if i = m ∧ j = n then 1 else if i = j ∧ i ≠ m ∧ i ≠ n ∧ i < nrows A then 1 else 0)) = nrows A"
by (metis (no_types) P_interchange_rows'_def P_interchange_rows'_nrows_A m n)
have 3: "nrows (Abs_matrix (λi j. if i = n ∧ j = m then 1::int else if i = m ∧ j = n then 1 else if i = j ∧ i ≠ m ∧ i ≠ n ∧ i < nrows A then 1 else 0))
= ncols (Abs_matrix (λi j. if i = n ∧ j = m then 1::int else if i = m ∧ j = n then 1 else if i = j ∧ i ≠ m ∧ i ≠ n ∧ i < nrows A then 1 else 0))"

using P_interchange_rows'_nrows_eq[of n A m, OF n m] unfolding P_interchange_rows'_def .
have "Rep_matrix (P_interchange_rows' A n m * P_interchange_rows' A n m) a b
= foldseq op + (λk. Rep_matrix (P_interchange_rows' A n m) a k * Rep_matrix (P_interchange_rows' A n m) k b)
(max (ncols (P_interchange_rows' A n m)) (nrows (P_interchange_rows' A n m)))"
using Rep_matrix_mult by auto
also have "... = Rep_matrix (one_matrix (max (nrows (P_interchange_rows' A n m)) (ncols (P_interchange_rows' A n m)))) a b"
unfolding Rep_one_matrix
unfolding P_interchange_rows'_def
proof (auto)
assume a_ne_b: "a≠b"
show "foldseq op +
(λk. Rep_matrix (Abs_matrix (λi j. if i = n ∧ j = m then 1::int else if i = m ∧ j = n then 1 else if i = j ∧ i ≠ m ∧ i ≠ n ∧ i < nrows A then 1 else 0)) a k *
Rep_matrix (Abs_matrix (λi j. if i = n ∧ j = m then 1 else if i = m ∧ j = n then 1 else if i = j ∧ i ≠ m ∧ i ≠ n ∧ i < nrows A then 1 else 0)) k b)
(max (ncols (Abs_matrix (λi j. if i = n ∧ j = m then 1::int else if i = m ∧ j = n then 1 else if i = j ∧ i ≠ m ∧ i ≠ n ∧ i < nrows A then 1 else 0)))
(nrows (Abs_matrix (λi j. if i = n ∧ j = m then 1::int else if i = m ∧ j = n then 1 else if i = j ∧ i ≠ m ∧ i ≠ n ∧ i < nrows A then 1 else 0)))) = 0"

apply (rule foldseq_zero)
apply simp
unfolding 1 using a_ne_b by force
next
assume "¬ b < max (nrows (Abs_matrix (λi j. if i = n ∧ j = m then 1::int else if i = m ∧ j = n then 1 else if i = j ∧ i ≠ m ∧ i ≠ n ∧ i < nrows A then 1 else 0)))
(ncols (Abs_matrix (λi j. if i = n ∧ j = m then 1::int else if i = m ∧ j = n then 1 else if i = j ∧ i ≠ m ∧ i ≠ n ∧ i < nrows A then 1 else 0)))"

hence b_ge_max:
"b ≥ max (nrows (Abs_matrix (λi j. if i = n ∧ j = m then 1::int else if i = m ∧ j = n then 1 else if i = j ∧ i ≠ m ∧ i ≠ n ∧ i < nrows A then 1 else 0)))
(ncols (Abs_matrix (λi j. if i = n ∧ j = m then 1::int else if i = m ∧ j = n then 1 else if i = j ∧ i ≠ m ∧ i ≠ n ∧ i < nrows A then 1 else 0)))"
by auto
show "foldseq op +
(λk. Rep_matrix (Abs_matrix (λi j. if i = n ∧ j = m then 1::int else if i = m ∧ j = n then 1 else if i = j ∧ i ≠ m ∧ i ≠ n ∧ i < nrows A then 1 else 0))
a k *
Rep_matrix (Abs_matrix (λi j. if i = n ∧ j = m then 1 else if i = m ∧ j = n then 1 else if i = j ∧ i ≠ m ∧ i ≠ n ∧ i < nrows A then 1 else 0)) k b)
(max (ncols (Abs_matrix (λi j. if i = n ∧ j = m then 1::int else if i = m ∧ j = n then 1 else if i = j ∧ i ≠ m ∧ i ≠ n ∧ i < nrows A then 1 else 0)))
(nrows (Abs_matrix (λi j. if i = n ∧ j = m then 1::int else if i = m ∧ j = n then 1 else if i = j ∧ i ≠ m ∧ i ≠ n ∧ i < nrows A then 1 else 0)))) = 0"

apply (rule foldseq_zero)
apply simp
unfolding 1
using b_ge_max
using n m 2 by auto
next
assume b:
"b < max (nrows (Abs_matrix (λi j. if i = n ∧ j = m then 1::int else if i = m ∧ j = n then 1 else if i = j ∧ i ≠ m ∧ i ≠ n ∧ i < nrows A then 1 else 0)))
(ncols (Abs_matrix (λi j. if i = n ∧ j = m then 1::int else if i = m ∧ j = n then 1 else if i = j ∧ i ≠ m ∧ i ≠ n ∧ i < nrows A then 1 else 0)))"

def P == "(λk. Rep_matrix
(Abs_matrix (λi j. if i = n ∧ j = m then 1 else if i = m ∧ j = n then 1::int else if i = j ∧ i ≠ m ∧ i ≠ n ∧ i < nrows A then 1 else 0)) b k *
Rep_matrix (Abs_matrix (λi j. if i = n ∧ j = m then 1 else if i = m ∧ j = n then 1::int else if i = j ∧ i ≠ m ∧ i ≠ n ∧ i < nrows A then 1 else 0)) k b)"

def max_n "(max (ncols (Abs_matrix (λi j. if i = n ∧ j = m then 1::int else if i = m ∧ j = n then 1 else if i = j ∧ i ≠ m ∧ i ≠ n ∧ i < nrows A
then 1 else 0)))
(nrows (Abs_matrix (λi j. if i = n ∧ j = m then 1 else if i = m ∧ j = n then 1::int else if i = j ∧ i ≠ m ∧ i ≠ n ∧ i < nrows A then 1 else 0))))"

have "∃k. (P k = 1 ∧ (∀i. i≠k --> P i = 0))"
proof (cases "b=n")
case True
have "P m = 1" unfolding P_def unfolding 1 unfolding True by simp
moreover have"(∀i. i≠m --> P i = 0)" unfolding P_def unfolding 1 unfolding True by auto
ultimately show ?thesis by auto
next
case False note b_not_n = False
show ?thesis
proof (cases "b=m")
case True
have "P n = 1" unfolding P_def unfolding 1 unfolding True by simp
moreover have"(∀i. i≠n --> P i = 0)" unfolding P_def unfolding 1 unfolding True by auto
ultimately show ?thesis by auto
next
case False
have "P b = 1" unfolding P_def unfolding 1 using False b_not_n using 2 using 3 b by auto
moreover have"(∀i. i≠b --> P i = 0)" unfolding P_def unfolding 1 using False b_not_n by auto
ultimately show ?thesis by auto
qed
qed
from this obtain k where P_k_1: "P k = 1" and zero: "(∀i. i≠k --> P i = 0)" by blast
have k_ke_max: "k ≤ max_n"
proof (rule ccontr)
assume k_g_max: "¬ k ≤ max_n"
hence "k > nrows (Abs_matrix (λi j. if i = n ∧ j = m then 1 else if i = m ∧ j = n then 1::int else if i = j ∧ i ≠ m ∧ i ≠ n ∧ i < nrows A then 1 else 0))"
unfolding max_n_def using 3 by simp
hence "P k = 0"
by (metis (no_types) "2" "3" P_def P_interchange_rows'_def P_k_1 less_imp_le_nat mult_zero_left ncols zero_neq_one)
thus False using P_k_1 by auto
qed
have "foldseq op + P max_n = (if k ≤ max_n then P k else 0)" using foldseq_almostzero [of "op +" _ P max_n, OF _ _ zero] by auto
also have "... = P k" using k_ke_max by simp
also have "...=1"using P_k_1 .
finally show "foldseq op +
(λk. Rep_matrix (Abs_matrix (λi j. if i = n ∧ j = m then 1 else if i = m ∧ j = n then 1::int
else if i = j ∧ i ≠ m ∧ i ≠ n ∧ i < nrows A then 1 else 0)) b k *
Rep_matrix (Abs_matrix (λi j. if i = n ∧ j = m then 1 else if i = m ∧ j = n then 1::int else if i = j ∧ i ≠ m ∧ i ≠ n ∧ i < nrows A then 1 else 0)) k b)
(max (ncols (Abs_matrix (λi j. if i = n ∧ j = m then 1::int else if i = m ∧ j = n then 1 else if i = j ∧ i ≠ m ∧ i ≠ n ∧ i < nrows A then 1 else 0)))
(nrows (Abs_matrix (λi j. if i = n ∧ j = m then 1::int else if i = m ∧ j = n then 1 else if i = j ∧ i ≠ m ∧ i ≠ n ∧ i < nrows A then 1 else 0)))) =
1"
unfolding P_def max_n_def by auto
qed
finally show ?thesis .
qed
qed
thus ?thesis using Rep_matrix_inject by auto
qed
thus "P_interchange_rows' A n m * P_interchange_rows' A n m = one_matrix (max (nrows (P_interchange_rows' A n m)) (ncols (P_interchange_rows' A n m)))" .
show "nrows (P_interchange_rows' A n m) ≤ ncols (P_interchange_rows' A n m)"
by (metis P_interchange_rows'_nrows_eq m n order_refl)
show "ncols (P_interchange_rows' A n m) ≤ nrows (P_interchange_rows' A n m)"
by (metis P_interchange_rows'_nrows_eq m n order_refl)
qed


lemma P_row_add_nrows:
assumes n:"n<nrows A"
and m:"m<nrows A"
and n_not_m: "n≠m"
shows "nrows (P_row_add (A::int matrix) n m q) = nrows A"

proof -
obtain a where a:"Suc a = nrows A" using n
by (metis gr_implies_not0 not0_implies_Suc)
have "Rep_matrix (P_row_add (A::int matrix) n m q) a a ≠ 0"
proof (cases "a=n")
case True thus ?thesis unfolding P_row_add_def row_add_matrix_def row_add_infmatrix_def Rep_aux_row_add using n m n_not_m by auto
next
case False thus ?thesis unfolding P_row_add_def row_add_matrix_def row_add_infmatrix_def Rep_aux_row_add using n m n_not_m a by auto
qed
hence "nrows (P_row_add A n m q) ≥ nrows A"
by (metis a not_less_eq_eq nrows)
moreover have "nrows (P_row_add A n m q) ≤ nrows A"
proof -
have "(∀j i. nrows A ≤ j --> Rep_matrix (P_row_add A n m q) j i = 0)"
unfolding P_row_add_def row_add_matrix_def row_add_infmatrix_def Rep_aux_row_add using n by auto
thus ?thesis
using nrows_le[of "(P_row_add (A::int matrix) n m q)" "nrows A"] by auto
qed
ultimately show ?thesis by auto
qed

lemma P_row_add_ncols:
assumes n:"n<nrows A"
and m:"m<nrows A"
and n_not_m: "n≠m"
shows "ncols (P_row_add (A::int matrix) n m q) = nrows A"

proof -
obtain a where a:"Suc a = nrows A" using n
by (metis gr_implies_not0 not0_implies_Suc)
have "Rep_matrix (P_row_add (A::int matrix) n m q) a a ≠ 0"
proof (cases "a=n")
case True thus ?thesis unfolding P_row_add_def row_add_matrix_def row_add_infmatrix_def Rep_aux_row_add using n m n_not_m by auto
next
case False thus ?thesis unfolding P_row_add_def row_add_matrix_def row_add_infmatrix_def Rep_aux_row_add using n m n_not_m a by auto
qed
hence "ncols (P_row_add A n m q) ≥ nrows A"
by (metis a ncols not_less_eq_eq)
moreover have "ncols (P_row_add A n m q) ≤ nrows A"
proof -
have "(∀j i. nrows A ≤ i --> Rep_matrix (P_row_add A n m q) j i = 0)"
unfolding P_row_add_def row_add_matrix_def row_add_infmatrix_def Rep_aux_row_add using n by auto
thus ?thesis
using ncols_le[of "(P_row_add (A::int matrix) n m q)" "nrows A"] by auto
qed
ultimately show ?thesis by auto
qed

corollary P_row_add_nrows_eq_ncols:
assumes n:"n<nrows A"
and m:"m<nrows A"
and n_not_m: "n≠m"
shows "nrows (P_row_add (A::int matrix) n m q) = ncols (P_row_add (A::int matrix) n m q)"

by (metis P_row_add_ncols P_row_add_nrows m n n_not_m)

corollary P_row_add'_nrows_eq_ncols:
assumes n:"n<nrows A"
and m:"m<nrows A"
and n_not_m: "n≠m"
shows "nrows (P_row_add' (A::int matrix) n m q) = ncols (P_row_add' (A::int matrix) n m q)"

by (metis P_row_add_eq P_row_add_ncols P_row_add_nrows m n n_not_m)


lemma Rep_abs_aux_P_row_add':
assumes "n<nrows A"
and "m<nrows A"
shows "Rep_matrix (Abs_matrix (λi j. if i = n ∧ j = m then q else if i = j ∧ i < nrows A then 1::int else 0))
=(λi j. if i = n ∧ j = m then q else if i = j ∧ i < nrows A then 1 else 0)"

proof (rule RepAbs_matrix)
show "∃ma. ∀j i. ma ≤ j --> (if j = n ∧ i = m then q else if j = i ∧ j < nrows A then 1 else 0) = 0"
by (metis assms(1) leD)
show "∃na. ∀j i. na ≤ i --> (if j = n ∧ i = m then q else if j = i ∧ j < nrows A then 1 else 0) = 0"
by (metis assms(2) leD)
qed

corollary P_row_add'_nrows_eq_nrows_inverse:
assumes n:"n<nrows A"
and m:"m<nrows A"
and n_not_m: "n≠m"
shows "nrows (P_row_add' (A::int matrix) n m q) = nrows (P_row_add' A n m (-q))"

by (metis P_row_add_eq P_row_add_nrows m n n_not_m)

corollary P_row_add'_nrows_eq_ncols_inverse:
assumes n:"n<nrows A"
and m:"m<nrows A"
and n_not_m: "n≠m"
shows "ncols (P_row_add' (A::int matrix) n m q) = ncols (P_row_add' A n m (-q))"

by (metis P_row_add_eq P_row_add_ncols m n n_not_m)


lemma foldseq_aux:
fixes s:: "nat => int"
shows "foldseq (op +) s (Suc n) = (op +) (s (Suc n)) (foldseq (op +) s n)"

proof (induct n)
case 0 show ?case by simp
next
case (Suc n)
have "foldseq op + s (Suc (Suc n)) = foldseq_transposed op + s (Suc (Suc n))"
by (metis associative_def foldseq_assoc zadd_assoc)
also have "... = (op +) (foldseq_transposed (op +) s (Suc n)) (s (Suc (Suc n)))" by simp
also have "... = s (Suc (Suc n)) + (foldseq_transposed (op +) s (Suc n))" by simp
also have "... = s (Suc (Suc n)) + (foldseq (op +) s (Suc n))"
by (metis associative_def foldseq_assoc zadd_assoc)
finally show ?case .
qed




lemma foldseq_finsum:
fixes s:: "nat => int"
shows "foldseq (op +) s n = finsum \<Z> s {..n}"

proof (induct n)
have 1: "{0::nat}={..0::nat}"
by (metis atMost_0)
show "foldseq op + s 0 = (\<Oplus>\<Z>k∈{..0}. s k)" apply auto unfolding 1
proof (rule Ring.abelian_monoid.finsum_0[of "\<Z>" s, symmetric], auto)
show "abelian_monoid \<Z>" by default
qed
next
case (Suc n)
have "foldseq op + s (Suc n) = s (Suc n) + (foldseq (op +) s n)" using foldseq_aux by auto
also have "...= s (Suc n) + finsum \<Z> s {..n}" using Suc.hyps by simp
also have "... = s (Suc n) ⊕\<Z> finsum \<Z> s {..n}" by auto
also have "...= finsum \<Z> s (insert (Suc n) {..n})"
proof (rule Ring.abelian_monoid.finsum_insert[of "\<Z>" "{..n}" "Suc n" "s", symmetric])
show "abelian_monoid \<Z>" by default
show "finite {..n}" by auto
show "Suc n ∉ {..n}" by auto
show "s ∈ {..n} -> carrier \<Z>" by auto
show "s (Suc n) ∈ carrier \<Z>" by auto
qed
also have "... = finsum \<Z> s {..(Suc n)}"
by (metis atMost_Suc)
finally show ?case .
qed

lemma abelian_monoid_z:
shows "abelian_monoid \<Z>"

by default

lemma finsum_singleton2:
shows "finsum \<Z> s {a} = s a"

proof -
have "finsum \<Z> s {a} = finsum \<Z> s (insert a ({a}-{a}))" by auto
also have "... = s a ⊕\<Z> finsum \<Z> s ({a}-{a})" proof (rule Ring.abelian_monoid.finsum_insert, auto)
show "abelian_monoid \<Z>" by default qed
also have "... =s a ⊕\<Z> finsum \<Z> s {}"by auto
also have "... =s a ⊕\<Z> \<zero>\<Z>" using Ring.abelian_monoid.finsum_empty[of \<Z> s] abelian_monoid_z by auto
also have "... = s a" by auto
finally show ?thesis .
qed

lemma foldseq_almostzero_2:
fixes s:: "nat => int"
assumes s0:"∀i. i ≠ j ∧ i ≠ k --> s i = 0"
and j:"j≤n" and k:"k≤n" and j_k: "j≠k"
shows "foldseq (op +) s n = (s j) + (s k)"

proof -
have eq_set:"{..n}-{j}={k} ∪ ({..n}-{j}-{k})"
by (metis Diff_iff atMost_iff insert_Diff_single insert_absorb insert_is_Un j_k k singletonE)
have 1: "finsum \<Z> s ({k} ∪ ({..n}-{j}-{k}))=finsum \<Z> s {k} ⊕\<Z> finsum \<Z> s ({..n}-{j}-{k})"
proof (rule Ring.abelian_monoid.finsum_Un_disjoint, auto)
show "abelian_monoid \<Z>" by default
qed

have eq_zero: "finsum \<Z> s ({..n}-{j}-{k}) = \<zero>\<Z>"
proof -
have "(\<Oplus>\<Z>a∈({..n}-{j}-{k}). s a) = (\<Oplus>\<Z>a∈({..n}-{j}-{k}). \<zero>\<Z>)"
proof (rule Ring.abelian_monoid.finsum_cong', auto)
show "abelian_monoid \<Z>" by default
show "!!i. [|i ≤ n; i ≠ j; i ≠ k|] ==> s i = 0" using s0 by auto
qed
also have "...=\<zero>\<Z>"
proof (rule Ring.abelian_monoid.finsum_zero, auto)
show "abelian_monoid \<Z>" by default
qed
finally show ?thesis .
qed

have "foldseq (op +) s n = finsum \<Z> s {..n}" using foldseq_finsum by simp find_theorems "finsum"
also have "...= finsum \<Z> s ({j} ∪ ({..n}-{j}))"
by (metis (full_types) Collect_def Un_def atMost_iff calculation foldseq_finsum insert_Diff_single insert_absorb insert_compr insert_is_Un j set_diff_eq)
also have "... = finsum \<Z> s {j} ⊕\<Z> finsum \<Z> s ({..n}-{j})" thm Ring.abelian_monoid.finsum_zero[of "\<Z>"]
proof (rule Ring.abelian_monoid.finsum_Un_disjoint, auto)
show "abelian_monoid \<Z>" by default
qed
also have "... = finsum \<Z> s {j} ⊕\<Z> finsum \<Z> s ({k} ∪ ({..n}-{j}-{k}))" using eq_set by simp
also have "...=finsum \<Z> s {j} ⊕\<Z> finsum \<Z> s {k} ⊕\<Z> finsum \<Z> s ({..n}-{j}-{k})" using 1 by auto
also have "...= finsum \<Z> s {j} ⊕\<Z> finsum \<Z> s {k} ⊕\<Z> \<zero>\<Z>" unfolding eq_zero ..
also have "... = finsum \<Z> s {j} ⊕\<Z> finsum \<Z> s {k}"
by (metis UNIV(1) int.a_comm int.l_zero ring.simps(1) ring.simps(2))
also have "... = s j ⊕\<Z> finsum \<Z> s {k}" using finsum_singleton2 by auto
also have "...=s j ⊕\<Z> s k" using finsum_singleton2 by auto
finally show ?thesis by auto
qed


lemma P_row_add_inverse:
assumes n:"n<nrows A"
and m:"m<nrows A"
and n_not_m: "n≠m"
shows "inverse_matrix (P_row_add' (A::int matrix) n m q) (P_row_add' A n m (-q))"

unfolding inverse_matrix_def apply (rule conjI) unfolding right_inverse_matrix_def left_inverse_matrix_def
proof (auto)
show "P_row_add' A n m q * P_row_add' A n m (- q) = one_matrix (max (nrows (P_row_add' A n m q)) (ncols (P_row_add' A n m (- q))))"
proof -
have" Rep_matrix (P_row_add' A n m q * P_row_add' A n m (- q)) = Rep_matrix (one_matrix (max (nrows (P_row_add' A n m q)) (ncols (P_row_add' A n m (- q)))))"
proof (rule,rule)
fix a b
have 1: "nrows (Abs_matrix (λi j. if i = n ∧ j = m then q else if i = j ∧ i < nrows A then 1 else 0))
= ncols (Abs_matrix (λi j. if i = n ∧ j = m then - q else if i = j ∧ i < nrows A then 1 else 0))"

by (metis (no_types) P_row_add'_def P_row_add_eq P_row_add_ncols P_row_add_nrows m n n_not_m)
have 2: "nrows (Abs_matrix (λi j. if i = n ∧ j = m then q else if i = j ∧ i < nrows A then 1 else 0)) = nrows A"
by (metis (no_types) P_row_add'_def P_row_add_eq P_row_add_nrows m n n_not_m)
have 3: "nrows (Abs_matrix (λi j. if i = n ∧ j = m then q else if i = j ∧ i < nrows A then 1 else 0))
= nrows (Abs_matrix (λi j. if i = n ∧ j = m then (-q) else if i = j ∧ i < nrows A then 1 else 0))"

by (metis (no_types) P_row_add'_def P_row_add_eq P_row_add_nrows m n n_not_m)
have "Rep_matrix (P_row_add' A n m q * P_row_add' A n m (- q)) a b = foldseq op + (λk. Rep_matrix (P_row_add' A n m q) a k
* Rep_matrix (P_row_add' A n m (- q)) k b) (max (ncols (P_row_add' A n m q)) (nrows (P_row_add' A n m (- q))))"
unfolding Rep_matrix_mult ..
also have"...=Rep_matrix (one_matrix (max (nrows (P_row_add' A n m q)) (ncols (P_row_add' A n m (- q))))) a b"
unfolding Rep_one_matrix
unfolding P_row_add'_def
unfolding Rep_abs_aux_P_row_add'[OF n m]
proof (auto) --"Tengo los mismos tres casos que en la demostracion con el intercambiar filas"
assume b:"¬ b < max (nrows (Abs_matrix (λi j. if i = n ∧ j = m then q else if i = j ∧ i < nrows A then 1 else 0)))
(ncols (Abs_matrix (λi j. if i = n ∧ j = m then - q else if i = j ∧ i < nrows A then 1 else 0)))"

show "foldseq op + (λk. (if a = n ∧ k = m then q else if a = k ∧ a < nrows A then 1 else 0) * (if k = n ∧ b = m then - q else if k = b ∧ k < nrows A then 1 else 0))
(max (ncols (Abs_matrix (λi j. if i = n ∧ j = m then q else if i = j ∧ i < nrows A then 1 else 0)))
(nrows (Abs_matrix (λi j. if i = n ∧ j = m then - q else if i = j ∧ i < nrows A then 1 else 0)))) =
0"

apply (rule foldseq_zero)
apply simp
using b
using 1 2 n m by auto
next
assume b: "b < max (nrows (Abs_matrix (λi j. if i = n ∧ j = m then q else if i = j ∧ i < nrows A then 1 else 0)))
(ncols (Abs_matrix (λi j. if i = n ∧ j = m then - q else if i = j ∧ i < nrows A then 1 else 0)))"

def P "(λk. (if b = n ∧ k = m then q else if b = k ∧ b < nrows A then 1 else 0) * (if k = n ∧ b = m then - q else if k = b ∧ k < nrows A then 1 else 0))"
def max_n "(max (ncols (Abs_matrix (λi j. if i = n ∧ j = m then q else if i = j ∧ i < nrows A then 1 else 0)))
(nrows (Abs_matrix (λi j. if i = n ∧ j = m then - q else if i = j ∧ i < nrows A then 1 else 0))))"

have "∃k. (P k = 1 ∧ (∀i. i≠k --> P i = 0))"
proof (rule exI[of _ "b"], unfold P_def, rule conjI)
show "(if b = n ∧ b = m then q else if b = b ∧ b < nrows A then 1 else 0) * (if b = n ∧ b = m then - q else if b = b ∧ b < nrows A then 1 else 0) = 1"
using n_not_m b 1 2 by auto
show " ∀i. i ≠ b --> (if b = n ∧ i = m then q else if b = i ∧ b < nrows A then 1 else 0) * (if i = n ∧ b = m then - q else if i = b ∧ i < nrows A then 1 else 0) = 0"
using n_not_m b 1 2 by auto
qed
from this obtain k where P_k_1: "P k = 1" and zero: "(∀i. i≠k --> P i = 0)" by blast
have k_ke_max: "k ≤ max_n"
proof (rule ccontr)
assume "¬ k ≤ max_n"
hence "k>nrows (Abs_matrix (λi j. if i = n ∧ j = m then q else if i = j ∧ i < nrows A then 1 else 0))" unfolding max_n_def using 3 by auto
hence "P k = 0" unfolding P_def max_n_def using 2 m by auto
thus False using P_k_1 by simp
qed
have "foldseq op + P max_n = (if k ≤ max_n then P k else 0)" using foldseq_almostzero [of "op +" _ P max_n, OF _ _ zero] by auto
also have "... = P k" using k_ke_max by simp
also have "...=1"using P_k_1 .
finally show "foldseq op + (λk. (if b = n ∧ k = m then q else if b = k ∧ b < nrows A then 1 else 0)
* (if k = n ∧ b = m then - q else if k = b ∧ k < nrows A then 1 else 0))
(max (ncols (Abs_matrix (λi j. if i = n ∧ j = m then q else if i = j ∧ i < nrows A then 1 else 0)))
(nrows (Abs_matrix (λi j. if i = n ∧ j = m then - q else if i = j ∧ i < nrows A then 1 else 0)))) = 1"

unfolding P_def max_n_def .
next
assume a_not_b: "a ≠ b"
show "foldseq op + (λk. (if a = n ∧ k = m then q else if a = k ∧ a < nrows A then 1 else 0) * (if k = n ∧ b = m then - q else if k = b ∧ k < nrows A then 1 else 0))
(max (ncols (Abs_matrix (λi j. if i = n ∧ j = m then q else if i = j ∧ i < nrows A then 1 else 0)))
(nrows (Abs_matrix (λi j. if i = n ∧ j = m then - q else if i = j ∧ i < nrows A then 1 else 0)))) = 0"

proof (cases "a=n ∧ b=m")
case False
show ?thesis
apply (rule foldseq_zero)
apply simp
using Rep_abs_aux_P_row_add'[OF n m] a_not_b n m n_not_m False
by auto
next
case True
def P "(λk. (if a = n ∧ k = m then q else if a = k ∧ a < nrows A then 1 else 0) * (if k = n ∧ b = m then - q else if k = b ∧ k < nrows A then 1 else 0))"
def max_n "(max (ncols (Abs_matrix (λi j. if i = n ∧ j = m then q else if i = j ∧ i < nrows A then 1 else 0)))
(nrows (Abs_matrix (λi j. if i = n ∧ j = m then - q else if i = j ∧ i < nrows A then 1 else 0))))"

have "foldseq op + P max_n = P a + P b"
proof (rule foldseq_almostzero_2)
show "∀i. i ≠ a ∧ i ≠ b --> P i = 0" unfolding P_def using True by auto
show "a ≤ max_n" unfolding max_n_def using True n using 2 3 by auto
show "b ≤ max_n"unfolding max_n_def using True m using 2 3 by auto
show "a ≠ b" using True n_not_m by simp
qed
also have "... = 0" unfolding P_def using True using n m n_not_m by auto
finally show ?thesis unfolding P_def max_n_def .
qed
qed
finally show "Rep_matrix (P_row_add' A n m q * P_row_add' A n m (- q)) a b
= Rep_matrix (one_matrix (max (nrows (P_row_add' A n m q)) (ncols (P_row_add' A n m (- q))))) a b"
.
qed
thus ?thesis using Rep_matrix_inject by blast
qed
next
show "nrows (P_row_add' A n m (- q)) ≤ ncols (P_row_add' A n m q)"
by (metis P_row_add_eq P_row_add_ncols P_row_add_nrows less_or_eq_imp_le m n n_not_m)
show "ncols (P_row_add' A n m (- q)) ≤ nrows (P_row_add' A n m q)"
by (metis P_row_add'_nrows_eq_ncols `nrows (P_row_add' A n m (- q)) ≤ ncols (P_row_add' A n m q)` m n n_not_m)
next
--"EL SIGUIENTE SHOW ES LA PARTE DE LA MULTIPLICACION A IZQUIERDA DE LA INVERSA, LA DEMOSTRACION ES MUY PARECIDA A LA ANTERIOR."

show "P_row_add' A n m (- q) * P_row_add' A n m q = one_matrix (max (nrows (P_row_add' A n m (- q))) (ncols (P_row_add' A n m q)))"
proof -
have" Rep_matrix (P_row_add' A n m (-q) * P_row_add' A n m q) = Rep_matrix (one_matrix (max (nrows (P_row_add' A n m (-q))) (ncols (P_row_add' A n m q))))"
proof (rule,rule)
fix a b
have 1: "nrows (Abs_matrix (λi j. if i = n ∧ j = m then q else if i = j ∧ i < nrows A then 1 else 0))
= ncols (Abs_matrix (λi j. if i = n ∧ j = m then - q else if i = j ∧ i < nrows A then 1 else 0))"

by (metis (no_types) P_row_add'_def P_row_add_eq P_row_add_ncols P_row_add_nrows m n n_not_m)
have 2: "nrows (Abs_matrix (λi j. if i = n ∧ j = m then q else if i = j ∧ i < nrows A then 1 else 0)) = nrows A"
by (metis (no_types) P_row_add'_def P_row_add_eq P_row_add_nrows m n n_not_m)
have 3: "nrows (Abs_matrix (λi j. if i = n ∧ j = m then q else if i = j ∧ i < nrows A then 1 else 0))
= nrows (Abs_matrix (λi j. if i = n ∧ j = m then (-q) else if i = j ∧ i < nrows A then 1 else 0))"

by (metis (no_types) P_row_add'_def P_row_add_eq P_row_add_nrows m n n_not_m)
have "Rep_matrix (P_row_add' A n m (-q) * P_row_add' A n m q) a b = foldseq op + (λk. Rep_matrix (P_row_add' A n m (-q)) a k
* Rep_matrix (P_row_add' A n m q) k b) (max (ncols (P_row_add' A n m (-q))) (nrows (P_row_add' A n m q)))"
unfolding Rep_matrix_mult ..
also have"...=Rep_matrix (one_matrix (max (nrows (P_row_add' A n m (-q))) (ncols (P_row_add' A n m q)))) a b"
unfolding Rep_one_matrix
unfolding P_row_add'_def
unfolding Rep_abs_aux_P_row_add'[OF n m]
proof (auto) --"Tengo los mismos tres casos que en la demostracion con el intercambiar filas"
assume b:"¬ b < max (nrows (Abs_matrix (λi j. if i = n ∧ j = m then -q else if i = j ∧ i < nrows A then 1 else 0)))
(ncols (Abs_matrix (λi j. if i = n ∧ j = m then q else if i = j ∧ i < nrows A then 1 else 0)))"

show "foldseq op + (λk. (if a = n ∧ k = m then -q else if a = k ∧ a < nrows A then 1 else 0) * (if k = n ∧ b = m then q else if k = b ∧ k < nrows A then 1 else 0))
(max (ncols (Abs_matrix (λi j. if i = n ∧ j = m then -q else if i = j ∧ i < nrows A then 1 else 0)))
(nrows (Abs_matrix (λi j. if i = n ∧ j = m then q else if i = j ∧ i < nrows A then 1 else 0)))) = 0"

apply (rule foldseq_zero)
apply simp
using b
using 1 2 n m 3 by force
next
assume b: "b < max (nrows (Abs_matrix (λi j. if i = n ∧ j = m then -q else if i = j ∧ i < nrows A then 1 else 0)))
(ncols (Abs_matrix (λi j. if i = n ∧ j = m then q else if i = j ∧ i < nrows A then 1 else 0)))"

def P "(λk. (if b = n ∧ k = m then -q else if b = k ∧ b < nrows A then 1 else 0) * (if k = n ∧ b = m then q else if k = b ∧ k < nrows A then 1 else 0))"
def max_n "(max (ncols (Abs_matrix (λi j. if i = n ∧ j = m then -q else if i = j ∧ i < nrows A then 1 else 0)))
(nrows (Abs_matrix (λi j. if i = n ∧ j = m then q else if i = j ∧ i < nrows A then 1 else 0))))"

have b2:"b<nrows A"
by (metis (no_types) "2" "3" P_row_add'_def P_row_add'_nrows_eq_ncols b m min_max.sup.idem n n_not_m)
have "∃k. (P k = 1 ∧ (∀i. i≠k --> P i = 0))"
proof (rule exI[of _ "b"], unfold P_def, rule conjI)
show "(if b = n ∧ b = m then -q else if b = b ∧ b < nrows A then 1 else 0) * (if b = n ∧ b = m then q else if b = b ∧ b < nrows A then 1 else 0) = 1"
using n_not_m b2 by auto
show " ∀i. i ≠ b --> (if b = n ∧ i = m then - q else if b = i ∧ b < nrows A then 1 else 0) * (if i = n ∧ b = m then q else if i = b ∧ i < nrows A then 1 else 0) = 0"
using n_not_m by auto
qed
from this obtain k where P_k_1: "P k = 1" and zero: "(∀i. i≠k --> P i = 0)" by blast
have k_ke_max: "k ≤ max_n"
proof (rule ccontr)
assume "¬ k ≤ max_n"
hence "k>nrows A"
by (metis (no_types) "1" "2" linorder_not_le max_n_def min_max.sup.idem)
hence "P k = 0" unfolding P_def using n m by auto
thus False using P_k_1 by simp
qed
have "foldseq op + P max_n = (if k ≤ max_n then P k else 0)" using foldseq_almostzero [of "op +" _ P max_n, OF _ _ zero] by auto
also have "... = P k" using k_ke_max by simp
also have "...=1"using P_k_1 .
finally show "foldseq op + (λk. (if b = n ∧ k = m then -q else if b = k ∧ b < nrows A then 1 else 0)
* (if k = n ∧ b = m then q else if k = b ∧ k < nrows A then 1 else 0))
(max (ncols (Abs_matrix (λi j. if i = n ∧ j = m then -q else if i = j ∧ i < nrows A then 1 else 0)))
(nrows (Abs_matrix (λi j. if i = n ∧ j = m then q else if i = j ∧ i < nrows A then 1 else 0)))) = 1"

unfolding P_def max_n_def .
next
assume a_not_b: "a ≠ b"
show "foldseq op + (λk. (if a = n ∧ k = m then -q else if a = k ∧ a < nrows A then 1 else 0) * (if k = n ∧ b = m then q else if k = b ∧ k < nrows A then 1 else 0))
(max (ncols (Abs_matrix (λi j. if i = n ∧ j = m then -q else if i = j ∧ i < nrows A then 1 else 0)))
(nrows (Abs_matrix (λi j. if i = n ∧ j = m then q else if i = j ∧ i < nrows A then 1 else 0)))) = 0"

proof (cases "a=n ∧ b=m")
case False
show ?thesis
apply (rule foldseq_zero)
apply simp
using Rep_abs_aux_P_row_add'[OF n m] a_not_b n m n_not_m False
by auto
next
case True
def P "(λk. (if a = n ∧ k = m then -q else if a = k ∧ a < nrows A then 1 else 0) * (if k = n ∧ b = m then q else if k = b ∧ k < nrows A then 1 else 0))"
def max_n "(max (ncols (Abs_matrix (λi j. if i = n ∧ j = m then -q else if i = j ∧ i < nrows A then 1 else 0)))
(nrows (Abs_matrix (λi j. if i = n ∧ j = m then q else if i = j ∧ i < nrows A then 1 else 0))))"

have "foldseq op + P max_n = P a + P b"
proof (rule foldseq_almostzero_2)
show "∀i. i ≠ a ∧ i ≠ b --> P i = 0" unfolding P_def using True by auto
show "a ≤ max_n" unfolding max_n_def using True n using 2 3 by auto
show "b ≤ max_n"unfolding max_n_def using True m using 2 3 by auto
show "a ≠ b" using True n_not_m by simp
qed
also have "... = 0" unfolding P_def using True using n m n_not_m by auto
finally show ?thesis unfolding P_def max_n_def .
qed
qed
finally show "Rep_matrix (P_row_add' A n m (-q)* P_row_add' A n m q) a b
= Rep_matrix (one_matrix (max (nrows (P_row_add' A n m (-q))) (ncols (P_row_add' A n m (q))))) a b"
.
qed
thus ?thesis using Rep_matrix_inject by blast
qed
qed

lemma Rep_matrix_mult_finsum:
shows "Rep_matrix (A * B) j i = finsum \<Z> (λk. Rep_matrix A j k * Rep_matrix B k i) {..(max (ncols A) (nrows B))}"

unfolding Rep_matrix_mult foldseq_finsum ..

definition is_invertible :: "int matrix => bool"
where "is_invertible A = (∃B. inverse_matrix A B)"


lemma interchange_rows_PA:
assumes n:"n<nrows A"
and m:"m<nrows A"
shows "interchange_rows_matrix (A::int matrix) n m = P_interchange_rows' A n m * A"

proof -
have "Rep_matrix (interchange_rows_matrix A n m) = Rep_matrix (P_interchange_rows' A n m * A)"
proof (rule, rule)
fix a b
have max_eq: "max (ncols (Abs_matrix (λi j. if i = n ∧ j = m then 1 else if i = m ∧ j = n
then 1::int else if i = j ∧ i ≠ m ∧ i ≠ n ∧ i < nrows A then 1 else 0))) (nrows A)
= nrows A"

using P_interchange_rows'_nrows_A_2[OF n m] unfolding P_interchange_rows'_def by auto
have "Rep_matrix (interchange_rows_matrix A n m) a b = (λi j. if i = n then Rep_matrix A m j else if i = m then Rep_matrix A n j else Rep_matrix A i j) a b"
unfolding interchange_rows_matrix_def interchange_rows_infmatrix_def Rep_Abs_aux_interchange2' ..
also have "... = Rep_matrix (P_interchange_rows' A n m * A) a b"
unfolding Rep_matrix_mult[of "P_interchange_rows' A n m" A] unfolding P_interchange_rows'_def
unfolding Rep_abs_aux_P_interchange_rows[OF n m]
unfolding max_eq
unfolding foldseq_finsum
proof (auto)
find_theorems "finsum"
assume n_noteq_m: "n≠m"
show "Rep_matrix A m b = (\<Oplus>\<Z>k∈{..nrows A}. (if k = m then 1 else if n = m ∧ k = n then 1 else if n = k ∧ n ≠ m ∧ n ≠ n ∧ n < nrows A then 1 else 0) * Rep_matrix A k b)"
proof -
have "finsum \<Z> (λk. (if k = m then 1 else if n = m ∧ k = n then 1 else if n = k ∧ n ≠ m ∧ n ≠ n ∧ n < nrows A then 1 else 0) * Rep_matrix A k b) {..nrows A}
= finsum \<Z> (λk. (if m = k then Rep_matrix A m b else \<zero>\<Z>)) {..nrows A}"

using Ring.abelian_monoid.finsum_cong'[of "\<Z>" "{..nrows A}" "{..nrows A}" "(λk. (if m = k then Rep_matrix A m b else \<zero>\<Z>))"
"(λk. (if k = m then 1 else if n = m ∧ k = n then 1 else if n = k ∧ n ≠ m ∧ n ≠ n ∧ n < nrows A then 1 else 0) * Rep_matrix A k b)"]
using abelian_monoid_z by auto
also have "...=Rep_matrix A m b"
using Ring.abelian_monoid.finsum_singleton[of \<Z> m "{..nrows A}"] using abelian_monoid_z m by auto
finally show ?thesis by simp
qed
next
assume " m ≠ n"
show "Rep_matrix A n b = (\<Oplus>\<Z>k∈{..nrows A}. (if k = n then 1 else if m = k ∧ m ≠ m ∧ m ≠ n ∧ m < nrows A then 1 else 0) * Rep_matrix A k b)"
proof -
have "finsum \<Z> (λk. (if k = n then 1 else if m = k ∧ m ≠ m ∧ m ≠ n ∧ m < nrows A then 1 else 0) * Rep_matrix A k b) {..nrows A}
= finsum \<Z> (λk. (if n = k then Rep_matrix A n b else \<zero>\<Z>)) {..nrows A}"
using Ring.abelian_monoid.finsum_cong'[OF abelian_monoid_z, of "{..nrows A}" "{..nrows A}"
"(λk. (if n = k then Rep_matrix A n b else \<zero>\<Z>))" "(λk. (if k = n then 1 else if m = k ∧ m ≠ m ∧ m ≠ n ∧ m < nrows A then 1 else 0) * Rep_matrix A k b)"]
by auto
also have "... = Rep_matrix A n b" using Ring.abelian_monoid.finsum_singleton[OF abelian_monoid_z, of n] using n by auto
finally show ?thesis by simp
qed
next
show "Rep_matrix A n b = (\<Oplus>\<Z>k∈{..nrows A}. (if k = n then 1 else if n = n ∧ k = n then 1 else if n = k ∧ n ≠ n ∧ n ≠ n ∧ n < nrows A then 1 else 0) * Rep_matrix A k b)"
proof -
have "finsum \<Z> (λk. (if k = n then 1 else if n = n ∧ k = n then 1 else if n = k ∧ n ≠ n ∧ n ≠ n ∧ n < nrows A then 1 else 0) * Rep_matrix A k b) {..nrows A} =
finsum \<Z> (λk. if n = k then Rep_matrix A n b else \<zero>\<Z>) {..nrows A}"

apply (rule Ring.abelian_monoid.finsum_cong') apply auto using abelian_monoid_z .
also have "... = Rep_matrix A n b" apply (rule Ring.abelian_monoid.finsum_singleton) apply auto using abelian_monoid_z n by auto
finally show ?thesis by simp
qed
next
assume a_not_m: "a≠m" and a_noteq_n: "a≠n"
show "Rep_matrix A a b = (\<Oplus>\<Z>k∈{..nrows A}. (if a = k ∧ a < nrows A then 1 else 0) * Rep_matrix A k b)"
proof (cases "a < nrows A")
case False
have "finsum \<Z> (λk. (if a = k ∧ a < nrows A then 1 else 0) * Rep_matrix A k b) {..nrows A}
= finsum \<Z> (λk. \<zero>\<Z>) {..nrows A}"

apply (rule Ring.abelian_monoid.finsum_cong') apply (rule abelian_monoid_z) apply rule apply simp
by (metis (full_types) False int_zero_eq mult_zero_left nrows_notzero zero_imp_mult_zero zmult_commute)
also have "...=\<zero>\<Z>" using Ring.abelian_monoid.finsum_zero[OF abelian_monoid_z] by auto
also have "...=0" by simp
also have "... = Rep_matrix A a b" using False nrows_notzero by auto
finally show ?thesis by simp
next
case True
have "(\<Oplus>\<Z>k∈{..nrows A}. (if a = k ∧ a < nrows A then 1 else 0) * Rep_matrix A k b) = (\<Oplus>\<Z>k∈{..nrows A}. (if a = k then 1 else 0) * Rep_matrix A k b)"
using Ring.abelian_monoid.finsum_cong'[OF abelian_monoid_z] using True by auto
also have "...=(\<Oplus>\<Z>k∈{..nrows A}. (if a = k then Rep_matrix A k b else \<zero>\<Z>))"
apply (rule abelian_monoid.finsum_cong') using abelian_monoid_z by auto
also have "...=Rep_matrix A a b" apply (rule Ring.abelian_monoid.finsum_singleton) using abelian_monoid_z True by auto
finally show ?thesis by simp
qed
qed
finally show "Rep_matrix (interchange_rows_matrix A n m) a b = Rep_matrix (P_interchange_rows' A n m * A) a b" .
qed
thus ?thesis using Rep_matrix_inject by auto
qed


lemma interchange_rows_PAQ:
assumes "n<nrows A"
and "m<nrows A"
shows "∃P Q. is_invertible P ∧ is_invertible Q ∧ interchange_rows_matrix A n m = P*A*Q"

proof (rule, rule, rule)
show "is_invertible (P_interchange_rows' A n m)" using P_interchange_rows_inverse assms unfolding is_invertible_def by auto
show "is_invertible (one_matrix (ncols A)) ∧ interchange_rows_matrix A n m = P_interchange_rows' A n m * A * (one_matrix(ncols A))"
by (metis assms(1) assms(2) is_invertible_def interchange_rows_PA ncols_mult one_matrix_inverse one_matrix_mult_right)
qed

lemma interchange_rows_PAQ':
assumes n:"n<nrows (A::int matrix)"
and m:"m<nrows A"
shows "interchange_rows_matrix A n m = (P_interchange_rows' A n m) *A*(one_matrix(ncols A))"

by (metis interchange_rows_PA m n ncols_mult one_matrix_mult_right)

lemma row_add_PA:
assumes n:"n<nrows A"
and m:"m<nrows A"
and n_not_m: "n≠m"
shows "row_add_matrix (A::int matrix) n m k = P_row_add' A n m k * A"

proof -
have "Rep_matrix (row_add_matrix A n m k) = Rep_matrix (P_row_add' A n m k * A)"
proof (rule, rule)
fix a b
have max_eq: "max (ncols (Abs_matrix (λi j. if i = n ∧ j = m then k else if i = j ∧ i < nrows A then 1 else 0))) (nrows A) = nrows A"
by (metis (no_types) P_row_add'_def P_row_add_eq P_row_add_ncols m min_max.sup.idem n n_not_m)
have "Rep_matrix (row_add_matrix A n m k) a b = (λa b. if a = n then Rep_matrix A n b + k * Rep_matrix A m b else Rep_matrix A a b) a b"
unfolding row_add_matrix_def row_add_infmatrix_def unfolding Rep_aux_row_add ..
also have "... = Rep_matrix (P_row_add' A n m k * A) a b"
unfolding P_row_add'_def Rep_matrix_mult unfolding Rep_abs_aux_P_row_add'[OF n m]
unfolding foldseq_finsum
unfolding max_eq
proof (auto)
show "Rep_matrix A n b + k * Rep_matrix A m b = (\<Oplus>\<Z>ka∈{..nrows A}. (if ka = m then k else if n = ka ∧ n < nrows A then 1 else 0) * Rep_matrix A ka b)"
proof -
def s"(λka. (if ka = m then k else if n = ka ∧ n < nrows A then 1 else 0) * Rep_matrix A ka b)"
--"Demuestro unas cuantas cosillas que nos serviran luego en el calculation"

have eq_set: "{..nrows A} = (insert m ({..nrows A}-{m}))" using m by auto
have eq_set2: "{..nrows A}-{m} = insert n ({..nrows A}-{m}-{n})"using n n_not_m by auto
have finsum_eq: "finsum \<Z> s ({..nrows A}-{m}) = s n ⊕\<Z> finsum \<Z> s ({..nrows A}-{m}-{n})"
proof -
have "finsum \<Z> s ({..nrows A}-{m}) = finsum \<Z> s (insert n ({..nrows A}-{m}-{n}))" using eq_set2 by auto
also have "... = s n ⊕\<Z> finsum \<Z> s ({..nrows A}-{m}-{n})" apply (rule Ring.abelian_monoid.finsum_insert) using abelian_monoid_z by auto
finally show ?thesis .
qed
have finsum_eq_zero: "finsum \<Z> s ({..nrows A}-{m}-{n}) = \<zero>\<Z> "
proof -
have "finsum \<Z> s ({..nrows A}-{m}-{n}) = finsum \<Z> (λk. \<zero>\<Z>) ({..nrows A}-{m}-{n})" apply (rule Ring.abelian_monoid.finsum_cong') using abelian_monoid_z s_def by auto
also have "...=\<zero>\<Z>" apply (rule Ring.abelian_monoid.finsum_zero) using abelian_monoid_z by auto
finally show ?thesis .
qed
have "finsum \<Z> s {..nrows A} = finsum \<Z> s (insert m ({..nrows A}-{m}))" using eq_set by simp
also have "...=s m ⊕\<Z> finsum \<Z> s ({..nrows A}-{m})" apply (rule Ring.abelian_monoid.finsum_insert) using abelian_monoid_z by auto
also have "... = s m ⊕\<Z> s n ⊕\<Z> finsum \<Z> s ({..nrows A}-{m}-{n})" using finsum_eq by auto
also have "... = s m ⊕\<Z> s n" using finsum_eq_zero by auto
also have "... = s n + s m" by simp
also have "... = Rep_matrix A n b + k * Rep_matrix A m b" unfolding s_def using n m n_not_m by auto
finally show ?thesis unfolding s_def by simp
qed
show "Rep_matrix A a b = (\<Oplus>\<Z>k∈{..nrows A}. (if a = k ∧ a < nrows A then 1 else 0) * Rep_matrix A k b)"
proof (cases "a<nrows A")
case False
have "finsum \<Z> (λk. (if a = k ∧ a < nrows A then 1 else 0) * Rep_matrix A k b) {..nrows A}
= finsum \<Z> (λk. \<zero>\<Z>) {..nrows A}"

apply (rule Ring.abelian_monoid.finsum_cong') apply (rule abelian_monoid_z) apply rule apply simp
by (metis (full_types) False int_zero_eq mult_zero_left nrows_notzero zero_imp_mult_zero zmult_commute)
also have "...=\<zero>\<Z>" using Ring.abelian_monoid.finsum_zero[OF abelian_monoid_z] by auto
also have "...=0" by simp
also have "... = Rep_matrix A a b" using False nrows_notzero by auto
finally show ?thesis by simp
next
case True
have "(\<Oplus>\<Z>k∈{..nrows A}. (if a = k ∧ a < nrows A then 1 else 0) * Rep_matrix A k b) = (\<Oplus>\<Z>k∈{..nrows A}. (if a = k then 1 else 0) * Rep_matrix A k b)"
using Ring.abelian_monoid.finsum_cong'[OF abelian_monoid_z] using True by auto
also have "...=(\<Oplus>\<Z>k∈{..nrows A}. (if a = k then Rep_matrix A k b else \<zero>\<Z>))"
apply (rule abelian_monoid.finsum_cong') using abelian_monoid_z by auto
also have "...=Rep_matrix A a b" apply (rule Ring.abelian_monoid.finsum_singleton) using abelian_monoid_z True by auto
finally show ?thesis by simp
qed
qed
finally show "Rep_matrix (row_add_matrix A n m k) a b = Rep_matrix (P_row_add' A n m k * A) a b" .
qed
thus ?thesis using Rep_matrix_inject by auto
qed

lemma rows_add_PAQ:
assumes "n<nrows A"
and "m<nrows A"
and "n≠m"
shows "∃P Q. is_invertible P ∧ is_invertible Q ∧ row_add_matrix A n m k= P*A*Q"

proof (rule, rule, rule)
show "is_invertible (P_row_add' A n m k)"
by (metis P_row_add_inverse assms(1) assms(2) assms(3) is_invertible_def)
next
show "is_invertible (one_matrix (ncols A)) ∧ row_add_matrix A n m k = P_row_add' A n m k * A * (one_matrix(ncols A))"
by (metis assms(1) assms(2) assms(3) is_invertible_def ncols_mult one_matrix_inverse one_matrix_mult_right row_add_PA)
qed

lemma mult_matrix_inverse:
assumes A:"is_invertible (A::int matrix)"
and B:"is_invertible B"
and n:"ncols A = nrows B" --"PARA LA MULTIPLICACION. NO SE SI SE PUEDE SACAR A PARTIR DE TEOREMAS YA HECHOS."
shows "is_invertible (A*B) ∧ nrows (A*B)=nrows A ∧ ncols(A*B)=ncols B"

proof-
from A obtain A' where inv_A: "inverse_matrix A A'" unfolding is_invertible_def by auto
from B obtain B' where inv_B: "inverse_matrix B B'" unfolding is_invertible_def by auto
show ?thesis
apply (rule conjI) prefer 2 apply (rule conjI) prefer 3 unfolding is_invertible_def
apply (rule exI[of _ "B'*A'"]) unfolding inverse_matrix_def unfolding right_inverse_matrix_def unfolding left_inverse_matrix_def
apply (rule conjI)+ prefer 3 apply (rule conjI)
proof -
show a:"A * B * (B' * A') = one_matrix (max (nrows (A * B)) (ncols (B' * A')))"
proof -
have "A * B * (B' * A') = A * (B * (B' * A'))" using mult_assoc[of A B "(B'*A')"] .
also have "...=A * ((B * B')*A')"using mult_assoc[of B "B'" "A'"] by auto
also have "...=(A*(B*B'))*A'" using mult_assoc[of A "B*B'" "A'"] by auto
also have "...=A*(one_matrix(max (nrows B) (ncols B')))*A'" using inv_B unfolding inverse_matrix_def right_inverse_matrix_def by simp
also have "...=A*A'" using one_matrix_mult_right[of A "max (nrows B) (ncols B')"] using n by auto
also have "... = one_matrix (max (nrows A) (ncols A'))" using inv_A unfolding inverse_matrix_def unfolding right_inverse_matrix_def by auto
also have "... = one_matrix (max (nrows (A * B)) (ncols (B' * A')))"
by (metis calculation inv_A inverse_matrix_def min_max.sup_absorb1 ncols_mult nrows_mult one_matrix_mult_left order_refl right_inverse_matrix_dim xt1(5))
finally show ?thesis .
qed
show "nrows (A*B) = nrows A" --"PARA ESTE ESTOY REPITIENDO GRAN PARTE DEL CALCULATION DE ANTES"
proof -
have "A * B * (B' * A') = A * (B * (B' * A'))" using mult_assoc[of A B "(B'*A')"] .
also have "...=A * ((B * B')*A')"using mult_assoc[of B "B'" "A'"] by auto
also have "...=(A*(B*B'))*A'" using mult_assoc[of A "B*B'" "A'"] by auto
also have "...=A*(one_matrix(max (nrows B) (ncols B')))*A'" using inv_B unfolding inverse_matrix_def right_inverse_matrix_def by simp
also have "...=A*A'" using one_matrix_mult_right[of A "max (nrows B) (ncols B')"] using n by auto
also have "... = one_matrix (max (nrows A) (ncols A'))" using inv_A unfolding inverse_matrix_def unfolding right_inverse_matrix_def by auto
also have "... = one_matrix (nrows A)"
by (metis inv_A inverse_matrix_def min_max.sup.idem right_inverse_matrix_dim)
finally have "nrows A = nrows (A * B * (B' * A'))" by auto
also have " ... ≤ nrows (A*B)"
by (metis nrows_mult)
also have "... ≤ nrows A" by (metis nrows_mult)
finally show ?thesis by simp
qed
show 1:"B' * A' * (A * B) = one_matrix (max (nrows (B' * A')) (ncols (A * B)))"
proof -
have "B' * A' * (A * B) = B' * (A' * (A * B))" using mult_assoc by auto
also have "... = B' * ((A' * A) * B)" using mult_assoc[of "A'" A B] by auto
also have "... = B' * (A' * A) * B" using mult_assoc[of "B'" "(A'*A)" B] by auto
also have "... = B' * B" using one_matrix_mult_right
by (metis `B' * (A' * A * B) = B' * (A' * A) * B` inv_A inverse_matrix_def le_maxI2 left_inverse_matrix_def n one_matrix_mult_left)
also have "... = one_matrix (max (nrows B') (ncols B))" using inv_B unfolding inverse_matrix_def unfolding left_inverse_matrix_def by auto
also have "... = one_matrix (max (nrows (B' * A')) (ncols (A * B)))"
by (metis `B' * (A' * (A * B)) = B' * (A' * A * B)` `B' * (A' * A * B)
= B' * (A' * A) * B`
`B' * (A' * A) * B = B' * B` `B' * A' * (A * B) = B' * (A' * (A * B))` `B' * B = one_matrix (max (nrows B') (ncols B))`
inv_B inverse_matrix_def left_inverse_matrix_dim min_max.sup_absorb1 min_max.sup_absorb2 ncols_mult ncols_one_matrix nrows_mult order_refl)

finally show ?thesis .
qed
show "ncols (A*B)=ncols B" --"LO MISMO; ESTOY REPITIENDO GRAN PARTE DEL CALCULATION DEL SHOW ANTERIOR"
proof -
have "B' * A' * (A * B) = B' * (A' * (A * B))" using mult_assoc by auto
also have "... = B' * ((A' * A) * B)" using mult_assoc[of "A'" A B] by auto
also have "... = B' * (A' * A) * B" using mult_assoc[of "B'" "(A'*A)" B] by auto
also have "... = B' * B" using one_matrix_mult_right
by (metis `B' * (A' * A * B) = B' * (A' * A) * B` inv_A inverse_matrix_def le_maxI2 left_inverse_matrix_def n one_matrix_mult_left)
also have "... = one_matrix (max (nrows B') (ncols B))" using inv_B unfolding inverse_matrix_def unfolding left_inverse_matrix_def by auto
also have "... = one_matrix (ncols B)"
by (metis inv_B inverse_matrix_def left_inverse_matrix_dim min_max.sup.idem)
finally have "ncols B = ncols (B' * A' * (A * B))" by auto
also have "... ≤ ncols (A * B)" using ncols_mult by auto
also have "... ≤ ncols B" using ncols_mult by auto
finally show ?thesis by simp
qed
have "ncols(B' * A' * (A * B)) = ncols (A*B)"
by (metis "1" le_antisym le_maxI1 min_max.le_iff_sup ncols_mult ncols_one_matrix nrows_mult nrows_one_matrix)
thus "nrows (B'*A')≤ncols (A*B)"
by (metis "1" min_max.sup.commute min_max.sup_absorb2 nat_le_linear ncols_one_matrix)
show "ncols (B' * A') ≤ nrows (A * B)"
by (metis a min_max.sup_absorb2 nat_le_linear nrows_mult nrows_one_matrix)
qed
qed

corollary mult_matrix_inverse2:
assumes A:"is_invertible (A::int matrix)"
and B:"is_invertible B"
and n:"ncols A = nrows B"
shows "is_invertible (A*B)"
using assms mult_matrix_inverse by auto

lemma Rep_abs_aux_Q_interchange_cols:
assumes "n<ncols A" and "m<ncols A"
shows "Rep_matrix (Abs_matrix (λi j. if i = n ∧ j = m then 1 else if i = m ∧ j = n then 1 else if i = j ∧ i ≠ m ∧ i ≠ n ∧ j < ncols A then 1 else 0))
=(λi j. if i = n ∧ j = m then 1 else if i = m ∧ j = n then 1 else if i = j ∧ i ≠ m ∧ i ≠ n ∧ j < ncols A then 1 else 0)"

proof (rule RepAbs_matrix)
show "∃ma. ∀j i. ma ≤ j --> (if j = n ∧ i = m then 1::'b else if j = m ∧ i = n then 1::'b else if j = i ∧ j ≠ m ∧ j ≠ n ∧ i < ncols A then 1::'b else (0::'b)) = (0::'b)"
by (metis assms(1) assms(2) leD linorder_linear order_le_less_trans)
thus " ∃na. ∀j i. na ≤ i --> (if j = n ∧ i = m then 1::'b else if j = m ∧ i = n then 1::'b else if j = i ∧ j ≠ m ∧ j ≠ n ∧ i < ncols A then 1::'b else (0::'b)) = (0::'b)"
by metis
qed




lemma one_matrix_exists_notnull_cols:
assumes "a<b"
shows "∃i. Rep_matrix ((one_matrix b)::int matrix) i a ≠ 0"

proof (rule exI[of _ a], unfold one_matrix_def)
have "Rep_matrix (Abs_matrix (λj i. if j = i ∧ j < b then 1::int else 0)) = (λj i. if j = i ∧ j < b then 1::int else 0)"
by (metis (no_types) Rep_one_matrix one_matrix_def)
hence "Rep_matrix (Abs_matrix (λj i. if j = i ∧ j < b then 1::int else 0)) a a = (λj i. if j = i ∧ j < b then 1::int else 0) a a" by auto
also have "... = 1" using assms by simp
also have "...≠ 0" by simp
finally show "Rep_matrix (Abs_matrix (λj i. if j = i ∧ j < b then 1::int else 0)) a a ≠ 0" .
qed



lemma interchange_columns_ncols:
assumes n:"n<ncols A"
and m:"m<ncols A"
and i: "∃i. Rep_matrix A i (min n m) ≠ 0"
shows "ncols (interchange_columns_matrix (A::int matrix) n m) = ncols A"

proof (cases "Suc m = ncols A")
case True
have n_le_m: "n≤m" using True n by auto
from this obtain b where not_null: "Rep_matrix A b n ≠ 0" using i by (metis min_max.inf_absorb2 min_max.inf_commute)
have "Rep_matrix (interchange_columns_matrix A n m) b m = Rep_matrix A b n" unfolding interchange_columns_matrix_def interchange_columns_infmatrix_def
unfolding Rep_Abs_aux_interchange' by simp
also have "...≠ 0" using not_null .
finally have "m < ncols (interchange_columns_matrix (A::int matrix) n m)"
by (metis ncols_notzero)
hence 1:"ncols A ≤ ncols (interchange_columns_matrix (A::int matrix) n m)" using True by simp
have 2:"ncols (interchange_columns_matrix (A::int matrix) n m) ≤ ncols A"
proof -
have "(∀j i. ncols A ≤ i --> Rep_matrix (interchange_columns_matrix A n m) j i = 0)"
unfolding interchange_columns_matrix_def interchange_columns_infmatrix_def Rep_Abs_aux_interchange'
by (metis Suc_leD True n_le_m not_less_eq_eq ncols)
thus ?thesis
using ncols_le[of "(interchange_columns_matrix (A::int matrix) n m)" "ncols A"] by simp
qed
show ?thesis using 1 and 2 by simp
next
case False note Suc_m_not_nrows_A=False
show ?thesis
proof (cases "Suc n = ncols A")
case True --"Va a ser casi repetir el caso anterior intercambiando n y m"
have m_le_n: "m≤n" using True m by simp
from this obtain b where not_null: "Rep_matrix A b m ≠ 0" using i by (metis min_max.inf_absorb2 min_max.inf_commute)
have "Rep_matrix (interchange_columns_matrix A n m) b n = Rep_matrix A b m" unfolding interchange_columns_matrix_def interchange_columns_infmatrix_def
unfolding Rep_Abs_aux_interchange' by simp
also have "...≠ 0" using not_null .
finally have "n < ncols (interchange_columns_matrix (A::int matrix) n m)" using less_ncols[of n "(interchange_columns_matrix (A::int matrix) n m)"] by auto
hence 1:"ncols A ≤ ncols (interchange_columns_matrix (A::int matrix) n m)" using True by simp
have 2:"ncols (interchange_columns_matrix (A::int matrix) n m) ≤ ncols A"
proof -
have "(∀j i. ncols A ≤ i --> Rep_matrix (interchange_columns_matrix A n m) j i = 0)"
unfolding interchange_columns_matrix_def interchange_columns_infmatrix_def Rep_Abs_aux_interchange'
by (metis Suc_leD True m_le_n not_less_eq_eq ncols)
thus ?thesis
using ncols_le[of "(interchange_columns_matrix (A::int matrix) n m)" "ncols A"] by simp
qed
show ?thesis using 1 and 2 by simp
next
case False
obtain a where a:"Suc a = ncols A" using n by (metis gr_implies_not0 nat.exhaust)
obtain c b where b:"Rep_matrix A b c ≠ 0" and "a ≤ c" using a less_ncols[of a A] by auto
have a_eq_c: "a=c"by (metis `a ≤ c` a b le_less_Suc_eq ncols_notzero)
hence b': "Rep_matrix A b a ≠ 0" using b by simp
have "Rep_matrix (interchange_columns_matrix A n m) b a = Rep_matrix A b a"
using False Suc_m_not_nrows_A a
unfolding interchange_columns_matrix_def interchange_columns_infmatrix_def
unfolding Rep_Abs_aux_interchange' by auto
also have "...≠0" using b' .
finally have "a < ncols (interchange_columns_matrix (A::int matrix) n m)" using less_ncols[of a "(interchange_columns_matrix (A::int matrix) n m)"] by auto
hence 1:"ncols A ≤ ncols (interchange_columns_matrix (A::int matrix) n m)" using a by simp
have 2:"ncols (interchange_columns_matrix (A::int matrix) n m) ≤ ncols A"
proof -
have "(∀j i. ncols A ≤ i --> Rep_matrix (interchange_columns_matrix A n m) j i = 0)"
unfolding interchange_columns_matrix_def interchange_columns_infmatrix_def Rep_Abs_aux_interchange'
by (metis less_not_refl less_ncols m n order_less_le_trans)
thus ?thesis
using ncols_le[of "(interchange_columns_matrix (A::int matrix) n m)" "ncols A"] by simp
qed
show ?thesis using 1 and 2 by simp
qed
qed



lemma interchange_columns_nrows:
assumes n:"n<ncols A"
and m:"m<ncols A"
shows "nrows (interchange_columns_matrix (A::int matrix) n m) = nrows A"

proof (cases "nrows A = 0")
case True show ?thesis
by (metis True less_ncols less_zeroE m nrows_notzero)
next
case False
from this obtain b where Suc_b:"b+1 = nrows A"
by (metis Suc_eq_plus1 not0_implies_Suc)
(*Lo mismo que antes, no sabe sacar directamente Rep_matrix A a b ≠ 0*)
obtain a c where not_null: "Rep_matrix A c a ≠ 0" and b_l_c: "b ≤ c"using Suc_b less_nrows[of b A] by auto
hence "c = b"
by (metis Suc_b Suc_eq_plus1 nrows not_less_eq_eq order_eq_iff)
hence not_null_ab:"Rep_matrix A b a ≠ 0" using not_null by simp
have "∃s. Rep_matrix (interchange_columns_matrix (A::int matrix) n m) b s ≠ 0" --"No puedo asegurar que s=a, solo en caso de que a distinto de n y a distinto de m"
proof (cases "a=n")
case True have "Rep_matrix (interchange_columns_matrix (A::int matrix) n m) b m = Rep_matrix A b a"
unfolding interchange_columns_matrix_def interchange_columns_infmatrix_def Rep_Abs_aux_interchange' True by simp
also have "...≠0" using not_null_ab .
finally show ?thesis by blast
next
case False note a_not_n = False
show ?thesis
proof (cases "a=m")
case True have "Rep_matrix (interchange_columns_matrix (A::int matrix) n m) b n = Rep_matrix A b a"
unfolding interchange_columns_matrix_def interchange_columns_infmatrix_def Rep_Abs_aux_interchange' True by simp
also have "...≠0" using not_null_ab .
finally show ?thesis by blast
next
case False
have "Rep_matrix (interchange_columns_matrix (A::int matrix) n m) b a = Rep_matrix A b a"
unfolding interchange_columns_matrix_def interchange_columns_infmatrix_def Rep_Abs_aux_interchange' using False a_not_n by simp
also have "...≠0" using not_null_ab .
finally show ?thesis by blast
qed
qed
from this obtain s where "Rep_matrix (interchange_columns_matrix (A::int matrix) n m) b s ≠ 0" by auto
hence "b < nrows (interchange_columns_matrix (A::int matrix) n m)" using less_nrows[of b "(interchange_columns_matrix (A::int matrix) n m)"] by auto
hence 1:"nrows (interchange_columns_matrix (A::int matrix) n m) ≥ nrows A" using Suc_b by auto
have 2:"nrows (interchange_columns_matrix (A::int matrix) n m) ≤ nrows A"
proof -
have "(∀j i. nrows A ≤ j --> Rep_matrix (interchange_columns_matrix A n m) j i = 0)"
unfolding interchange_columns_matrix_def interchange_columns_infmatrix_def Rep_Abs_aux_interchange'
by (metis nrows)
thus ?thesis
by (metis nrows_le)
qed
show ?thesis using 1 and 2 by simp
qed


lemma Q_interchange_columns_nrows_eq:
assumes "n<ncols A"
and "m<ncols A"
shows "nrows (Q_interchange_columns (A::int matrix) n m) = ncols (Q_interchange_columns (A::int matrix) n m)"

unfolding Q_interchange_columns_def using interchange_columns_nrows interchange_columns_ncols assms ncols_one_matrix nrows_one_matrix
by auto

thm P_interchange_row_eq

lemma Q_interchange_rows'_nrows_eq:
assumes "n<ncols A"
and "m<ncols A"
shows "nrows (Q_interchange_columns' (A::int matrix) n m) = ncols (Q_interchange_columns' (A::int matrix) n m)"

by (metis Q_interchange_column_eq Q_interchange_columns_nrows_eq assms(1) assms(2))


thm P_interchange_rows_nrows_eq2

lemma Q_interchange_rows_ncols_eq:
assumes "n<ncols A"
and "m<ncols A"
shows "ncols (Q_interchange_columns (A::int matrix) n m) = ncols A"

by (metis (no_types) Q_interchange_columns_def Q_interchange_columns_nrows_eq assms(1) assms(2) interchange_columns_nrows ncols_one_matrix nrows_one_matrix)


corollary Q_interchange_columns'_nrows:
assumes "n<ncols A"
and "m<ncols A"
shows "nrows (Q_interchange_columns' (A::int matrix) n m)=nrows (Q_interchange_columns (A::int matrix) n m)"

by (metis Q_interchange_column_eq assms(1) assms(2))

corollary Q_interchange_columns'_ncols_A:
assumes "n<ncols A"
and "m<ncols A"
shows "ncols (Q_interchange_columns' (A::int matrix) n m)=ncols A"

by (metis Q_interchange_column_eq Q_interchange_rows_ncols_eq assms(1) assms(2))



corollary Q_interchange_columns'_ncols_A_2:
assumes "n<ncols A"
and "m<ncols A"
shows "nrows (Q_interchange_columns' (A::int matrix) n m)=ncols A"

by (metis Q_interchange_column_eq Q_interchange_columns'_ncols_A Q_interchange_columns_nrows_eq assms(1) assms(2))


lemma Q_interchange_columns_inverse:
assumes n:"n<ncols A"
and m: "m<ncols (A::int matrix)"
shows "inverse_matrix (Q_interchange_columns' A n m) (Q_interchange_columns' A n m)"

unfolding inverse_matrix_def apply (rule conjI) unfolding right_inverse_matrix_def left_inverse_matrix_def
proof (auto)
show " nrows (Q_interchange_columns' A n m) ≤ ncols (Q_interchange_columns' A n m)"
by (metis Q_interchange_rows'_nrows_eq assms(1) assms(2) order_refl)
show "ncols (Q_interchange_columns' A n m) ≤ nrows (Q_interchange_columns' A n m)"
by (metis Q_interchange_rows'_nrows_eq assms(1) assms(2) order_refl)
show "Q_interchange_columns' A n m * Q_interchange_columns' A n m = one_matrix (max (nrows (Q_interchange_columns' A n m)) (ncols (Q_interchange_columns' A n m)))"
proof -
have "Rep_matrix (Q_interchange_columns' A n m * Q_interchange_columns' A n m)
= Rep_matrix (one_matrix (max (nrows (Q_interchange_columns' A n m)) (ncols (Q_interchange_columns' A n m))))"

proof (rule, rule)
fix a b
have 1:"(max (ncols (Q_interchange_columns' A n m)) (nrows (Q_interchange_columns' A n m))) = ncols (Q_interchange_columns' A n m)"
by (metis `nrows (Q_interchange_columns' A n m) ≤ ncols (Q_interchange_columns' A n m)` min_max.sup_absorb1)
have 2: "max (nrows (Q_interchange_columns' A n m)) (ncols (Q_interchange_columns' A n m)) = ncols (Q_interchange_columns' A n m)"
by (metis "1" min_max.sup.commute)
have "Rep_matrix (Q_interchange_columns' (A::int matrix) n m * Q_interchange_columns' A n m) a b = foldseq op +
(λk. Rep_matrix (Q_interchange_columns' A n m) a k * Rep_matrix (Q_interchange_columns' A n m) k b)
(max (ncols (Q_interchange_columns' A n m)) (nrows (Q_interchange_columns' A n m)))"
unfolding Rep_matrix_mult ..
also have "... = Rep_matrix (one_matrix (max (nrows (Q_interchange_columns' A n m)) (ncols (Q_interchange_columns' A n m)))) a b"
unfolding 1 2 unfolding Q_interchange_columns'_def unfolding Rep_abs_aux_Q_interchange_cols[OF n m]
proof (auto)
assume a_not_b: "a ≠ b"
show "foldseq op +
(λk. (if a = n ∧ k = m then 1::int else if a = m ∧ k = n then 1 else if a = k ∧ a ≠ m ∧ a ≠ n ∧ k < ncols A then 1 else 0) *
(if k = n ∧ b = m then 1 else if k = m ∧ b = n then 1 else if k = b ∧ k ≠ m ∧ k ≠ n ∧ b < ncols A then 1 else 0))
(ncols (Abs_matrix (λi j. if i = n ∧ j = m then 1 else if i = m ∧ j = n then 1::int else if i = j ∧ i ≠ m ∧ i ≠ n ∧ j < ncols A then 1 else 0))) = 0"

apply (rule foldseq_zero) apply simp using a_not_b by force
next
assume b_ge_cols: "¬ b < ncols (Abs_matrix (λi j. if i = n ∧ j = m then 1::int else if i = m ∧ j = n then 1 else if i = j ∧ i ≠ m ∧ i ≠ n ∧ j < ncols A then 1 else 0))"
show "foldseq op +
(λk. (if a = n ∧ k = m then 1::int else if a = m ∧ k = n then 1 else if a = k ∧ a ≠ m ∧ a ≠ n ∧ k < ncols A then 1 else 0) *
(if k = n ∧ b = m then 1 else if k = m ∧ b = n then 1 else if k = b ∧ k ≠ m ∧ k ≠ n ∧ b < ncols A then 1 else 0))
(ncols (Abs_matrix (λi j. if i = n ∧ j = m then 1 else if i = m ∧ j = n then 1::int else if i = j ∧ i ≠ m ∧ i ≠ n ∧ j < ncols A then 1 else 0))) = 0"

apply (rule foldseq_zero) apply simp using b_ge_cols using n m Q_interchange_columns'_ncols_A unfolding Q_interchange_columns'_def by force
next
assume b:" b < ncols (Abs_matrix (λi j. if i = n ∧ j = m then 1::int else if i = m ∧ j = n then 1 else if i = j ∧ i ≠ m ∧ i ≠ n ∧ j < ncols A then 1 else 0))"
def P == " (λk. (if b = n ∧ k = m then 1 else if b = m ∧ k = n then 1 else if b = k ∧ b ≠ m ∧ b ≠ n ∧ k < ncols A then 1::int else 0) *
(if k = n ∧ b = m then 1 else if k = m ∧ b = n then 1 else if k = b ∧ k ≠ m ∧ k ≠ n ∧ b < ncols A then 1::int else 0))"

have "∃k. (P k = 1 ∧ (∀i. i≠k --> P i = 0))"
proof (cases "b=n")
case True
have "P m = 1" unfolding P_def unfolding True by simp
moreover have"(∀i. i≠m --> P i = 0)" unfolding P_def unfolding 1 unfolding True by auto
ultimately show ?thesis by auto
next
case False note b_not_n = False
show ?thesis
proof (cases "b=m")
case True
have "P n = 1" unfolding P_def unfolding True by simp
moreover have"(∀i. i≠n --> P i = 0)" unfolding P_def unfolding True by auto
ultimately show ?thesis by auto
next
case False
have "P b = 1" unfolding P_def using False b_not_n b using Q_interchange_columns'_ncols_A[OF n m] unfolding Q_interchange_columns'_def by auto
moreover have"(∀i. i≠b --> P i = 0)" unfolding P_def using False b_not_n by auto
ultimately show ?thesis by auto
qed
qed
from this obtain k where P_k_1: "P k = 1" and zero: "(∀i. i≠k --> P i = 0)" by blast
have k_le_ncols: "k < ncols (Abs_matrix (λi j. if i = n ∧ j = m then 1::int else if i = m ∧ j = n then 1 else if i = j ∧ i ≠ m ∧ i ≠ n ∧ j < ncols A then 1 else 0))"
proof (rule ccontr)
assume k_g_max: "¬ k < ncols (Abs_matrix (λi j. if i = n ∧ j = m then 1::int else if i = m ∧ j = n
then 1 else if i = j ∧ i ≠ m ∧ i ≠ n ∧ j < ncols A then 1 else 0))"

hence "k ≥ ncols (Abs_matrix (λi j. if i = n ∧ j = m then 1::int else if i = m ∧ j = n
then 1 else if i = j ∧ i ≠ m ∧ i ≠ n ∧ j < ncols A then 1 else 0))"
by auto
hence "P k = 0" unfolding P_def using n m using Q_interchange_columns'_ncols_A[OF n m] unfolding Q_interchange_columns'_def by auto
thus False using P_k_1 by auto
qed
def aux=="(ncols (Abs_matrix (λi j. if i = n ∧ j = m then 1 else if i = m ∧ j = n then 1::int else if i = j ∧ i ≠ m ∧ i ≠ n ∧ j < ncols A then 1 else 0)))"
have "foldseq op + P aux = (if k ≤ aux then P k else 0)" using foldseq_almostzero [of "op +" _ P aux, OF _ _ zero] by auto
also have "... = P k" using k_le_ncols unfolding aux_def by simp
also have "...=1"using P_k_1 .
finally show "foldseq op +
(λk. (if b = n ∧ k = m then 1 else if b = m ∧ k = n then 1 else if b = k ∧ b ≠ m ∧ b ≠ n ∧ k < ncols A then 1::int else 0) *
(if k = n ∧ b = m then 1 else if k = m ∧ b = n then 1 else if k = b ∧ k ≠ m ∧ k ≠ n ∧ b < ncols A then 1::int else 0))
(ncols (Abs_matrix (λi j. if i = n ∧ j = m then 1 else if i = m ∧ j = n then 1 else if i = j ∧ i ≠ m ∧ i ≠ n ∧ j < ncols A then 1::int else 0))) = 1"

unfolding P_def aux_def by auto
qed
finally show "Rep_matrix (Q_interchange_columns' A n m * Q_interchange_columns' A n m) a b =
Rep_matrix (one_matrix (max (nrows (Q_interchange_columns' A n m)) (ncols (Q_interchange_columns' A n m)))) a b"
.
qed
thus ?thesis using Rep_matrix_inject by auto
qed
thus " Q_interchange_columns' A n m * Q_interchange_columns' A n m = one_matrix (max (nrows (Q_interchange_columns' A n m)) (ncols (Q_interchange_columns' A n m)))" .
qed


lemma interchange_columns_AQ:
assumes n:"n<ncols A"
and m: "m<ncols (A::int matrix)"
shows "interchange_columns_matrix A n m = A * Q_interchange_columns' A n m"

proof -
have "Rep_matrix (interchange_columns_matrix A n m) = Rep_matrix (A * Q_interchange_columns' A n m)"
proof (rule, rule)
fix a b
have 1:"max (ncols A) (nrows (Abs_matrix (λi j. if i = n ∧ j = m then 1::int else if i = m ∧ j = n then 1 else if i = j ∧ i ≠ m ∧ i ≠ n ∧ j < ncols A then 1 else 0)))
= ncols A"

by (metis (no_types) Q_interchange_columns'_def Q_interchange_columns'_ncols_A_2 assms(1) assms(2) min_max.sup.idem)
show "Rep_matrix (interchange_columns_matrix A n m) a b = Rep_matrix (A * Q_interchange_columns' A n m) a b"
unfolding interchange_columns_matrix_def interchange_columns_infmatrix_def unfolding Rep_Abs_aux_interchange' unfolding Rep_matrix_mult
unfolding Q_interchange_columns'_def 1 unfolding Rep_abs_aux_Q_interchange_cols[OF n m] unfolding foldseq_finsum
proof (auto)
show "Rep_matrix A a n
= (\<Oplus>\<Z>k∈{..ncols A}. Rep_matrix A a k * (if k = n then 1 else if k = m ∧ m = n then 1 else if k = m ∧ k ≠ m ∧ k ≠ n ∧ m < ncols A then 1 else 0))"

proof -
have "finsum \<Z> (λk. Rep_matrix A a k * (if k = n then 1 else if k = m ∧ m = n then 1 else if k = m ∧ k ≠ m ∧ k ≠ n ∧ m < ncols A then 1 else 0)) {..ncols A}=
finsum \<Z> (λk. if n=k then Rep_matrix A a n else \<zero>\<Z>) {..ncols A}"
apply (rule Ring.abelian_monoid.finsum_cong') using abelian_monoid_z by auto
also have "... = Rep_matrix A a n" apply (rule Ring.abelian_monoid.finsum_singleton) using n abelian_monoid_z by auto
finally show ?thesis by simp
qed
assume n_not_m: "n ≠ m"
have "finsum \<Z> (λk. Rep_matrix A a k * (if k = m then 1 else if k = n ∧ k ≠ m ∧ k ≠ n ∧ n < ncols A then 1 else 0)) {..ncols A}
= finsum \<Z> (λk. if m=k then Rep_matrix A a m else \<zero>\<Z>) {..ncols A} "
apply(rule Ring.abelian_monoid.finsum_cong'[OF abelian_monoid_z]) by auto
also have "...=Rep_matrix A a m" apply(rule Ring.abelian_monoid.finsum_singleton[OF abelian_monoid_z]) using m by auto
finally show "Rep_matrix A a m = (\<Oplus>\<Z>k∈{..ncols A}. Rep_matrix A a k * (if k = m then 1 else if k = n ∧ k ≠ m ∧ k ≠ n ∧ n < ncols A then 1 else 0))" by simp
next
assume b_not_m: "b ≠ m" and b_not_n: "b ≠ n"
show "Rep_matrix A a b = (\<Oplus>\<Z>k∈{..ncols A}. Rep_matrix A a k * (if k = b ∧ k ≠ m ∧ k ≠ n ∧ b < ncols A then 1 else 0))"
proof (cases "b < ncols A")
case False
have "finsum \<Z> (λk. Rep_matrix A a k* (if k = b ∧ k ≠ m ∧ k ≠ n ∧ b < ncols A then 1 else 0)) {..ncols A}
= finsum \<Z> (λk. \<zero>\<Z>) {..ncols A}"

apply (rule Ring.abelian_monoid.finsum_cong') apply (rule abelian_monoid_z) apply rule apply simp
by (metis (full_types) False comm_semiring_1_class.normalizing_semiring_rules(10) ring.simps(1) zmult_commute)
also have "...=\<zero>\<Z>" using Ring.abelian_monoid.finsum_zero[OF abelian_monoid_z] by auto
also have "...=0" by simp
also have "... = Rep_matrix A a b" using False ncols_notzero by auto
finally show ?thesis by simp
next
case True
have "finsum \<Z> (λk. Rep_matrix A a k * (if k = b ∧ k ≠ m ∧ k ≠ n ∧ b < ncols A then 1 else 0)) {..ncols A}
= finsum \<Z> (λk. if b = k then Rep_matrix A a b else \<zero>\<Z>) {..ncols A}"

apply (rule Ring.abelian_monoid.finsum_cong'[OF abelian_monoid_z]) using n m b_not_m b_not_n apply auto using ncols by auto
also have "...=Rep_matrix A a b" apply (rule Ring.abelian_monoid.finsum_singleton) using abelian_monoid_z True by auto
finally show ?thesis by simp
qed
qed
qed
thus ?thesis using Rep_matrix_inject by auto
qed

lemma interchange_columns_PAQ:
assumes "n<ncols A"
and "m<ncols A"
shows "∃P Q. is_invertible P ∧ is_invertible Q ∧ interchange_columns_matrix A n m = P*A*Q"

by (metis Q_interchange_columns_inverse assms(1) assms(2) is_invertible_def interchange_columns_AQ le_maxI1 one_matrix_inverse one_matrix_mult_left)


lemma interchange_columns_PAQ':
assumes n:"n<ncols (A::int matrix)"
and m:"m<ncols A"
shows "interchange_columns_matrix A n m = one_matrix (nrows A)*A*Q_interchange_columns' A n m"

using interchange_columns_AQ[OF n m] one_matrix_mult_left by auto


lemma Q_column_add_ncols:
assumes (*q:"q≠0"
and*)
n:"n<ncols A"
and m:"m<ncols A"
and n_not_m: "n≠m"
shows "ncols (Q_column_add (A::int matrix) n m q) = ncols A"

proof -
obtain a where a:"Suc a = ncols A" using n
by (metis gr_implies_not0 not0_implies_Suc)
have "Rep_matrix (Q_column_add (A::int matrix) n m q) a a ≠ 0"
proof (cases "a=n")
case True thus ?thesis unfolding Q_column_add_def column_add_matrix_def column_add_infmatrix_def Rep_aux_column_add using n m n_not_m by auto
next
case False thus ?thesis unfolding Q_column_add_def column_add_matrix_def column_add_infmatrix_def Rep_aux_column_add using n m n_not_m a by auto
qed
hence "ncols (Q_column_add A n m q) ≥ ncols A"
by (metis a not_less_eq_eq ncols)
moreover have "ncols (Q_column_add A n m q) ≤ ncols A"
proof -
have "(∀j i. ncols A ≤ i --> Rep_matrix (Q_column_add A n m q) j i = 0)"
unfolding Q_column_add_def column_add_matrix_def column_add_infmatrix_def Rep_aux_column_add using n by auto
thus ?thesis
using ncols_le[of "(Q_column_add (A::int matrix) n m q)" "ncols A"] by auto
qed
ultimately show ?thesis by auto
qed

lemma Q_column_add_nrows:
assumes (*q:"q≠0"
and*)
n:"n<ncols A"
and m:"m<ncols A"
and n_not_m: "n≠m"
shows "nrows (Q_column_add (A::int matrix) n m q) = ncols A"

proof -
obtain a where a:"Suc a = ncols A" using n
by (metis gr_implies_not0 not0_implies_Suc)
have "Rep_matrix (Q_column_add (A::int matrix) n m q) a a ≠ 0"
proof (cases "a=n")
case True thus ?thesis unfolding Q_column_add_def column_add_matrix_def column_add_infmatrix_def Rep_aux_column_add using n m n_not_m by auto
next
case False thus ?thesis unfolding Q_column_add_def column_add_matrix_def column_add_infmatrix_def Rep_aux_column_add using n m n_not_m a by auto
qed
hence "nrows (Q_column_add A n m q) ≥ ncols A"
by (metis a nrows not_less_eq_eq)
moreover have "nrows (Q_column_add A n m q) ≤ ncols A"
proof -
have "(∀j i. ncols A ≤ j --> Rep_matrix (Q_column_add A n m q) j i = 0)"
unfolding Q_column_add_def column_add_matrix_def column_add_infmatrix_def Rep_aux_column_add using n by auto
thus ?thesis
using nrows_le[of "(Q_column_add (A::int matrix) n m q)" "ncols A"] by auto
qed
ultimately show ?thesis by auto
qed

corollary Q_column_add_nrows_eq_ncols:
assumes n:"n<ncols A"
and m:"m<ncols A"
and n_not_m: "n≠m"
shows "nrows (Q_column_add (A::int matrix) n m q) = ncols (Q_column_add (A::int matrix) n m q)"

by (metis Q_column_add_ncols Q_column_add_nrows m n n_not_m)

corollary Q_column_add'_nrows_eq_ncols:
assumes n:"n<ncols A"
and m:"m<ncols A"
and n_not_m: "n≠m"
shows "nrows (Q_column_add' (A::int matrix) n m q) = ncols (Q_column_add' (A::int matrix) n m q)"

by (metis Q_column_add_eq Q_column_add_ncols Q_column_add_nrows m n n_not_m)


lemma Rep_abs_aux_Q_column_add':
assumes "n<ncols A"
and "m<ncols A"
shows "Rep_matrix (Abs_matrix (λi j. if i = m ∧ j = n then q else if i = j ∧ j < ncols A then 1 else 0))
= (λi j. if i = m ∧ j = n then q else if i = j ∧ j < ncols A then 1 else 0)"

proof (rule RepAbs_matrix)
show "∃ma. ∀j i. ma ≤ j --> (if j = m ∧ i = n then q else if j = i ∧ i < ncols A then 1::'b else (0::'b)) = (0::'b)"
by (metis assms(2) leD)
show "∃na. ∀j i. na ≤ i --> (if j = m ∧ i = n then q else if j = i ∧ i < ncols A then 1::'b else (0::'b)) = (0::'b)"
by (metis assms(1) leD)
qed

corollary Q_column_add'_ncols_eq_ncols_inverse:
assumes n:"n<ncols A"
and m:"m<ncols A"
and n_not_m: "n≠m"
shows "ncols (Q_column_add' (A::int matrix) n m q) = ncols (Q_column_add' A n m (-q))"

by (metis Q_column_add_eq Q_column_add_ncols m n n_not_m)

corollary Q_column_add'_ncols_eq_nrows_inverse:
assumes n:"n<ncols A"
and m:"m<ncols A"
and n_not_m: "n≠m"
shows "nrows (Q_column_add' (A::int matrix) n m q) = nrows (Q_column_add' A n m (-q))"

by (metis Q_column_add_eq Q_column_add_nrows m n n_not_m)



lemma Q_column_add_inverse:
assumes n:"n<ncols A"
and m:"m<ncols A"
and n_not_m: "n≠m"
shows "inverse_matrix (Q_column_add' (A::int matrix) n m q) (Q_column_add' A n m (-q))"

unfolding inverse_matrix_def apply (rule conjI) unfolding right_inverse_matrix_def left_inverse_matrix_def
proof (auto)
show "nrows (Q_column_add' A n m (- q)) ≤ ncols (Q_column_add' A n m q)"
by (metis Q_column_add'_ncols_eq_nrows_inverse Q_column_add'_nrows_eq_ncols m n n_not_m order_refl)
show "ncols (Q_column_add' A n m (- q)) ≤ nrows (Q_column_add' A n m q)"
by (metis Q_column_add'_ncols_eq_nrows_inverse Q_column_add'_nrows_eq_ncols m n n_not_m order_refl)
show "Q_column_add' A n m q * Q_column_add' A n m (- q) = one_matrix (max (nrows (Q_column_add' A n m q)) (ncols (Q_column_add' A n m (- q))))"
proof -
have" Rep_matrix (Q_column_add' A n m q * Q_column_add' A n m (- q)) = Rep_matrix (one_matrix (max (nrows (Q_column_add' A n m q)) (ncols (Q_column_add' A n m (- q)))))"
proof (rule,rule)
fix a b
have 1:"max (nrows (Q_column_add' A n m q)) (ncols (Q_column_add' A n m (- q))) = ncols (Q_column_add' A n m (- q))"
by (metis Q_column_add'_ncols_eq_nrows_inverse Q_column_add'_nrows_eq_ncols m min_max.sup.idem n n_not_m)
have 2: "max (ncols (Q_column_add' A n m q)) (nrows (Q_column_add' A n m (- q))) = ncols (Q_column_add' A n m (- q))"
by (metis Q_column_add'_ncols_eq_ncols_inverse `nrows (Q_column_add' A n m (- q)) ≤ ncols (Q_column_add' A n m q)` m min_max.sup_absorb1 n n_not_m)
have 3: "ncols (Abs_matrix (λi j. if i = m ∧ j = n then - q else if i = j ∧ j < ncols A then 1 else 0)) = ncols A"
by (metis (no_types) Q_column_add'_def Q_column_add'_nrows_eq_ncols Q_column_add_eq Q_column_add_nrows equation_minus_iff m n n_not_m zminus_zminus)
have "Rep_matrix (Q_column_add' A n m q * Q_column_add' A n m (- q)) a b = foldseq op + (λk. Rep_matrix (Q_column_add' A n m q) a k
* Rep_matrix (Q_column_add' A n m (- q)) k b) (max (ncols (Q_column_add' A n m q)) (nrows (Q_column_add' A n m (- q))))"
unfolding Rep_matrix_mult ..
also have"...=Rep_matrix (one_matrix (max (nrows (Q_column_add' A n m q)) (ncols (Q_column_add' A n m (- q))))) a b"
unfolding Rep_one_matrix 1 2
unfolding Q_column_add'_def
unfolding Rep_abs_aux_Q_column_add'[OF n m]
proof (auto, unfold 3) --"Tengo los mismos tres casos que en la demostracion con el intercambiar filas"
assume b:"¬ b < ncols A"
show "foldseq op + (λk. (if a = m ∧ k = n then q else if a = k ∧ k < ncols A then 1 else 0) * (if k = m ∧ b = n then - q else if k = b ∧ b < ncols A then 1 else 0))
(ncols A) = 0"

apply (rule foldseq_zero)
apply simp
using b
using n m by auto
next
assume b: "b < ncols A"

def P "(λk. (if b = m ∧ k = n then q else if b = k ∧ k < ncols A then 1 else 0) * (if k = m ∧ b = n then - q else if k = b ∧ b < ncols A then 1 else 0))"
have "∃k. (P k = 1 ∧ (∀i. i≠k --> P i = 0))"
proof (rule exI[of _ "b"], unfold P_def, rule conjI)
show "(if b = m ∧ b = n then q else if b = b ∧ b < ncols A then 1 else 0) * (if b = m ∧ b = n then - q else if b = b ∧ b < ncols A then 1 else 0) = 1"
using n_not_m b by auto
show "∀i. i ≠ b --> (if b = m ∧ i = n then q else if b = i ∧ i < ncols A then 1 else 0) * (if i = m ∧ b = n then - q else if i = b ∧ b < ncols A then 1 else 0) = 0"
using n_not_m b by auto
qed
from this obtain k where P_k_1: "P k = 1" and zero: "(∀i. i≠k --> P i = 0)" by blast
have k_le_ncols: "k < ncols A"
proof (rule ccontr)
assume "¬ k < ncols A"
hence "k≥ncols A" by auto
hence "P k = 0" unfolding P_def using m ncols by auto
thus False using P_k_1 by simp
qed
have "foldseq op + P (ncols A) = (if k ≤ ncols A then P k else 0)" using foldseq_almostzero [of "op +" _ P "ncols A", OF _ _ zero] by auto
also have "... = P k" using k_le_ncols by simp
also have "...=1"using P_k_1 .
finally show "foldseq op + (λk. (if b = m ∧ k = n then q else if b = k ∧ k < ncols A then 1 else 0)
* (if k = m ∧ b = n then - q else if k = b ∧ b < ncols A then 1 else 0)) (ncols A) = 1"

unfolding P_def .
next
assume a_not_b: "a ≠ b"
show "foldseq op + (λk. (if a = m ∧ k = n then q else if a = k ∧ k < ncols A then 1 else 0)
* (if k = m ∧ b = n then - q else if k = b ∧ b < ncols A then 1 else 0)) (ncols A) = 0"

proof (cases "a=m ∧ b=n")
case False
show ?thesis
apply (rule foldseq_zero)
apply simp
using a_not_b n m n_not_m False by simp
next
case True
def P "(λk. (if a = m ∧ k = n then q else if a = k ∧ k < ncols A then 1 else 0)
* (if k = m ∧ b = n then - q else if k = b ∧ b < ncols A then 1 else 0))"

have "foldseq op + P (ncols A) = P a + P b"
proof (rule foldseq_almostzero_2)
show "∀i. i ≠ a ∧ i ≠ b --> P i = 0" unfolding P_def using True by auto
show "a ≤ ncols A" using True m by auto
show "b ≤ ncols A" using True n by auto
show "a ≠ b" using True n_not_m by simp
qed
also have "... = 0" unfolding P_def using True using n m n_not_m by auto
finally show ?thesis unfolding P_def .
qed
qed
finally show "Rep_matrix (Q_column_add' A n m q * Q_column_add' A n m (- q)) a b
= Rep_matrix (one_matrix (max (nrows (Q_column_add' A n m q)) (ncols (Q_column_add' A n m (- q))))) a b"
.
qed
thus ?thesis using Rep_matrix_inject by blast
qed
next
--"EL SIGUIENTE SHOW ES LA PARTE DE LA MULTIPLICACION A IZQUIERDA DE LA INVERSA, LA DEMOSTRACION ES MUY PARECIDA A LA ANTERIOR."

show "Q_column_add' A n m (- q) * Q_column_add' A n m q = one_matrix (max (nrows (Q_column_add' A n m (- q))) (ncols (Q_column_add' A n m q)))"
proof -
have" Rep_matrix (Q_column_add' A n m (-q) * Q_column_add' A n m q) = Rep_matrix (one_matrix (max (nrows (Q_column_add' A n m (-q))) (ncols (Q_column_add' A n m q))))"
proof (rule,rule)
fix a b
have 1:"max (nrows (Q_column_add' A n m (-q))) (ncols (Q_column_add' A n m q)) = ncols (Q_column_add' A n m q)"
by (metis Q_column_add'_ncols_eq_nrows_inverse Q_column_add'_nrows_eq_ncols m min_max.sup.idem n n_not_m)
have 2: "max (ncols (Q_column_add' A n m (-q))) (nrows (Q_column_add' A n m q)) = ncols (Q_column_add' A n m q)"
by (metis Max.idem Q_column_add_eq Q_column_add_ncols Q_column_add_nrows m n n_not_m)
have 3: "ncols (Abs_matrix (λi j. if i = m ∧ j = n then q else if i = j ∧ j < ncols A then 1 else 0)) = ncols A"
by (metis (no_types) Q_column_add'_def Q_column_add'_nrows_eq_ncols Q_column_add_eq Q_column_add_nrows m n n_not_m)
have "Rep_matrix (Q_column_add' A n m (-q) * Q_column_add' A n m q) a b = foldseq op + (λk. Rep_matrix (Q_column_add' A n m (-q)) a k
* Rep_matrix (Q_column_add' A n m q) k b) (max (ncols (Q_column_add' A n m (-q))) (nrows (Q_column_add' A n m q)))"
unfolding Rep_matrix_mult ..
also have"...=Rep_matrix (one_matrix (max (nrows (Q_column_add' A n m (-q))) (ncols (Q_column_add' A n m q)))) a b"
unfolding Rep_one_matrix 1 2
unfolding Q_column_add'_def
unfolding Rep_abs_aux_Q_column_add'[OF n m]
proof (auto, unfold 3)
assume b:"¬ b < ncols A"
show "foldseq op + (λk. (if a = m ∧ k = n then (-q) else if a = k ∧ k < ncols A then 1 else 0) * (if k = m ∧ b = n then q else if k = b ∧ b < ncols A then 1 else 0))
(ncols A) = 0"

apply (rule foldseq_zero)
apply simp
using b
using n m by auto
next
assume b: "b < ncols A"
def P "(λk. (if b = m ∧ k = n then (-q) else if b = k ∧ k < ncols A then 1 else 0) * (if k = m ∧ b = n then q else if k = b ∧ b < ncols A then 1 else 0))"
have "∃k. (P k = 1 ∧ (∀i. i≠k --> P i = 0))"
proof (rule exI[of _ "b"], unfold P_def, rule conjI)
show "(if b = m ∧ b = n then -q else if b = b ∧ b < ncols A then 1 else 0) * (if b = m ∧ b = n then q else if b = b ∧ b < ncols A then 1 else 0) = 1"
using n_not_m b by auto
show "∀i. i ≠ b --> (if b = m ∧ i = n then -q else if b = i ∧ i < ncols A then 1 else 0) * (if i = m ∧ b = n then q else if i = b ∧ b < ncols A then 1 else 0) = 0"
using n_not_m b by auto
qed
from this obtain k where P_k_1: "P k = 1" and zero: "(∀i. i≠k --> P i = 0)" by blast
have k_le_ncols: "k < ncols A"
proof (rule ccontr)
assume "¬ k < ncols A"
hence "k≥ncols A" by auto
hence "P k = 0" unfolding P_def using m ncols by auto
thus False using P_k_1 by simp
qed
have "foldseq op + P (ncols A) = (if k ≤ ncols A then P k else 0)" using foldseq_almostzero [of "op +" _ P "ncols A", OF _ _ zero] by auto
also have "... = P k" using k_le_ncols by simp
also have "...=1"using P_k_1 .
finally show "foldseq op + (λk. (if b = m ∧ k = n then -q else if b = k ∧ k < ncols A then 1 else 0)
* (if k = m ∧ b = n then q else if k = b ∧ b < ncols A then 1 else 0)) (ncols A) = 1"

unfolding P_def .
next
assume a_not_b: "a ≠ b"
show "foldseq op + (λk. (if a = m ∧ k = n then -q else if a = k ∧ k < ncols A then 1 else 0)
* (if k = m ∧ b = n then q else if k = b ∧ b < ncols A then 1 else 0)) (ncols A) = 0"

proof (cases "a=m ∧ b=n")
case False
show ?thesis
apply (rule foldseq_zero)
apply simp
using a_not_b n m n_not_m False by simp
next
case True
def P "(λk. (if a = m ∧ k = n then -q else if a = k ∧ k < ncols A then 1 else 0)
* (if k = m ∧ b = n then q else if k = b ∧ b < ncols A then 1 else 0))"

have "foldseq op + P (ncols A) = P a + P b"
proof (rule foldseq_almostzero_2)
show "∀i. i ≠ a ∧ i ≠ b --> P i = 0" unfolding P_def using True by auto
show "a ≤ ncols A" using True m by auto
show "b ≤ ncols A" using True n by auto
show "a ≠ b" using True n_not_m by simp
qed
also have "... = 0" unfolding P_def using True using n m n_not_m by auto
finally show ?thesis unfolding P_def .
qed
qed
finally show "Rep_matrix (Q_column_add' A n m (-q) * Q_column_add' A n m q) a b
= Rep_matrix (one_matrix (max (nrows (Q_column_add' A n m (-q))) (ncols (Q_column_add' A n m q)))) a b"
.
qed
thus ?thesis using Rep_matrix_inject by blast
qed
qed



lemma column_add_AQ:
assumes n:"n<ncols A"
and m:"m<ncols A"
and n_not_m: "n≠m"
shows "column_add_matrix (A::int matrix) n m k = A * Q_column_add' A n m k"

proof -
have "Rep_matrix (column_add_matrix A n m k) = Rep_matrix (A* Q_column_add' A n m k)"
proof (rule, rule)
fix a b
have max_eq: "max (ncols A) (nrows (Abs_matrix (λi j. if i = m ∧ j = n then k else if i = j ∧ j < ncols A then 1 else 0))) = ncols A"
by (metis (no_types) Q_column_add'_def Q_column_add_eq Q_column_add_nrows m min_max.sup.idem n n_not_m)
have "Rep_matrix (column_add_matrix A n m k) a b = (if b = n then Rep_matrix A a n + k * Rep_matrix A a m else Rep_matrix A a b)"
unfolding column_add_matrix_def column_add_infmatrix_def unfolding Rep_aux_column_add ..
also have "... = Rep_matrix (A * Q_column_add' A n m k) a b"
unfolding Q_column_add'_def Rep_matrix_mult unfolding Rep_abs_aux_Q_column_add'[OF n m]
unfolding foldseq_finsum
unfolding max_eq
proof (auto)
show " Rep_matrix A a n + k * Rep_matrix A a m = (\<Oplus>\<Z>ka∈{..ncols A}. Rep_matrix A a ka * (if ka = m then k else if ka = n ∧ n < ncols A then 1 else 0))"
proof -
def s"(λka. Rep_matrix A a ka * (if ka = m then k else if ka = n ∧ n < ncols A then 1 else 0))"
--"Demuestro unas cuantas cosillas que nos serviran luego en el calculation"

have eq_set: "{..ncols A} = (insert m ({..ncols A}-{m}))" using m by auto
have eq_set2: "{..ncols A}-{m} = insert n ({..ncols A}-{m}-{n})"using n n_not_m by auto
have finsum_eq: "finsum \<Z> s ({..ncols A}-{m}) = s n ⊕\<Z> finsum \<Z> s ({..ncols A}-{m}-{n})"
proof -
have "finsum \<Z> s ({..ncols A}-{m}) = finsum \<Z> s (insert n ({..ncols A}-{m}-{n}))" using eq_set2 by auto
also have "... = s n ⊕\<Z> finsum \<Z> s ({..ncols A}-{m}-{n})" apply (rule Ring.abelian_monoid.finsum_insert) using abelian_monoid_z by auto
finally show ?thesis .
qed
have finsum_eq_zero: "finsum \<Z> s ({..ncols A}-{m}-{n}) = \<zero>\<Z> "
proof -
have "finsum \<Z> s ({..ncols A}-{m}-{n}) = finsum \<Z> (λk. \<zero>\<Z>) ({..ncols A}-{m}-{n})" apply (rule Ring.abelian_monoid.finsum_cong') using abelian_monoid_z s_def by auto
also have "...=\<zero>\<Z>" apply (rule Ring.abelian_monoid.finsum_zero) using abelian_monoid_z by auto
finally show ?thesis .
qed
have "finsum \<Z> s {..ncols A} = finsum \<Z> s (insert m ({..ncols A}-{m}))" using eq_set by simp
also have "...=s m ⊕\<Z> finsum \<Z> s ({..ncols A}-{m})" apply (rule Ring.abelian_monoid.finsum_insert) using abelian_monoid_z by auto
also have "... = s m ⊕\<Z> s n ⊕\<Z> finsum \<Z> s ({..ncols A}-{m}-{n})" using finsum_eq by auto
also have "... = s m ⊕\<Z> s n" using finsum_eq_zero by auto
also have "... = s n + s m" by simp
also have "... = Rep_matrix A a n + k * Rep_matrix A a m" unfolding s_def using n m n_not_m by simp
finally show ?thesis unfolding s_def by simp
qed
assume b_not_n: "b ≠ n"
show "Rep_matrix A a b = (\<Oplus>\<Z>k∈{..ncols A}. Rep_matrix A a k * (if k = b ∧ b < ncols A then 1 else 0))"
proof (cases "b<ncols A")
case False
have "finsum \<Z> (λk. Rep_matrix A a k * (if k = b ∧ b < ncols A then 1 else 0)) {..ncols A}
= finsum \<Z> (λk. \<zero>\<Z>) {..ncols A}"

apply (rule Ring.abelian_monoid.finsum_cong') apply (rule abelian_monoid_z) apply rule apply simp
by (metis (full_types) False comm_semiring_1_class.normalizing_semiring_rules(10) ring.simps(1))
also have "...=\<zero>\<Z>" using Ring.abelian_monoid.finsum_zero[OF abelian_monoid_z] by auto
also have "...=0" by simp
also have "... = Rep_matrix A a b" using False ncols_notzero by auto
finally show ?thesis by simp
next
case True
have "(\<Oplus>\<Z>k∈{..ncols A}. Rep_matrix A a k * (if k = b ∧ b < ncols A then 1 else 0)) = (\<Oplus>\<Z>k∈{..ncols A}. Rep_matrix A a k * (if k = b then 1 else 0))"
using Ring.abelian_monoid.finsum_cong'[OF abelian_monoid_z] using True by auto
also have "...=(\<Oplus>\<Z>k∈{..ncols A}. (if b = k then Rep_matrix A a k else \<zero>\<Z>))"
apply (rule abelian_monoid.finsum_cong') using abelian_monoid_z by auto
also have "...=Rep_matrix A a b" apply (rule Ring.abelian_monoid.finsum_singleton) using abelian_monoid_z True by auto
finally show ?thesis by simp
qed
qed
finally show "Rep_matrix (column_add_matrix A n m k) a b = Rep_matrix (A * Q_column_add' A n m k) a b" .
qed
thus ?thesis using Rep_matrix_inject by auto
qed

lemma column_add_PAQ:
assumes "n<ncols A"
and "m<ncols A"
and "n≠m"
shows "∃P Q. is_invertible P ∧ is_invertible Q ∧ column_add_matrix A n m k= P*A*Q"

apply (rule, rule, rule) prefer 2 proof (rule)
show "is_invertible (Q_column_add' A n m k)"
by (metis Q_column_add_inverse assms(1) assms(2) assms(3) is_invertible_def)
show "is_invertible (one_matrix (nrows A))"
by (metis is_invertible_def one_matrix_inverse)
show "column_add_matrix A n m k = (one_matrix(nrows A)) * A * Q_column_add' A n m k"
by (metis assms(1) assms(2) assms(3) column_add_AQ one_matrix_mult_left order_refl)
qed

(* PROBLEM, WE CAN'T DEMONSTRATE IT. IT IS FALSE!
lemma
assumes i:"i<nrows (A::int matrix)" and k:"k<nrows A" and j:"j<ncols A" and l:"l<ncols A"
shows "∃P Q. is_invertible P ∧ is_invertible Q ∧ move_element A i j k l = P*A*Q" (*PENSAR EN EL NROWS Y NCOLS QUE DEBERIAN TENER P y Q*)
proof -
def A1 == "(P_interchange_rows' A i k * A * (one_matrix (ncols A)))"
(*have 1: "nrows A1 = nrows A" unfolding A1_def sorry*) (*ESTO ES FALSO SERA CIERTO SI A ES INVERTIBLEEEE!*)
show ?thesis
proof (rule exI[of _ "(one_matrix (nrows (A1))) *(P_interchange_rows' A i k)"], rule exI[of _ "(one_matrix (ncols A))* Q_interchange_columns' A1 j l"], auto)
show "is_invertible (one_matrix (nrows A1) * P_interchange_rows' A i k)"
by (metis P_interchange_rows'_nrows_A P_interchange_rows_inverse is_invertible_def i k one_matrix_mult_left order_refl)
show "is_invertible (one_matrix (ncols A) * Q_interchange_columns' A1 j l)"
proof -
have 2: "one_matrix (ncols A) * Q_interchange_columns' A1 j l = Q_interchange_columns' A1 j l"
by (metis A1_def Q_interchange_columns'_ncols_A_2 i interchange_rows_PAQ' interchange_rows_ncols j k l le_refl one_matrix_mult_left)
have "is_invertible (Q_interchange_columns' A1 j l)"
by (metis A1_def Q_interchange_columns_inverse is_invertible_def i interchange_rows_PAQ' interchange_rows_ncols j k l)
thus ?thesis unfolding 2 .
qed

show " move_element A i j k l = one_matrix (nrows A) * P_interchange_rows' A i k * A * (one_matrix (ncols A) * Q_interchange_columns' A1 j l)"
proof -
have "move_element A i j k l = interchange_columns_matrix (interchange_rows_matrix A i k) j l" unfolding move_element_def ..
also have "... = interchange_columns_matrix A1 j l" unfolding interchange_rows_PAQ'[OF i k] unfolding A1_def ..
also have "... = one_matrix (nrows A1) * A1 * Q_interchange_columns' A1 j l"
proof (rule interchange_columns_PAQ')
show "j < ncols A1"
by (metis A1_def i interchange_rows_PAQ' interchange_rows_ncols j k)
show "l < ncols A1"
by (metis A1_def i interchange_rows_PAQ' interchange_rows_ncols k l)
qed
also have "... = one_matrix (nrows A1) * (P_interchange_rows' A i k * A * one_matrix (ncols A)) * Q_interchange_columns' A1 j l" unfolding A1_def ..
also have "... = one_matrix (nrows A) * (P_interchange_rows' A i k * A * one_matrix (ncols A)) * Q_interchange_columns' A1 j l" unfolding 1 ..
finally show ?thesis
by (metis "1" A1_def P_interchange_rows'_nrows_A Q_interchange_columns'_ncols_A_2 i
interchange_rows_PAQ' interchange_rows_ncols j k l le_refl ncols_mult one_matrix_mult_left one_matrix_mult_right)
qed

have "ncols (one_matrix (nrows A)) = nrows (P_interchange_rows A i j)"
proof -
have "ncols (one_matrix (nrows A)) = nrows A" using ncols_one_matrix[of "nrows A"] apply auto


obtain P1 and Q1 where P1:"is_invertible (P1)" and Q1:"is_invertible (Q1)" and eq1:"(interchange_rows_matrix A i k) = P1 * A * Q1" using interchange_rows_PAQ[OF i k] by blast
obtain P2 and Q2 where P2:"is_invertible (P2)" and Q2:"is_invertible (Q2)" and eq2:"interchange_columns_matrix (interchange_rows_matrix A i k) j l = P2 * (P1 * A * Q1)*Q2"
using interchange_columns_PAQ
by (metis `interchange_rows_matrix A i k = P1 * A * Q1` i interchange_rows_ncols j k l)
show ?thesis
proof (rule exI[of _ "(P2*P1)"], rule exI[of _ "(Q1*Q2)"], auto)
show "is_invertible (P2 * P1)" using mult_matrix_inverse[OF P2 P1] apply auto
proof (rule mult_matrix_inverse)
qed
*)




(*We have to define new P and Q matrices keeping the dimension (based on the nrows and ncols of the original matrix*)

definition P_interchange_rows_new :: "nat => nat => nat => int matrix"
where "P_interchange_rows_new a n m == interchange_rows_matrix (one_matrix a) n m"



definition P_interchange_rows'_new :: "nat => nat => nat => int matrix"
where "P_interchange_rows'_new a n m ≡ Abs_matrix(λi j. if (i=n ∧ j=m) then 1 else (if (i=m ∧ j=n) then 1 else (if (i=j ∧ i≠m ∧ i≠n ∧ i < a) then 1 else 0)))"


lemma P_interchange_row_new_eq:
assumes n:"n < a"
and m:"m < a"
shows "P_interchange_rows'_new a n m = P_interchange_rows_new a n m"

proof -
have 1:"nrows ((one_matrix a)::int matrix) = a"
using nrows_one_matrix[of a] by auto
show ?thesis
using P_interchange_row_eq[of n "(one_matrix a)" m]
unfolding 1 using n m unfolding P_interchange_rows_new_def P_interchange_rows'_new_def P_interchange_rows_def P_interchange_rows'_def unfolding 1 by auto
qed

lemma P_interchange_rows_new_ncols:
assumes "n<a" and "m<a"
shows "ncols (P_interchange_rows'_new a n m) = a"

by (metis P_interchange_row_new_eq P_interchange_rows_new_def assms(1) assms(2) interchange_rows_ncols ncols_one_matrix nrows_one_matrix)

lemma P_interchange_rows_new_nrows:
assumes "n<a" and "m<a"
shows "nrows (P_interchange_rows'_new a n m) = a"

by (metis P_interchange_row_new_eq P_interchange_rows_def P_interchange_rows_new_def P_interchange_rows_nrows_eq2 assms(1) assms(2) nrows_one_matrix)



lemma P_interchange_rows_new_inverse:
assumes n:"n<a"
and m:"m<a"
shows "inverse_matrix (P_interchange_rows'_new a n m) (P_interchange_rows'_new a n m)"

proof -
have 1:"nrows ((one_matrix a)::int matrix) = a"
using nrows_one_matrix[of a] by auto
show ?thesis using P_interchange_rows_inverse[of n "one_matrix a" m] unfolding P_interchange_rows'_def unfolding 1 using n m unfolding P_interchange_rows'_new_def by auto
qed


lemma Rep_Abs_P_interchange_rows_new:
assumes n:"n<z" and m:"m<z"
shows " Rep_matrix (Abs_matrix (λi j. if i = n ∧ j = m then 1::int else if i = m ∧ j = n then 1 else if i = j ∧ i ≠ m ∧ i ≠ n ∧ i < z then 1 else 0))
= (λi j. if i = n ∧ j = m then 1 else if i = m ∧ j = n then 1::int else if i = j ∧ i ≠ m ∧ i ≠ n ∧ i < z then 1 else 0)"

proof (rule RepAbs_matrix)
show " ∃ma. ∀j i. ma ≤ j --> (if j = n ∧ i = m then 1 else if j = m ∧ i = n then 1 else if j = i ∧ j ≠ m ∧ j ≠ n ∧ j < z then 1 else 0) = 0"
by (metis leD m n)
show "∃na. ∀j i. na ≤ i --> (if j = n ∧ i = m then 1 else if j = m ∧ i = n then 1 else if j = i ∧ j ≠ m ∧ j ≠ n ∧ j < z then 1 else 0) = 0"
by (metis linorder_not_le m n)
qed

lemma interchange_rows_PA_new:
assumes n:"n<nrows A"
and m: "m<nrows A"
and na: "nrows A ≤ z"
shows "interchange_rows_matrix (A::int matrix) n m = P_interchange_rows'_new z n m * A"

proof -
have "Rep_matrix (interchange_rows_matrix (A::int matrix) n m) = Rep_matrix (P_interchange_rows'_new z n m * A)"
proof (rule, rule)
fix a b
have aux: "max (ncols (Abs_matrix (λi j. if i = n ∧ j = m then 1::int else if i = m ∧ j = n then 1 else if i = j ∧ i ≠ m ∧ i ≠ n ∧ i < z then 1 else 0))) (nrows A) = z"
by (metis P_interchange_rows'_new_def P_interchange_rows_new_ncols linorder_not_less m min_max.sup.commute min_max.sup_absorb1 n na not_less_iff_gr_or_eq xt1(5) xt1(6))
have nz:"n<z" and mz:"m<z" using n m na by auto
show "Rep_matrix (interchange_rows_matrix A n m) a b = Rep_matrix (P_interchange_rows'_new z n m * A) a b"
unfolding interchange_rows_matrix_def interchange_rows_infmatrix_def
P_interchange_rows'_new_def
unfolding Rep_matrix_mult unfolding Rep_Abs_aux_interchange2'[of n A m] unfolding Rep_Abs_P_interchange_rows_new[OF nz mz]
unfolding foldseq_finsum unfolding aux
proof auto
show "Rep_matrix A n b = (\<Oplus>\<Z>k∈{..z}. (if k = n then 1 else if n = n ∧ k = n then 1 else if n = k ∧ n ≠ n ∧ n ≠ n ∧ n < z then 1 else 0) * Rep_matrix A k b)"
proof -
have "finsum \<Z> (λk. (if k = n then 1 else if n = n ∧ k = n then 1 else if n = k ∧ n ≠ n ∧ n ≠ n ∧ n < z then 1 else 0) * Rep_matrix A k b) {..z}=
finsum \<Z> (λk. if n=k then Rep_matrix A n b else \<zero>\<Z>) {..z}"
apply (rule Ring.abelian_monoid.finsum_cong') using abelian_monoid_z by auto
also have "... = Rep_matrix A n b" apply (rule Ring.abelian_monoid.finsum_singleton) using nz using abelian_monoid_z by auto
finally show ?thesis by simp
qed
show "Rep_matrix A n b = (\<Oplus>\<Z>k∈{..z}. (if k = n then 1 else if m = k ∧ m ≠ m ∧ m ≠ n ∧ m < z then 1 else 0) * Rep_matrix A k b)"
proof -
have "finsum \<Z> (λk. (if k = n then 1 else if m = k ∧ m ≠ m ∧ m ≠ n ∧ m < z then 1 else 0) * Rep_matrix A k b) {..z}
= finsum \<Z> (λk. if n=k then Rep_matrix A n b else \<zero>\<Z>) {..z}"
apply (rule Ring.abelian_monoid.finsum_cong') using abelian_monoid_z by auto
also have "... = Rep_matrix A n b" apply (rule Ring.abelian_monoid.finsum_singleton) using nz using abelian_monoid_z by auto
finally show ?thesis by simp
qed
show " Rep_matrix A m b = (\<Oplus>\<Z>k∈{..z}. (if k = m then 1 else if n = m ∧ k = n then 1 else if n = k ∧ n ≠ m ∧ n ≠ n ∧ n < z then 1 else 0) * Rep_matrix A k b)"
proof -
have "finsum \<Z> (λk. (if k = m then 1 else if n = m ∧ k = n then 1 else if n = k ∧ n ≠ m ∧ n ≠ n ∧ n < z then 1 else 0) * Rep_matrix A k b) {..z}
= finsum \<Z> (λk. if m=k then Rep_matrix A m b else \<zero>\<Z>) {..z}"
apply (rule Ring.abelian_monoid.finsum_cong') using abelian_monoid_z by auto
also have "... = Rep_matrix A m b" apply (rule Ring.abelian_monoid.finsum_singleton) using mz using abelian_monoid_z by auto
finally show ?thesis by simp
qed
assume a_not_m: "a ≠ m" and a_not_n: "a ≠ n"
show " Rep_matrix A a b = (\<Oplus>\<Z>k∈{..z}. (if a = k ∧ a < z then 1 else 0) * Rep_matrix A k b)"
proof (cases "a < z")
case False
have "finsum \<Z> (λk. (if a = k ∧ a < z then 1 else 0) * Rep_matrix A k b) {..z}
= finsum \<Z> (λk. \<zero>\<Z>) {..z}"

apply (rule Ring.abelian_monoid.finsum_cong') apply (rule abelian_monoid_z) apply rule apply simp
by (metis (full_types) False int_zero_eq mult_eq_0_iff)
also have "...=\<zero>\<Z>" using Ring.abelian_monoid.finsum_zero[OF abelian_monoid_z] by auto
also have "...=0" by simp
also have "... = Rep_matrix A a b" by (metis False linorder_not_less na nrows_le)
finally show ?thesis by simp
next
case True
have "(\<Oplus>\<Z>k∈{..z}. (if a = k ∧ a < z then 1 else 0) * Rep_matrix A k b) = (\<Oplus>\<Z>k∈{..z}. (if a = k then 1 else 0) * Rep_matrix A k b)"
using Ring.abelian_monoid.finsum_cong'[OF abelian_monoid_z] using True by auto
also have "...=(\<Oplus>\<Z>k∈{..z}. (if a = k then Rep_matrix A k b else \<zero>\<Z>))"
apply (rule abelian_monoid.finsum_cong') using abelian_monoid_z by auto
also have "...=Rep_matrix A a b" apply (rule Ring.abelian_monoid.finsum_singleton) using abelian_monoid_z True by auto
finally show ?thesis by simp
qed
qed
qed
thus ?thesis using Rep_matrix_inject by auto
qed





definition Q_interchange_columns_new :: "nat => nat => nat => int matrix"
where "Q_interchange_columns_new a n m == interchange_columns_matrix (one_matrix a) n m"


definition Q_interchange_columns'_new :: "nat => nat => nat => int matrix"
where "Q_interchange_columns'_new a n m ≡ Abs_matrix(λi j. if (i=n ∧ j=m) then 1 else (if (i=m ∧ j=n) then 1 else (if (i=j ∧ i≠m ∧ i≠n ∧ j < a) then 1 else 0)))"


lemma Q_interchange_columns_new_eq:
assumes n:"n < a"
and m:"m < a"
shows "Q_interchange_columns'_new a n m = Q_interchange_columns_new a n m"

proof -
have 1: "ncols ((one_matrix a)::int matrix) = a" using ncols_one_matrix[of a] by auto
show ?thesis using Q_interchange_column_eq[of n "one_matrix a" m]
unfolding Q_interchange_columns'_def Q_interchange_columns_def Q_interchange_columns'_new_def Q_interchange_columns_new_def interchange_columns_matrix_def interchange_columns_infmatrix_def using n m unfolding 1 by simp
qed


lemma Q_interchange_columns_new_inverse:
assumes n:"n < a"
and m: "m < a"
shows "inverse_matrix (Q_interchange_columns'_new a n m) (Q_interchange_columns'_new a n m)"

proof -
have 1: "ncols ((one_matrix a)::int matrix) = a" using ncols_one_matrix[of a] by auto
show ?thesis using Q_interchange_columns_inverse[of n "one_matrix a" m]
unfolding Q_interchange_columns'_def Q_interchange_columns_def Q_interchange_columns'_new_def Q_interchange_columns_new_def interchange_columns_matrix_def interchange_columns_infmatrix_def using n m unfolding 1 by simp
qed

lemma Rep_Abs_Q_interchange_columns_new:
assumes n:"n<z" and m:"m<z"
shows "Rep_matrix (Abs_matrix (λi j. if i = n ∧ j = m then 1::int else if i = m ∧ j = n then 1 else if i = j ∧ i ≠ m ∧ i ≠ n ∧ j < z then 1 else 0)) =
(λi j. if i = n ∧ j = m then 1 else if i = m ∧ j = n then 1::int else if i = j ∧ i ≠ m ∧ i ≠ n ∧ j < z then 1 else 0)"

proof (rule RepAbs_matrix)
show "∃ma. ∀j i. ma ≤ j --> (if j = n ∧ i = m then 1 else if j = m ∧ i = n then 1 else if j = i ∧ j ≠ m ∧ j ≠ n ∧ i < z then 1 else 0) = 0"
by (metis m n order_less_le order_less_le_trans)
show "∃na. ∀j i. na ≤ i --> (if j = n ∧ i = m then 1 else if j = m ∧ i = n then 1 else if j = i ∧ j ≠ m ∧ j ≠ n ∧ i < z then 1 else 0) = 0"
by (metis leD m n)
qed


lemma Q_interchange_columns_new_nrows:
assumes "n<a" and "m<a"
shows "nrows (Q_interchange_columns'_new a n m) = a"

by (metis Q_interchange_columns_new_def Q_interchange_columns_new_eq assms(1) assms(2) interchange_columns_nrows ncols_one_matrix nrows_one_matrix)


lemma interchange_columns_AQ_new:
assumes n:"n<ncols A"
and m: "m<ncols A"
and na: "ncols A ≤ z"
shows "interchange_columns_matrix (A::int matrix) n m = A*Q_interchange_columns'_new z n m"

proof -
have "Rep_matrix (interchange_columns_matrix (A::int matrix) n m) = Rep_matrix (A*Q_interchange_columns'_new z n m)"
proof (rule, rule)
fix a b
have max_eq: "max (ncols A) (nrows (Abs_matrix (λi j. if i = n ∧ j = m then 1::int else if i = m ∧ j = n then 1 else if i = j ∧ i ≠ m ∧ i ≠ n ∧ j < z then 1 else 0)))
= z"

by (metis Q_interchange_columns'_new_def Q_interchange_columns_new_nrows linorder_not_less m min_max.sup.commute min_max.sup_absorb1 n na xt1(6))
have nz: "n<z" and mz: "m<z"using n m na by auto
show "Rep_matrix (interchange_columns_matrix A n m) a b = Rep_matrix (A * Q_interchange_columns'_new z n m) a b"
unfolding interchange_columns_matrix_def interchange_columns_infmatrix_def Q_interchange_columns'_new_def
unfolding Rep_matrix_mult unfolding Rep_Abs_aux_interchange'[of n A m] unfolding Rep_Abs_Q_interchange_columns_new[OF nz mz] unfolding foldseq_finsum
unfolding max_eq
proof (auto)
show "Rep_matrix A a n = (\<Oplus>\<Z>k∈{..z}. Rep_matrix A a k * (if k = n then 1 else if k = m ∧ m = n then 1 else if k = m ∧ k ≠ m ∧ k ≠ n ∧ m < z then 1 else 0))"
proof -
have "finsum \<Z> (λk. Rep_matrix A a k * (if k = n then 1 else if k = m ∧ m = n then 1 else if k = m ∧ k ≠ m ∧ k ≠ n ∧ m < z then 1 else 0)) {..z} =
finsum \<Z> (λk. if n=k then Rep_matrix A a n else \<zero>\<Z>) {..z}"
apply (rule Ring.abelian_monoid.finsum_cong') using abelian_monoid_z by auto
also have "... = Rep_matrix A a n" apply (rule Ring.abelian_monoid.finsum_singleton) using nz using abelian_monoid_z by auto
finally show ?thesis by simp
qed
show "Rep_matrix A a m = (\<Oplus>\<Z>k∈{..z}. Rep_matrix A a k * (if k = m then 1 else if k = n ∧ k ≠ m ∧ k ≠ n ∧ n < z then 1 else 0))"
proof -
have "finsum \<Z> (λk. Rep_matrix A a k * (if k = m then 1 else if k = n ∧ k ≠ m ∧ k ≠ n ∧ n < z then 1 else 0)) {..z} =
finsum \<Z> (λk. if m=k then Rep_matrix A a m else \<zero>\<Z>) {..z}"
apply (rule Ring.abelian_monoid.finsum_cong') using abelian_monoid_z by auto
also have "... = Rep_matrix A a m" apply (rule Ring.abelian_monoid.finsum_singleton) using mz using abelian_monoid_z by auto
finally show ?thesis by simp
qed next
assume b_not_m: "b ≠ m" and b_not_n:"b ≠ n"
show "Rep_matrix A a b = (\<Oplus>\<Z>k∈{..z}. Rep_matrix A a k * (if k = b ∧ k ≠ m ∧ k ≠ n ∧ b < z then 1 else 0))"
proof -
have "(\<Oplus>\<Z>k∈{..z}. Rep_matrix A a k * (if k = b ∧ k ≠ m ∧ k ≠ n ∧ b < z then 1 else 0)) = finsum \<Z> (λk. Rep_matrix A a k * (if k = b ∧ b < z then 1 else 0)) {..z}"
apply (rule Ring.abelian_monoid.finsum_cong') apply (rule abelian_monoid_z) using b_not_m b_not_n by auto
also have "... = Rep_matrix A a b"
proof (cases "b < z")
case False
have "finsum \<Z> (λk. Rep_matrix A a k * (if k = b ∧ b < z then 1 else 0)) {..z}
= finsum \<Z> (λk. \<zero>\<Z>) {..z}"

apply (rule Ring.abelian_monoid.finsum_cong') apply (rule abelian_monoid_z) apply rule apply simp
by (metis (full_types) False UNIV(1) int.m_comm mult_zero_left ring.simps(1))
also have "...=\<zero>\<Z>" using Ring.abelian_monoid.finsum_zero[OF abelian_monoid_z] by auto
also have "...=0" by simp
also have "... = Rep_matrix A a b" by (metis False linorder_not_less na ncols_le)
finally show ?thesis by simp
next
case True
have "(\<Oplus>\<Z>k∈{..z}. Rep_matrix A a k * (if k = b ∧ b < z then 1 else 0)) = (\<Oplus>\<Z>k∈{..z}. Rep_matrix A a k * (if b = k then 1 else 0))"
apply (rule Ring.abelian_monoid.finsum_cong'[OF abelian_monoid_z]) apply auto using True by (metis na ncols)
also have "...=(\<Oplus>\<Z>k∈{..z}. (if b = k then Rep_matrix A a b else \<zero>\<Z>))"
apply (rule abelian_monoid.finsum_cong') using abelian_monoid_z by auto
also have "...=Rep_matrix A a b" apply (rule Ring.abelian_monoid.finsum_singleton) using abelian_monoid_z True by auto
finally show ?thesis by simp
qed
finally show ?thesis by simp
qed
qed
qed
thus ?thesis using Rep_matrix_inject by auto
qed


definition P_row_add_new :: "nat => nat => nat => int => int matrix"
where "P_row_add_new a i j q = row_add_matrix (one_matrix a) i j q"


definition P_row_add'_new :: "nat => nat => nat => int => int matrix"
where "P_row_add'_new a n m q = Abs_matrix(λi j. if (i=n ∧ j=m) then q else (if (i=j ∧ i<a) then 1 else 0))"



lemma P_row_add_new_eq:
assumes n:"n < a"
and m:"m < a"
and n_not_m: "n≠m"
shows "P_row_add'_new a n m q = P_row_add_new a n m q"

proof -
have 1: "nrows ((one_matrix a)::int matrix) = a" by auto
show ?thesis using P_row_add_eq[of n "one_matrix a" m q]
unfolding P_row_add'_def P_row_add_def P_row_add'_new_def P_row_add_new_def row_add_matrix_def row_add_infmatrix_def unfolding 1 using n m n_not_m by simp
qed


lemma P_row_add_new_inverse:
assumes n:"n<a"
and m:"m<a"
and n_not_m: "n≠m"
shows "inverse_matrix (P_row_add'_new a n m q) (P_row_add'_new a n m (-q))"

proof -
have 1: "nrows ((one_matrix a)::int matrix) = a" by auto
show ?thesis using P_row_add_inverse[of n "one_matrix a" m q]
unfolding P_row_add'_def P_row_add_def P_row_add'_new_def P_row_add_new_def row_add_matrix_def row_add_infmatrix_def unfolding 1 using n m n_not_m by simp
qed

lemma RepAbs_row_add_new:
assumes "n<z" and "m<z"
shows " Rep_matrix (Abs_matrix (λi j. if i = n ∧ j = m then k else if i = j ∧ i < z then 1::int else 0))
= (λi j. if i = n ∧ j = m then k else if i = j ∧ i < z then 1::int else 0) "

proof (rule RepAbs_matrix)
show "∃ma. ∀j i. ma ≤ j --> (if j = n ∧ i = m then k else if j = i ∧ j < z then 1 else 0) = 0"
by (metis assms(1) leD)
show "∃na. ∀j i. na ≤ i --> (if j = n ∧ i = m then k else if j = i ∧ j < z then 1 else 0) = 0"
by (metis assms(2) leD)
qed


lemma P_row_add'_new_ncols:
assumes n: "n<a" and m: "m<a" and n_not_m: "n≠m"
shows "ncols (P_row_add'_new a n m q) = a"

proof -
have 1: "nrows ((one_matrix a)::int matrix) = a" by auto
show ?thesis using P_row_add_ncols[of n "one_matrix a" m q]
using P_row_add_new_eq[of n a m q]
unfolding P_row_add_def
unfolding P_row_add'_new_def P_row_add_new_def
unfolding row_add_matrix_def row_add_infmatrix_def
using n m n_not_m unfolding 1 by auto
qed

lemma row_add_new_PA:
assumes n:"n<nrows (A::int matrix)"
and m:"m<nrows A"
and na: "nrows A ≤ z "
and n_not_m: "n≠m"
shows "row_add_matrix A n m q = P_row_add'_new z n m q * A"

proof -
have "Rep_matrix (row_add_matrix A n m q) = Rep_matrix (P_row_add'_new z n m q * A)"
proof (rule, rule)
fix a b
have max_eq: "max (ncols (Abs_matrix (λi j. if i = n ∧ j = m then q else if i = j ∧ i < z then 1 else 0))) (nrows A) = z"
by (metis P_row_add'_new_def P_row_add'_new_ncols m min_max.sup_absorb1 n n_not_m na order_less_le_trans)
have nz: "n<z" and mz: "m<z" using n m na by auto
show "Rep_matrix (row_add_matrix A n m q) a b = Rep_matrix (P_row_add'_new z n m q * A) a b"
unfolding row_add_matrix_def P_row_add'_new_def row_add_infmatrix_def unfolding Rep_aux_row_add unfolding Rep_matrix_mult
unfolding RepAbs_row_add_new[OF nz mz] unfolding max_eq unfolding foldseq_finsum
proof (auto)
show "Rep_matrix A n b + q * Rep_matrix A m b = (\<Oplus>\<Z>k∈{..z}. (if k = m then q else if n = k ∧ n < z then 1 else 0) * Rep_matrix A k b)"
proof -
def s"(λk. (if k = m then q else if n = k ∧ n < z then 1 else 0) * Rep_matrix A k b)"
--"Demuestro unas cuantas cosillas que nos serviran luego en el calculation"

have eq_set: "{..z} = (insert m ({..z}-{m}))" using mz by auto
have eq_set2: "{..z}-{m} = insert n ({..z}-{m}-{n})"using nz n_not_m by auto
have finsum_eq: "finsum \<Z> s ({..z}-{m}) = s n ⊕\<Z> finsum \<Z> s ({..z}-{m}-{n})"
proof -
have "finsum \<Z> s ({..z}-{m}) = finsum \<Z> s (insert n ({..z}-{m}-{n}))" using eq_set2 by auto
also have "... = s n ⊕\<Z> finsum \<Z> s ({..z}-{m}-{n})" apply (rule Ring.abelian_monoid.finsum_insert) using abelian_monoid_z by auto
finally show ?thesis .
qed
have finsum_eq_zero: "finsum \<Z> s ({..z}-{m}-{n}) = \<zero>\<Z> "
proof -
have "finsum \<Z> s ({..z}-{m}-{n}) = finsum \<Z> (λk. \<zero>\<Z>) ({..z}-{m}-{n})" apply (rule Ring.abelian_monoid.finsum_cong') using abelian_monoid_z s_def by auto
also have "...=\<zero>\<Z>" apply (rule Ring.abelian_monoid.finsum_zero) using abelian_monoid_z by auto
finally show ?thesis .
qed
have "finsum \<Z> s {..z} = finsum \<Z> s (insert m ({..z}-{m}))" using eq_set by simp
also have "...=s m ⊕\<Z> finsum \<Z> s ({..z}-{m})" apply (rule Ring.abelian_monoid.finsum_insert) using abelian_monoid_z by auto
also have "... = s m ⊕\<Z> s n ⊕\<Z> finsum \<Z> s ({..z}-{m}-{n})" using finsum_eq by auto
also have "... = s m ⊕\<Z> s n" using finsum_eq_zero by auto
also have "... = s n + s m" by simp
also have "... = Rep_matrix A n b + q * Rep_matrix A m b" unfolding s_def using nz mz n_not_m by auto
finally show ?thesis unfolding s_def by simp
qed
show "Rep_matrix A a b = (\<Oplus>\<Z>k∈{..z}. (if a = k ∧ a < z then 1 else 0) * Rep_matrix A k b)"
proof (cases "a<z")
case False
have "finsum \<Z> (λk. (if a = k ∧ a < z then 1 else 0) * Rep_matrix A k b) {..z}
= finsum \<Z> (λk. \<zero>\<Z>) {..z}"

apply (rule Ring.abelian_monoid.finsum_cong') apply (rule abelian_monoid_z) apply rule apply simp
by (metis (full_types) False int_zero_eq mult_eq_0_iff)
also have "...=\<zero>\<Z>" using Ring.abelian_monoid.finsum_zero[OF abelian_monoid_z] by auto
also have "...=0" by simp
also have "... = Rep_matrix A a b" by (metis False linorder_not_less na nrows_le)
finally show ?thesis by simp
next
case True
have "(\<Oplus>\<Z>k∈{..z}. (if a = k ∧ a < z then 1 else 0) * Rep_matrix A k b) = (\<Oplus>\<Z>k∈{..z}. (if a = k then 1 else 0) * Rep_matrix A k b)"
using Ring.abelian_monoid.finsum_cong'[OF abelian_monoid_z] using True by auto
also have "...=(\<Oplus>\<Z>k∈{..z}. (if a = k then Rep_matrix A k b else \<zero>\<Z>))"
apply (rule abelian_monoid.finsum_cong') using abelian_monoid_z by auto
also have "...=Rep_matrix A a b" apply (rule Ring.abelian_monoid.finsum_singleton) using abelian_monoid_z True by auto
finally show ?thesis by simp
qed
qed
qed
thus ?thesis using Rep_matrix_inject by auto
qed



definition Q_column_add_new :: "nat => nat => nat => int => int matrix"
where "Q_column_add_new a i j q = column_add_matrix (one_matrix a) i j q"


definition Q_column_add'_new :: "nat => nat => nat => int => int matrix"
where "Q_column_add'_new a n m q = Abs_matrix(λi j. if (i=m ∧ j=n) then q else (if (i=j ∧ j < a) then 1 else 0))"


lemma Q_column_add_new_eq:
assumes n:"n < a"
and m:"m < a"
and n_not_m: "n≠m"
shows "Q_column_add_new a n m q = Q_column_add'_new a n m q"

proof -
have 1: "ncols ((one_matrix a)::int matrix)=a" by auto
show ?thesis using Q_column_add_eq[of n "one_matrix a" m q] unfolding Q_column_add'_def Q_column_add_def Q_column_add_new_def Q_column_add'_new_def unfolding 1
using n m n_not_m by simp
qed


lemma Q_column_add_new_inverse:
assumes n:"n<a"
and m:"m<a"
and n_not_m: "n≠m"
shows "inverse_matrix (Q_column_add'_new a n m q) (Q_column_add'_new a n m (-q))"

proof -
have 1: "ncols ((one_matrix a)::int matrix)=a" by auto
show ?thesis using Q_column_add_inverse[of n "one_matrix a" m q] unfolding Q_column_add'_def Q_column_add_def Q_column_add_new_def Q_column_add'_new_def unfolding 1
using n m n_not_m by simp
qed


lemma Q_column_add'_new_nrows:
assumes n: "n<a" and m: "m<a" and n_not_m: "n≠m"
shows "nrows (Q_column_add'_new a n m q) = a"

proof -
have 1: "ncols ((one_matrix a)::int matrix) = a" by auto
show ?thesis using Q_column_add_nrows[of n "one_matrix a" m q]
using Q_column_add_new_eq[of n a m q]
unfolding Q_column_add_def
unfolding Q_column_add'_new_def Q_column_add_new_def
unfolding column_add_matrix_def column_add_infmatrix_def
using n m n_not_m unfolding 1 by auto
qed


lemma RepAbs_column_add_new:
assumes n:"n<z" and m:"m<z"
shows " Rep_matrix (Abs_matrix (λi j. if i = m ∧ j = n then k else if i = j ∧ j < z then 1::int else 0))
= (λi j. if i = m ∧ j = n then k else if i = j ∧ j < z then 1::int else 0) "

proof (rule RepAbs_matrix)
show "∃ma. ∀j i. ma ≤ j --> (if j = m ∧ i = n then k else if j = i ∧ i < z then 1 else 0) = 0"
by (metis assms(2) leD)
show "∃na. ∀j i. na ≤ i --> (if j = m ∧ i = n then k else if j = i ∧ i < z then 1 else 0) = 0"
by (metis leD n)
qed


lemma column_add_new_AQ:
assumes n:"n<ncols (A::int matrix)"
and m:"m<ncols A"
and na: "ncols A ≤ z "
and n_not_m: "n≠m"
shows "column_add_matrix A n m q = A * Q_column_add'_new z n m q"

proof -
have "Rep_matrix (column_add_matrix A n m q) = Rep_matrix (A * Q_column_add'_new z n m q)"
proof (rule, rule)
fix a b
have nz:"n<z" and mz: "m<z" using n m na by auto
have max_eq: "max (ncols A) (nrows (Abs_matrix (λi j. if i = m ∧ j = n then q else if i = j ∧ j < z then 1 else 0))) = z"
by (metis (no_types) Q_column_add'_new_def Q_column_add_def Q_column_add_new_def Q_column_add_new_eq
Q_column_add_nrows min_max.sup.commute min_max.sup_absorb1 mz n_not_m na ncols_one_matrix nz)

show"Rep_matrix (column_add_matrix A n m q) a b = Rep_matrix (A * Q_column_add'_new z n m q) a b"
unfolding column_add_matrix_def column_add_infmatrix_def Rep_matrix_mult Q_column_add'_new_def unfolding max_eq unfolding RepAbs_column_add_new[OF nz mz, of q]
unfolding foldseq_finsum unfolding Rep_aux_column_add
proof auto
show " Rep_matrix A a n + q * Rep_matrix A a m = (\<Oplus>\<Z>k∈{..z}. Rep_matrix A a k * (if k = m then q else if k = n ∧ n < z then 1 else 0))"
proof -
def s"(λk. Rep_matrix A a k * (if k = m then q else if k = n ∧ n < z then 1 else 0))"
--"Demuestro unas cuantas cosillas que nos serviran luego en el calculation"

have eq_set: "{..z} = (insert m ({..z}-{m}))" using mz by auto
have eq_set2: "{..z}-{m} = insert n ({..z}-{m}-{n})"using nz n_not_m by auto
have finsum_eq: "finsum \<Z> s ({..z}-{m}) = s n ⊕\<Z> finsum \<Z> s ({..z}-{m}-{n})"
proof -
have "finsum \<Z> s ({..z}-{m}) = finsum \<Z> s (insert n ({..z}-{m}-{n}))" using eq_set2 by auto
also have "... = s n ⊕\<Z> finsum \<Z> s ({..z}-{m}-{n})" apply (rule Ring.abelian_monoid.finsum_insert) using abelian_monoid_z by auto
finally show ?thesis .
qed
have finsum_eq_zero: "finsum \<Z> s ({..z}-{m}-{n}) = \<zero>\<Z> "
proof -
have "finsum \<Z> s ({..z}-{m}-{n}) = finsum \<Z> (λk. \<zero>\<Z>) ({..z}-{m}-{n})" apply (rule Ring.abelian_monoid.finsum_cong') using abelian_monoid_z s_def by auto
also have "...=\<zero>\<Z>" apply (rule Ring.abelian_monoid.finsum_zero) using abelian_monoid_z by auto
finally show ?thesis .
qed
have "finsum \<Z> s {..z} = finsum \<Z> s (insert m ({..z}-{m}))" using eq_set by simp
also have "...=s m ⊕\<Z> finsum \<Z> s ({..z}-{m})" apply (rule Ring.abelian_monoid.finsum_insert) using abelian_monoid_z by auto
also have "... = s m ⊕\<Z> s n ⊕\<Z> finsum \<Z> s ({..z}-{m}-{n})" using finsum_eq by auto
also have "... = s m ⊕\<Z> s n" using finsum_eq_zero by auto
also have "... = s n + s m" by simp
also have "... = Rep_matrix A a n + q * Rep_matrix A a m" unfolding s_def using nz mz n_not_m by simp
finally show ?thesis unfolding s_def by simp
qed
assume b_not_n: "b ≠ n"
show "Rep_matrix A a b = (\<Oplus>\<Z>k∈{..z}. Rep_matrix A a k * (if k = b ∧ b < z then 1 else 0))"
proof (cases "b<z")
case False
have "finsum \<Z> (λk. Rep_matrix A a k * (if k = b ∧ b < z then 1 else 0)) {..z}
= finsum \<Z> (λk. \<zero>\<Z>) {..z}"

apply (rule Ring.abelian_monoid.finsum_cong') apply (rule abelian_monoid_z) apply rule apply simp
by (metis (full_types) False comm_semiring_1_class.normalizing_semiring_rules(10) ring.simps(1))
also have "...=\<zero>\<Z>" using Ring.abelian_monoid.finsum_zero[OF abelian_monoid_z] by auto
also have "...=0" by simp
also have "... = Rep_matrix A a b" using False ncols_notzero
by (metis linorder_not_less na ncols_le)
finally show ?thesis by simp
next
case True
have "(\<Oplus>\<Z>k∈{..z}. Rep_matrix A a k * (if k = b ∧ b < z then 1 else 0)) = (\<Oplus>\<Z>k∈{..z}. Rep_matrix A a k * (if k = b then 1 else 0))"
using Ring.abelian_monoid.finsum_cong'[OF abelian_monoid_z] using True by auto
also have "...=(\<Oplus>\<Z>k∈{..z}. (if b = k then Rep_matrix A a k else \<zero>\<Z>))"
apply (rule abelian_monoid.finsum_cong') using abelian_monoid_z by auto
also have "...=Rep_matrix A a b" apply (rule Ring.abelian_monoid.finsum_singleton) using abelian_monoid_z True by auto
finally show ?thesis by simp
qed
qed
qed
thus ?thesis using Rep_matrix_inject by auto
qed

lemma PQ_move_element:
assumes i:"i<nrows (A::int matrix)" and k:"k<nrows A" and j:"j<ncols A" and l:"l<ncols A"
shows "∃P Q. is_invertible P ∧ is_invertible Q ∧ move_element A i j k l = P*A*Q"

proof (rule exI[of _ "(one_matrix (nrows (A))) *(P_interchange_rows'_new (nrows A) i k)"],
rule exI[of _ "(one_matrix (ncols A))* Q_interchange_columns'_new (ncols A) j l"], auto)

show "is_invertible (one_matrix (nrows A) * P_interchange_rows'_new (nrows A) i k)"
by (metis P_interchange_row_new_eq P_interchange_rows_def P_interchange_rows_new_def P_interchange_rows_new_inverse
P_interchange_rows_nrows_eq2 is_invertible_def i k one_matrix_mult_left order_eq_refl)

show "is_invertible (one_matrix (ncols A) * Q_interchange_columns'_new (ncols A) j l)"
by (metis Q_interchange_columns_new_def Q_interchange_columns_new_eq Q_interchange_columns_new_inverse
is_invertible_def interchange_columns_AQ_new j l ncols_one_matrix order_eq_refl)

show " move_element A i j k l = one_matrix (nrows A) * P_interchange_rows'_new (nrows A) i k * A * (one_matrix (ncols A) * Q_interchange_columns'_new (ncols A) j l)"
by (metis (no_types) P_interchange_row_eq P_interchange_row_new_eq P_interchange_rows_def P_interchange_rows_new_def
P_interchange_rows_nrows_eq2 Q_interchange_columns_new_def Q_interchange_columns_new_eq i interchange_columns_AQ_new interchange_rows_PA
interchange_rows_ncols j k l move_element_def ncols_one_matrix one_matrix_mult_left order_eq_refl)

qed


definition is_square :: "(int matrix) => nat => bool"
where "is_square A n = (nrows A = n ∧ ncols A = n)"


lemma PQ_move_element':
assumes i:"i<nrows (A::int matrix)" and k:"k<nrows A" and j:"j<ncols A" and l:"l<ncols A"
shows "∀z t. z≥nrows A ∧ t≥ncols A--> (∃P Q. is_invertible P ∧ is_invertible Q ∧ move_element A i j k l = P*A*Q ∧ is_square P z ∧ is_square Q t)"

proof (rule, rule, rule)
fix z t assume "nrows A ≤ z ∧ ncols A ≤ t" hence z: "nrows A ≤ z" and t: "ncols A ≤ t" by auto
show " ∃P Q. is_invertible P ∧ is_invertible Q ∧ move_element A i j k l = P * A * Q ∧ is_square P z ∧ is_square Q t"
proof (rule exI[of _ "((one_matrix z)::int matrix) *(P_interchange_rows'_new z i k)"],
rule exI[of _ "((one_matrix t)::int matrix)* Q_interchange_columns'_new t j l"], auto)

(*show "nrows (one_matrix z * P_interchange_rows'_new z i k) = z"
by (metis P_interchange_rows_new_nrows i k linorder_not_le one_matrix_mult_left order_less_asym xt1(6) z)*)

show "is_invertible (one_matrix z * P_interchange_rows'_new z i k)"
proof (rule mult_matrix_inverse2)
show "is_invertible (one_matrix z)"
by (metis is_invertible_def one_matrix_inverse)
show " is_invertible (P_interchange_rows'_new z i k)"
by (metis P_interchange_rows_new_inverse z is_invertible_def i k leD nat_le_linear order_le_less_trans order_neq_le_trans)
show "ncols ((one_matrix z)::int matrix) = nrows (P_interchange_rows'_new z i k)" using P_interchange_rows_new_nrows using i k z by auto
--"CURIOSO, LE TENGO QUE RECORDAR EL TIPO, SI QUITO EL INT MATRIX AL (ONE MATRIX a) NO LO SABE HACER."

qed
show "move_element A i j k l = one_matrix z * P_interchange_rows'_new z i k * A * (one_matrix t * Q_interchange_columns'_new t j l)"
proof -
have "move_element A i j k l = interchange_columns_matrix (interchange_rows_matrix A i k) j l" unfolding move_element_def ..
also have "... = interchange_columns_matrix (P_interchange_rows'_new z i k * A * (one_matrix t)) j l"
by (metis `nrows A ≤ z ∧ ncols A ≤ t` i interchange_rows_PA_new interchange_rows_ncols k one_matrix_mult_right)
also have "... = (P_interchange_rows'_new z i k * A * (one_matrix t)) * Q_interchange_columns'_new t j l"
by (metis `nrows A ≤ z ∧ ncols A ≤ t` i interchange_columns_AQ_new interchange_rows_PA_new interchange_rows_ncols j k l one_matrix_mult_right)
also have "... = P_interchange_rows'_new z i k * A * (one_matrix t * Q_interchange_columns'_new t j l)"
by (metis (no_types) Q_interchange_columns_new_def Q_interchange_columns_new_eq calculation i interchange_columns_AQ_new interchange_rows_PA_new
interchange_rows_ncols j k l move_element_def ncols_one_matrix order_eq_refl order_less_le_trans t z)

also have "... = one_matrix z * (P_interchange_rows'_new z i k * A * (one_matrix t * Q_interchange_columns'_new t j l))"
by (metis P_interchange_rows_new_nrows `nrows A ≤ z ∧ ncols A ≤ t` calculation i k nrows_mult one_matrix_mult_left order_less_le_trans order_trans)
also have "... = one_matrix z * P_interchange_rows'_new z i k * A * (one_matrix t * Q_interchange_columns'_new t j l)"
by (metis P_interchange_rows_new_nrows `P_interchange_rows'_new z i k * A * (one_matrix t * Q_interchange_columns'_new t j l)
= one_matrix z * (P_interchange_rows'_new z i k * A * (one_matrix t * Q_interchange_columns'_new t j l))`
`nrows A ≤ z ∧ ncols A ≤ t` i k
one_matrix_mult_left order_less_le order_refl xt1(7))

finally show ?thesis .
qed
(*show "ncols (one_matrix t * Q_interchange_columns'_new t j l) = t"
by (metis Q_interchange_columns_new_def Q_interchange_columns_new_eq Q_interchange_columns_new_inverse
Q_interchange_columns_new_nrows interchange_columns_AQ_new inverse_matrix_def j l ncols_one_matrix order_refl right_inverse_matrix_dim t xt1(8))*)

show "is_invertible (one_matrix t * Q_interchange_columns'_new t j l)" using one_matrix_mult_left[of "Q_interchange_columns'_new t j l" t]
by (metis Q_interchange_columns_new_def Q_interchange_columns_new_eq Q_interchange_columns_new_inverse is_invertible_def
interchange_columns_AQ_new j l le_refl ncols_one_matrix t xt1(8))

show " is_square (one_matrix z * P_interchange_rows'_new z i k) z"
by (metis P_interchange_rows_new_ncols P_interchange_rows_new_nrows `nrows A ≤ z ∧ ncols A ≤ t` is_square_def i k one_matrix_mult_left order_refl xt1(8))
show "is_square (one_matrix t * Q_interchange_columns'_new t j l) t"
proof -
have 1:"one_matrix t * Q_interchange_columns'_new t j l = Q_interchange_columns'_new t j l"
by (metis (no_types) Q_interchange_columns_new_def Q_interchange_columns_new_eq interchange_columns_AQ_new j l leD le_refl linorder_less_linear ncols_one_matrix order_less_trans order_neq_le_trans t)
have 2:"nrows (Q_interchange_columns'_new t j l) = ncols (Q_interchange_columns'_new t j l)"
by (metis Q_interchange_columns_new_inverse inverse_matrix_def j l order_neq_le_trans right_inverse_matrix_dim t xt1(10))
show ?thesis using 1 2 unfolding is_square_def
by (metis Q_interchange_columns_new_nrows j l t xt1(8))
qed
qed
qed


lemma snd_minNonzero_less_nrows:
assumes minNonzero: "fst_bis (minNonzero_nc A k) = True"
shows "snd_bis (minNonzero_nc A k) < nrows A"

proof -
have notall: "¬ (∀(i::nat) j::nat. i ∈ {k..<nrows A} ∧ j ∈ {k..<ncols A} --> Rep_matrix A i j = 0)"
using minNonzero_true_inv[OF minNonzero] .
def P=="(λpos. k ≤ fst pos ∧ k ≤ snd pos ∧ Rep_matrix A (fst pos) (snd pos) ≠ 0 ∧
(∀m n. k ≤ m ∧ k ≤ n ∧ Rep_matrix A m n ≠ 0 --> ¦Rep_matrix A (fst pos) (snd pos)¦ ≤ ¦Rep_matrix A m n¦))"

def pos=="LEAST k. P k"
have "P (pos)" using minNonzero_least[of P k A] minNonzero unfolding P_def pos_def by auto
have "Rep_matrix A (snd_bis (minNonzero_nc A k)) (trd_bis (minNonzero_nc A k))
=Rep_matrix A (fst pos) (snd pos)"
using notall unfolding minNonzero_nc_def snd_bis_def trd_bis_def pos_def P_def by auto
also have "... ≠ 0"
by (metis P_def `P pos`)
finally show ?thesis
by (metis nrows_notzero)
qed



lemma trd_minNonzero_less_ncols:
assumes minNonzero: "fst_bis (minNonzero_nc A k) = True"
shows "trd_bis (minNonzero_nc A k) < ncols A"

proof -
have notall: "¬ (∀(i::nat) j::nat. i ∈ {k..<nrows A} ∧ j ∈ {k..<ncols A} --> Rep_matrix A i j = 0)"
using minNonzero_true_inv[OF minNonzero] .
def P=="(λpos. k ≤ fst pos ∧ k ≤ snd pos ∧ Rep_matrix A (fst pos) (snd pos) ≠ 0 ∧
(∀m n. k ≤ m ∧ k ≤ n ∧ Rep_matrix A m n ≠ 0 --> ¦Rep_matrix A (fst pos) (snd pos)¦ ≤ ¦Rep_matrix A m n¦))"

def pos=="LEAST k. P k"
have "P (pos)" using minNonzero_least[of P k A] minNonzero unfolding P_def pos_def by auto
have "Rep_matrix A (snd_bis (minNonzero_nc A k)) (trd_bis (minNonzero_nc A k))
=Rep_matrix A (fst pos) (snd pos)"
using notall unfolding minNonzero_nc_def snd_bis_def trd_bis_def pos_def P_def by auto
also have "... ≠ 0"
by (metis P_def `P pos`)
finally show ?thesis
by (metis ncols_notzero)
qed


lemma move_minNonzero_PAQ:
assumes minNonzero: "fst_bis (minNonzero_nc A k) = True"
shows "∀z t. z≥nrows A ∧ t ≥ ncols A --> (∃P Q. is_invertible P ∧ is_invertible Q ∧ (move_minNonzero_nc A k) = P*A*Q ∧ is_square P z ∧ is_square Q t)"

unfolding move_minNonzero_nc_def
proof (rule PQ_move_element')
show "snd_bis (minNonzero_nc A k) < nrows A"
by (metis assms snd_minNonzero_less_nrows)
show "k < nrows A" by (metis `snd_bis (minNonzero_nc A k) < nrows A` assms k_le_snd_minNonzero xt1(7))
show "trd_bis (minNonzero_nc A k) < ncols A"
by (metis assms trd_minNonzero_less_ncols)
show "k < ncols A"
by (metis `trd_bis (minNonzero_nc A k) < ncols A` assms k_le_trd_minNonzero order_le_less_trans)
qed

thm partRowReduce_matrix_def thm partRowReduce_infmatrix_def thm iterate_partRowReduce.simps thm row_add_matrix_def


lemma nrows_row_add_matrix_le:
assumes i: "i<nrows A"
shows "nrows (row_add_matrix A i j k) ≤ nrows A"

proof -
have "Rep_matrix (row_add_matrix A i j k) (nrows A) j = Rep_matrix A (nrows A) j"
by (metis Rep_matrix_row_add_matrix_preserve assms less_not_refl)
also have "... = 0"
by (metis less_not_refl nrows_notzero)
finally show ?thesis
by (metis Rep_matrix_row_add_matrix_preserve assms leD nrows nrows_le)
qed



partial_function (tailrec) row_add_matrix_tail_rec' :: "int matrix => nat => nat => nat => int matrix"
where
"row_add_matrix_tail_rec' A i j k=
(if nrows A ≤ i then A
else row_add_matrix_tail_rec'
(row_add_matrix A (i + 1) k (- (((Rep_matrix A) (i + 1) j) div ((Rep_matrix A) k j))))
(i + 1)
j
k)"




lemma aux_nrows:
"nrows (row_add_matrix A (nrows (A::int matrix)) j (- (Rep_matrix A (nrows A) j div Rep_matrix A i j))) ≤ nrows A"

proof -
find_theorems "nrows ?A ≤ ?B"
have "(∀a b. (nrows A) ≤ a --> Rep_matrix (row_add_matrix A (nrows A) j (- (Rep_matrix A (nrows A) j div Rep_matrix A i j))) a b = 0)"
proof (rule, rule, rule)
fix a b
assume nr: "nrows A ≤ a"
have "Rep_matrix (row_add_matrix A (nrows A) j (- (Rep_matrix A (nrows A) j div Rep_matrix A i j))) a b = Rep_matrix
(Abs_matrix (λa b. if a = nrows A then Rep_matrix A (nrows A) b + - (Rep_matrix A (nrows A) j div Rep_matrix A i j) * Rep_matrix A j b else Rep_matrix A a b)) a b"

unfolding row_add_matrix_def row_add_infmatrix_def ..
also have "... = (λa b. if a = nrows A then Rep_matrix A (nrows A) b + - (Rep_matrix A (nrows A) j div Rep_matrix A i j) * Rep_matrix A j b else Rep_matrix A a b) a b"
using Rep_aux_row_add[of "nrows A" A "- (Rep_matrix A (nrows A) j div Rep_matrix A i j)" j] by auto
also have "... = 0"
proof (cases "a=nrows A")
case True
have "(λa b. if a = nrows A then Rep_matrix A (nrows A) b + - (Rep_matrix A (nrows A) j div Rep_matrix A i j) * Rep_matrix A j b else Rep_matrix A a b) a b
= Rep_matrix A (nrows A) b + - (Rep_matrix A (nrows A) j div Rep_matrix A i j) * Rep_matrix A j b"
using True by simp
also have "... = 0"
by (metis add.comm_neutral div_0 minus_zero mult_zero_right nrows order_refl zmult_commute)
finally show ?thesis by simp
next
case False
have "(if a = nrows A then Rep_matrix A (nrows A) b + - (Rep_matrix A (nrows A) j div Rep_matrix A i j) * Rep_matrix A j b else Rep_matrix A a b) = Rep_matrix A a b"
using False nr by auto
also have "... = 0" using False nr
by (metis nrows)
finally show ?thesis .
qed
finally show "Rep_matrix (row_add_matrix A (nrows A) j (- (Rep_matrix A (nrows A) j div Rep_matrix A i j))) a b = 0" .
qed
thus ?thesis using nrows_le by auto
qed

lemma Rep_matrix_row_add_matrix_preserve':
assumes "s≠i" shows "Rep_matrix (row_add_matrix A i j q) s t = Rep_matrix A s t"

using Rep_matrix_row_add_matrix_preserve assms by force

lemma row_add_infmatrix_preserve:
assumes "s≠i" shows "(row_add_infmatrix A i j q) s t = A s t"

unfolding row_add_infmatrix_def using assms by force


lemma aux_nrows':
"nrows (row_add_matrix A (nrows (A::int matrix)) z (- (Rep_matrix A (nrows A) j div Rep_matrix A z j))) ≤ nrows A"

by (metis aux_nrows less_irrefl_nat nrows_notzero zdiv_zero)


lemma
row_add_matrix_tail_rec_preserves'':
assumes s_le_i: "s ≤ i"
and z:"z≤i"
shows "Rep_matrix (row_add_matrix_tail_rec' (A::int matrix) i j z) s t = (Rep_matrix A) s t"

proof -
obtain p where "nrows A = p" and nr: "nrows A ≤ p" by simp
show ?thesis
using s_le_i nr z
proof (induct i arbitrary: A rule: measure_induct_rule [of "(λk. p - k)"])
case (less i)
show ?case
proof (cases "nrows A ≤ i")
case True
show ?thesis
unfolding row_add_matrix_tail_rec'.simps [of A i] by (simp add: True)
next
case False
hence i: "i < nrows A" by arith
show ?thesis
proof (cases "Suc i = nrows A")
case False
have rewrite_hyps: "Rep_matrix
(row_add_matrix_tail_rec' (row_add_matrix A (Suc i) z (- (Rep_matrix A (Suc i) j div Rep_matrix A z j))) (Suc i) j z)
s t =
Rep_matrix (row_add_matrix A (Suc i) z (- (Rep_matrix A (Suc i) j div Rep_matrix A z j)))
s t"

proof (rule less.hyps)
show "p - Suc i < p - i" using less.prems False i by auto
show "s ≤ Suc i" by (metis le_Suc_eq less(2))
show "nrows (row_add_matrix A (Suc i) z (- (Rep_matrix A (Suc i) j div Rep_matrix A z j))) ≤ p"
by (metis False Suc_lessI i le_trans less(3) nrows_row_add_matrix_le)
show "z ≤ Suc i" using less(4) by auto
qed
show ?thesis
unfolding row_add_matrix_tail_rec'.simps [of A i j z]
apply (simp add: False)
unfolding rewrite_hyps
using less.prems (1)
using Rep_matrix_row_add_matrix_preserve [of "Suc i" A z "(- (Rep_matrix A (Suc i) j div Rep_matrix A z j))"]
by simp
next
case True
have aux: "nrows (row_add_matrix A (Suc i) z (- (Rep_matrix A (Suc i) j div Rep_matrix A z j))) ≤ Suc i" unfolding True using aux_nrows' by metis
have "row_add_matrix_tail_rec' A i j z = (row_add_matrix_tail_rec' (row_add_matrix A (Suc i) z (- (Rep_matrix A (Suc i) j div Rep_matrix A z j))) (Suc i) j z)"
using row_add_matrix_tail_rec'.simps[of A i j] True by auto
also have "... = row_add_matrix A (Suc i) z (- (Rep_matrix A (Suc i) j div Rep_matrix A z j))"
using row_add_matrix_tail_rec'.simps[of "(row_add_matrix A (Suc i) z (- (Rep_matrix A (Suc i) j div Rep_matrix A z j)))" "Suc i" j] using aux by auto
finally show ?thesis
by (metis False Rep_matrix_row_add_matrix_preserve True less(2))
qed
qed
qed
qed


lemma row_add_matrix_tail_rec_eq_2'':
assumes nr_le_k: "nrows (A::int matrix) ≤ k"
and z: "z≤i"
shows "Rep_matrix (row_add_matrix_tail_rec' A i j z) k l = 0"

proof -
obtain p where p:"nrows A = p" and nr: "nrows A ≤ p" by simp
show ?thesis
using nr nr_le_k z
proof (induct i arbitrary: A rule: measure_induct_rule [of "(λi. p- i)"])
case (less i)
show "Rep_matrix (row_add_matrix_tail_rec' A i j z) k l = 0"
proof (cases "i<nrows A")
case False
show ?thesis
unfolding row_add_matrix_tail_rec'.simps [of A i j]
using False using nrows [of A k l] using less.prems by simp
next
case True note i=True
have rewr: "Rep_matrix (row_add_matrix_tail_rec' (row_add_matrix A (i + 1) z
(- (Rep_matrix A (i + 1) j div Rep_matrix A z j)))
(i + 1) j z)
k l = 0"

proof (cases "Suc i = nrows A")
case True
have aux: "nrows (row_add_matrix A (Suc i) z (- (Rep_matrix A (Suc i) j div Rep_matrix A z j))) ≤ Suc i" unfolding True using aux_nrows' by metis
have "row_add_matrix_tail_rec' A i j z = (row_add_matrix_tail_rec' (row_add_matrix A (Suc i) z (- (Rep_matrix A (Suc i) j div Rep_matrix A z j))) (Suc i) j z)"
using row_add_matrix_tail_rec'.simps[of A i j] True by auto
also have "... = row_add_matrix A (Suc i) z (- (Rep_matrix A (Suc i) j div Rep_matrix A z j))"
using row_add_matrix_tail_rec'.simps[of "(row_add_matrix A (Suc i) z (- (Rep_matrix A (Suc i) j div Rep_matrix A z j)))" "Suc i"] using aux by auto
finally show ?thesis
by (metis Rep_matrix_row_add_matrix_preserve Suc_eq_plus1 `row_add_matrix_tail_rec' A i j z
= row_add_matrix_tail_rec' (row_add_matrix A (Suc i) z (- (Rep_matrix A (Suc i) j div Rep_matrix A z j))) (Suc i) j z`

aux less(3) linorder_not_less nrows_notzero)

next
case False
show ?thesis
proof (rule less.hyps)
show "p - (i + 1) < p - i"
by (metis True diff_less_mono2 leD less(2) less_add_one linorder_le_less_linear xt1(6))
show "nrows (row_add_matrix A (i + 1) z (- (Rep_matrix A (i + 1) j div Rep_matrix A z j))) ≤ p"
by (metis False Suc_eq_plus1 Suc_lessI True less(2) nrows_row_add_matrix_le xt1(6))
show "nrows (row_add_matrix A (i + 1) z (- (Rep_matrix A (i + 1) j div Rep_matrix A z j))) ≤ k"
by (metis False Suc_eq_plus1 Suc_lessI True less(3) nrows_row_add_matrix_le xt1(6))
show "z≤i+1" using less(4) by simp
qed
qed
thus ?thesis
by (metis less(3) nrows row_add_matrix_tail_rec'.simps)
qed
qed
qed



lemma row_add_matrix_tail_rec_eq':
assumes i_l_k: "i<k"
and z_le_i: "z≤i"
shows "Rep_matrix (row_add_matrix_tail_rec' A i j z) k l =
row_add_infmatrix (Rep_matrix A) k z (- (Rep_matrix A k j div Rep_matrix A z j)) k l"

proof -
obtain p where "nrows A = p" and nr: "nrows A ≤ p" by simp
show ?thesis
using i_l_k nr z_le_i
proof (induct i arbitrary: A rule: measure_induct_rule [of "(λk. p - k)"])
case (less i)
show ?case
proof (cases "nrows A ≤ i")
case True
show ?thesis
unfolding row_add_matrix_tail_rec'.simps [of A i]
apply (simp add: True)
unfolding row_add_infmatrix_def using nrows [of A k] using less.prems (1) using True by simp
next
case False
hence i: "i<nrows A" by simp
show ?thesis
proof (cases "Suc i = nrows A")
case True
have "row_add_matrix_tail_rec' A i j z = row_add_matrix_tail_rec' (row_add_matrix A (i + 1) z (- (Rep_matrix A (i + 1) j div Rep_matrix A z j))) (i + 1) j z"
using row_add_matrix_tail_rec'.simps[of A i j z] using i by simp
also have "... = row_add_matrix A (i + 1) z (- (Rep_matrix A (i + 1) j div Rep_matrix A z j))" using aux_nrows'[of A z j]
using True using row_add_matrix_tail_rec'.simps by auto
finally have 1: "row_add_matrix_tail_rec' A i j z = row_add_matrix A (i + 1) z (- (Rep_matrix A (i + 1) j div Rep_matrix A z j))" .
show ?thesis
proof (cases "Suc i = k")
case True thus ?thesis unfolding 1 unfolding row_add_matrix_def unfolding row_add_infmatrix_def
unfolding Rep_aux_row_add by simp
next
case False
have k: "Suc i < k" using False i less.prems by auto
have k_nrows: "nrows A ≤ k" by (metis True k order_less_imp_le)
have 2: "Rep_matrix A k l + - (Rep_matrix A k j div Rep_matrix A z j) * Rep_matrix A z l = 0"
by (metis True add_0 div_0 less(2) minus_mult_commute mult_zero_right not_less_eq nrows_notzero zmult_commute)
show ?thesis unfolding 1 unfolding row_add_matrix_def unfolding row_add_infmatrix_def
unfolding Rep_aux_row_add using False unfolding 2 using nrows[OF k_nrows, of l] by auto
qed
next
case False
note Suc_i_not_nrows = False
def A'"(row_add_matrix A (i + 1) z (- (Rep_matrix A (i + 1) j div Rep_matrix A z j)))"
show ?thesis
proof (cases "i+1<k")
case True
have "Rep_matrix (row_add_matrix_tail_rec' A i j z) k l
= Rep_matrix (row_add_matrix_tail_rec' (row_add_matrix A (i + 1) z (- (Rep_matrix A (i + 1) j div Rep_matrix A z j))) (i + 1) j z) k l"

using row_add_matrix_tail_rec'.simps[of A i j z] using i by simp
also have "... = row_add_infmatrix (Rep_matrix A') k z (- (Rep_matrix A' k j div Rep_matrix A' z j)) k l"
proof (unfold A'_def, rule less.hyps)
show "p - (i + 1) < p - i"
by (metis Suc_eq_plus1 diff_less_mono2 i less(3) lessI order_less_le_trans)
show "nrows (row_add_matrix A (i + 1) z (- (Rep_matrix A (i + 1) j div Rep_matrix A z j))) ≤ p"
using nrows_row_add_matrix_le[of "i+1" A z] using Suc_i_not_nrows i
using less(3)
by (metis Suc_eq_plus1 Suc_lessI order_trans)
show "i + 1 < k" using True .
show "z ≤ i + 1" using less (4) by simp
qed
also have "... = row_add_infmatrix (Rep_matrix A) k z (- (Rep_matrix A k j div Rep_matrix A z j)) k l"
unfolding A'_def unfolding row_add_matrix_def row_add_infmatrix_def
unfolding Rep_aux_row_add using True using True less(4) by simp
finally show ?thesis .
next
case False
have i_k: "i+1=k"
by (metis False Suc_eq_plus1 Suc_lessI less(2))
have "Rep_matrix (row_add_matrix_tail_rec' A i j z) k l
= Rep_matrix (row_add_matrix_tail_rec' (row_add_matrix A (i + 1) z (- (Rep_matrix A (i + 1) j div Rep_matrix A z j))) (i + 1) j z) k l"

using row_add_matrix_tail_rec'.simps[of A i j z] using i by simp
also have "... = Rep_matrix A' k l" unfolding A'_def using row_add_matrix_tail_rec_preserves''
by (metis A'_def i_k le_Suc_eq less(2) less(4) less_eq_Suc_le order_eq_refl xt1(6))
finally show ?thesis unfolding A'_def unfolding row_add_matrix_def row_add_infmatrix_def unfolding Rep_aux_row_add using i_k by auto
qed
qed
qed
qed
qed



lemma
row_add_matrix_tail_rec_preserves':
assumes s_le_i: "s ≤ i"
shows "Rep_matrix (row_add_matrix_tail_rec' (A::int matrix) i j i) s t = (Rep_matrix A) s t"

proof -
obtain p where "nrows A = p" and nr: "nrows A ≤ p" by simp
def z == "i"
show ?thesis apply (subst (2) z_def [symmetric])
using s_le_i nr
proof (induct i arbitrary: A rule: measure_induct_rule [of "(λk. p - k)"])
case (less i)
show ?case
proof (cases "nrows A ≤ i")
case True
show ?thesis
unfolding row_add_matrix_tail_rec'.simps [of A i] by (simp add: True)
next
case False
hence i: "i < nrows A" by arith
show ?thesis
proof (cases "Suc i = nrows A")
case False
have rewrite_hyps: "Rep_matrix
(row_add_matrix_tail_rec' (row_add_matrix A (Suc i) z (- (Rep_matrix A (Suc i) j div Rep_matrix A z j))) (Suc i) j z)
s t =
Rep_matrix (row_add_matrix A (Suc i) z (- (Rep_matrix A (Suc i) j div Rep_matrix A z j)))
s t"

proof (rule less.hyps)
show "p - Suc i < p - i" using less.prems False i by auto
show "s ≤ Suc i" by (metis le_Suc_eq less(2))
show "nrows (row_add_matrix A (Suc i) z (- (Rep_matrix A (Suc i) j div Rep_matrix A z j))) ≤ p"
by (metis False Suc_lessI i le_trans less(3) nrows_row_add_matrix_le)
qed
show ?thesis
unfolding row_add_matrix_tail_rec'.simps [of A i j z]
apply (simp add: False)
unfolding rewrite_hyps
using less.prems (1)
using Rep_matrix_row_add_matrix_preserve [of "Suc i" A z "(- (Rep_matrix A (Suc i) j div Rep_matrix A z j))"]
by simp
next
case True
have aux: "nrows (row_add_matrix A (Suc i) z (- (Rep_matrix A (Suc i) j div Rep_matrix A z j))) ≤ Suc i" unfolding True using aux_nrows' by metis
have "row_add_matrix_tail_rec' A i j z = (row_add_matrix_tail_rec' (row_add_matrix A (Suc i) z (- (Rep_matrix A (Suc i) j div Rep_matrix A z j))) (Suc i) j z)"
using row_add_matrix_tail_rec'.simps[of A i j] True by auto
also have "... = row_add_matrix A (Suc i) z (- (Rep_matrix A (Suc i) j div Rep_matrix A z j))"
using row_add_matrix_tail_rec'.simps[of "(row_add_matrix A (Suc i) z (- (Rep_matrix A (Suc i) j div Rep_matrix A z j)))" "Suc i" j] using aux by auto
finally show ?thesis
by (metis False Rep_matrix_row_add_matrix_preserve True less(2))
qed
qed
qed
qed


lemma Rep_matrix_row_add_matrix_tail_rec'_eq_partRowReduce_matrix:
shows "Rep_matrix (row_add_matrix_tail_rec' A i j i) = Rep_matrix (partRowReduce_matrix A i j)"

proof (rule, rule)
fix a b
show "Rep_matrix (row_add_matrix_tail_rec' A i j i) a b = Rep_matrix (partRowReduce_matrix A i j) a b"
proof (cases "i<a")
case True
have "Rep_matrix (row_add_matrix_tail_rec' A i j i) a b=
row_add_infmatrix (Rep_matrix A) a i (- (Rep_matrix A a j div Rep_matrix A i j)) a b"

using row_add_matrix_tail_rec_eq'[OF True, of i] by auto
also have "... = Rep_matrix (partRowReduce_matrix A i j) a b" unfolding partRowReduce_matrix_def partRowReduce_infmatrix_def
unfolding aux_partRowReduce using True by simp
finally show ?thesis .
next
case False
have "Rep_matrix (row_add_matrix_tail_rec' A i j i) a b = Rep_matrix A a b"
by (metis False linorder_not_less row_add_matrix_tail_rec_preserves')
also have "... = Rep_matrix (partRowReduce_matrix A i j) a b"
by (metis False linorder_not_less partRowReduce')
finally show ?thesis .
qed
qed

corollary row_add_matrix_tail_rec'_eq_partRowReduce_matrix:
shows "row_add_matrix_tail_rec' A i j i = partRowReduce_matrix A i j"

by (metis Rep_matrix_inverse Rep_matrix_row_add_matrix_tail_rec'_eq_partRowReduce_matrix)


lemma row_add_new_PA':
assumes n:"n<nrows (A::int matrix)"
and m:"m<nrows A"
and na: "nrows A ≤ z "
and n_not_m: "n≠m"
shows "∃P. is_invertible P ∧ row_add_matrix A n m q = P * A ∧ is_square P z"

proof (rule exI[of _ "P_row_add'_new z n m q"], auto)
show "is_invertible (P_row_add'_new z n m q)"
by (metis P_row_add_new_inverse is_invertible_def m n n_not_m na xt1(8))
show "row_add_matrix A n m q = P_row_add'_new z n m q * A"
by (metis m n n_not_m na row_add_new_PA)
show "is_square (P_row_add'_new z n m q) z"
by (metis (no_types) P_row_add'_new_ncols P_row_add_def P_row_add_new_def P_row_add_new_eq P_row_add_nrows is_square_def m n n_not_m na nrows_one_matrix xt1(8))
qed


lemma row_add_matrix_tail_rec'_PA:
assumes z:"z≥nrows A"
and not_zero: "Rep_matrix A k k ≠ 0"
and i: "k≤i"
shows "(∃P. is_invertible P ∧ (row_add_matrix_tail_rec' A i k k) = P*A ∧ is_square P z)"

proof -
obtain p where "nrows A = p" and nr: "nrows A ≤ p" by simp
show ?thesis
using z not_zero i nr
proof (induct i arbitrary: A rule: measure_induct_rule [of "(λk. p - k)"])
case (less i)
show ?case
proof (cases "nrows A ≤ i")
case True
show ?thesis
unfolding row_add_matrix_tail_rec'.simps [of A i]
apply (simp add: True)
by (metis is_square_def is_invertible_def less(2) ncols_one_matrix nrows_one_matrix one_matrix_inverse one_matrix_mult_left)
next
case False
hence i: "i<nrows A" by simp
show ?thesis
proof (cases "Suc i = nrows A")
case True
have 1: "row_add_matrix_tail_rec' A i k k = A"
proof -
have "Rep_matrix (row_add_matrix_tail_rec' A i k k) = Rep_matrix (row_add_matrix A (i + 1) k (- (Rep_matrix A (i + 1) k div Rep_matrix A k k)))"
using aux_nrows'[of A k k]
using row_add_matrix_tail_rec'.simps[of A i k k] using i True
by (metis False Suc_eq_plus1 row_add_matrix_tail_rec'.simps)
also have "... = Rep_matrix A"
proof (rule, rule)
fix a b
show " Rep_matrix (row_add_matrix A (i + 1) k (- (Rep_matrix A (i + 1) k div Rep_matrix A k k))) a b = Rep_matrix A a b"
proof (cases "a = i+1")
case False
show ?thesis
by (metis False Rep_matrix_row_add_matrix_preserve')
next
case True
have "Rep_matrix (row_add_matrix A (i + 1) k (- (Rep_matrix A (i + 1) k div Rep_matrix A k k))) a b
= Rep_matrix A (i + 1) b + - (Rep_matrix A (i + 1) k div Rep_matrix A k k) * Rep_matrix A k b"

unfolding row_add_matrix_def row_add_infmatrix_def unfolding Rep_aux_row_add using True by auto
also have "... = 0"
by (metis Suc_eq_plus1 `Suc i = nrows A` comm_semiring_1_class.normalizing_semiring_rules(10)
div_0 nrows order_refl zadd_0_right zminus_0 zmult_commute)

also have "... = Rep_matrix A a b"
by (metis Suc_eq_plus1 True `Suc i = nrows A` eq_imp_le nrows)
finally show ?thesis .
qed
qed
finally show ?thesis using Rep_matrix_inject by auto
qed

show ?thesis
unfolding 1
proof (rule exI[of _ "one_matrix z"], auto)
show "is_invertible (one_matrix z)"
by (metis is_invertible_def one_matrix_inverse)
show "A = one_matrix z * A"
by (metis less(2) one_matrix_mult_left)
show "is_square (one_matrix z) z"
by (metis is_square_def ncols_one_matrix nrows_one_matrix)
qed

next
case False
def A' "(row_add_matrix A (i + 1) k (- (Rep_matrix A (i + 1) k div Rep_matrix A k k)))"
have "row_add_matrix_tail_rec' A i k k = row_add_matrix_tail_rec' A' (i+1) k k"
unfolding A'_def using row_add_matrix_tail_rec'.simps using i by simp
have "∃P. is_invertible P ∧ row_add_matrix_tail_rec' A' (i + 1) k k = P * A' ∧ is_square P z"
proof (unfold A'_def, rule less.hyps)
show "p - (i + 1) < p - i"
by (metis Suc_eq_plus1 diff_less_mono2 i less(5) lessI linorder_not_less xt1(7))
show "nrows (row_add_matrix A (i + 1) k (- (Rep_matrix A (i + 1) k div Rep_matrix A k k))) ≤ z"
proof -
have "nrows (row_add_matrix A (i + 1) k (- (Rep_matrix A (i + 1) k div Rep_matrix A k k))) ≤ nrows A "
by (metis False Suc_eq_plus1 Suc_lessI i nrows_row_add_matrix_le)
also have "... ≤ z" using less(2) .
finally show ?thesis .
qed
show " Rep_matrix (row_add_matrix A (i + 1) k (- (Rep_matrix A (i + 1) k div Rep_matrix A k k))) k k ≠ 0"
by (metis Rep_matrix_row_add_matrix_preserve less(3) less(4) less_add_one linorder_not_less)
show "k ≤ i + 1"
by (metis less(4) trans_le_add1)
show "nrows (row_add_matrix A (i + 1) k (- (Rep_matrix A (i + 1) k div Rep_matrix A k k))) ≤ p"
proof -
have "nrows (row_add_matrix A (i + 1) k (- (Rep_matrix A (i + 1) k div Rep_matrix A k k))) ≤ nrows A "
by (metis False Suc_eq_plus1 Suc_lessI i nrows_row_add_matrix_le)
also have "... ≤ p" using less(5) .
finally show ?thesis .
qed
qed
from this obtain P where inv_P: "is_invertible P" and eq: "row_add_matrix_tail_rec' A' (i + 1) k k = P * A'" and cuadrada: "is_square P z"
by blast

have "∃P'. is_invertible P' ∧ A' = P' * A ∧ is_square P' z" --"Hago primero el existencial con el teorema para no complicar el obtain"
proof (unfold A'_def, rule row_add_new_PA')
show "i + 1 < nrows A" using False i by simp
show "k < nrows A"
by (metis less(3) nrows_notzero)
show "nrows A ≤ z" using less(2) .
show "i + 1 ≠ k"
by (metis Suc_eq_plus1 Suc_n_not_le_n less(4))
qed

from this obtain P' where inv_P': "is_invertible P'" and eq_P': "A'= P' * A" and cuadradaP': "is_square P' z" by blast

show ?thesis
proof (rule exI[of _ "P*P'"], auto)
show "is_invertible (P * P')"
by (metis cuadrada cuadradaP' is_square_def inv_P inv_P' mult_matrix_inverse2)
show "is_square (P * P') z"
by (metis (no_types) cuadrada cuadradaP' is_square_def inv_P inv_P' mult_matrix_inverse)
have "row_add_matrix_tail_rec' A i k k = row_add_matrix_tail_rec' A' (i+1) k k"
unfolding A'_def using row_add_matrix_tail_rec'.simps using i by simp
also have "... = P * A'" using eq .
also have "... = P * (P' * A)" unfolding eq_P' ..
also have "... = P * P' * A" using mult_assoc by metis
finally show "row_add_matrix_tail_rec' A i k k = P * P' * A" .
qed
qed
qed
qed
qed



lemma partRowReduce_matrix_PA:
assumes z:"z≥nrows A"
and not_zero: "Rep_matrix A k k ≠ 0"
shows "(∃P. is_invertible P ∧ (partRowReduce_matrix A k k) = P*A ∧ is_square P z)"

using row_add_matrix_tail_rec'_PA[of A z k k] unfolding row_add_matrix_tail_rec'_eq_partRowReduce_matrix[of A k k] using assms by simp



lemma PQ_partRowReduce_move_minNonzero:
assumes z: "z≥nrows A"
and t: "t ≥ ncols A"
and not_all_zero: "fst_bis (minNonzero_nc A k) = True"
shows "∃P Q. is_invertible P ∧ is_invertible Q ∧ (partRowReduce_matrix(move_minNonzero_nc A k) k k) = P * A * Q ∧ is_square P z ∧ is_square Q t"

proof -
obtain P1 and Q1 where P1: "is_invertible P1" and Q1: "is_invertible Q1" and move: "(move_minNonzero_nc A k) = P1*A*Q1"
and cuadradaP1: "is_square P1 z" and cuadradaQ1: "is_square Q1 t"

using move_minNonzero_PAQ [OF not_all_zero] using assms by metis
have "(∃P. is_invertible P ∧ (partRowReduce_matrix (move_minNonzero_nc A k) k k) = P*(move_minNonzero_nc A k) ∧ is_square P z)"
proof (rule partRowReduce_matrix_PA)
show "nrows (move_minNonzero_nc A k) ≤ z"
by (metis cuadradaP1 is_square_def move nrows_mult order_trans)
show "Rep_matrix (move_minNonzero_nc A k) k k ≠ 0"
by (metis move_minNonzero_not_zero not_all_zero)
qed
from this obtain P2 where "is_invertible P2" and "is_square P2 z" and p:"(partRowReduce_matrix (move_minNonzero_nc A k) k k) = P2*(move_minNonzero_nc A k)" by auto
show ?thesis
proof (rule exI[of _ "P2*P1"], rule exI[of _ "Q1"], auto)
show "is_invertible (P2 * P1)"
by (metis P1 `is_square P2 z` `is_invertible P2` cuadradaP1 is_square_def mult_matrix_inverse2)
show " is_invertible Q1" using Q1 .
show "partRowReduce_matrix (move_minNonzero_nc A k) k k = P2 * P1 * A * Q1"
proof -
have "partRowReduce_matrix (move_minNonzero_nc A k) k k = P2*(move_minNonzero_nc A k)" using p .
also have "... = P2*(P1*A*Q1)" unfolding move ..
also have "... = P2*P1*A*Q1" using mult_assoc by metis
finally show ?thesis .
qed
show "is_square (P2 * P1) z"
by (metis (no_types) P1 `is_square P2 z` `is_invertible P2` cuadradaP1 is_square_def mult_matrix_inverse)
show "is_square Q1 t" using cuadradaQ1 .
qed
qed


lemma PQ_iterate_partRowReduce:
assumes z: "z≥nrows A"
and t: "t≥ncols A"
and k: "k<nrows A"
and k2: "k<ncols A"
shows "∃P Q. is_invertible P ∧ is_invertible Q ∧ iterate_partRowReduce A k = P*A*Q ∧ is_square P z ∧ is_square Q t"

using z t k k2
proof (induct "(A,k)" arbitrary: A k rule: measure_induct_rule [of f'])
fix A k assume h1:"!!Aa ka. [|f' (Aa, ka) < f' (A, k); nrows Aa ≤ z; ncols Aa ≤ t; ka < nrows Aa; ka < ncols Aa|]
==> ∃P Q. is_invertible P ∧ is_invertible Q ∧ iterate_partRowReduce Aa ka = P * Aa * Q ∧ is_square P z ∧ is_square Q t"

and z: "nrows A ≤ z" and t:"ncols A ≤ t" and k:"k < nrows A" and k2:"k < ncols A"

show "∃P Q. is_invertible P ∧ is_invertible Q ∧ iterate_partRowReduce A k = P*A*Q ∧ is_square P z ∧ is_square Q t"
proof (cases "∀i>k. Rep_matrix A i k = 0")
case True show ?thesis apply (subst iterate_partRowReduce.simps) apply (rule exI[of _ "(one_matrix z)::int matrix"]) apply (rule exI[of _ "(one_matrix t)::int matrix"])
using True apply simp
by (metis is_square_def is_invertible_def ncols_one_matrix nrows_one_matrix one_matrix_inverse one_matrix_mult_left one_matrix_mult_right t z)
next
case False note False1=False
have "∃P Q. is_invertible P ∧ is_invertible Q ∧ (partRowReduce_matrix(move_minNonzero_nc A k) k k) = P * A * Q ∧ is_square P z ∧ is_square Q t"
proof (rule PQ_partRowReduce_move_minNonzero)
show "fst_bis (minNonzero_nc A k) = True"
proof (rule minNonzero_true)
from False1 obtain i where i: "i>k" and rep: "Rep_matrix A i k ≠ 0" by auto
have aux: "k<ncols A" using rep ncols_notzero by auto
have "i <nrows A"
by (metis nrows_notzero rep)
thus " ¬ (∀i j. i ∈ {k..<nrows A} ∧ j ∈ {k..<ncols A} --> Rep_matrix A i j = 0)" using i rep aux by auto
qed
show "nrows A ≤ z" using z .
show "ncols A ≤ t" using t .
qed
from this obtain P Q where invertible_P: "is_invertible P" and invertible_Q: "is_invertible Q" and part_move: "(partRowReduce_matrix(move_minNonzero_nc A k) k k) = P*A*Q"
and cuadrada_P: "is_square P z"
and cuadrada_Q: "is_square Q t"
by auto
have 1: "∃P Q. is_invertible P ∧ is_invertible Q ∧ iterate_partRowReduce (partRowReduce_matrix(move_minNonzero_nc A k) k k) k
= P*(partRowReduce_matrix(move_minNonzero_nc A k) k k)*Q ∧ is_square P z ∧ is_square Q t"

proof (cases "∀i>k. Rep_matrix (partRowReduce_matrix(move_minNonzero_nc A k) k k) i k = 0")
case True
show ?thesis
proof (rule exI[of _ "(one_matrix z)::int matrix"], rule exI[of _ "(one_matrix t)::int matrix"], auto)
show "iterate_partRowReduce (partRowReduce_matrix (move_minNonzero_nc A k) k k) k = one_matrix z * partRowReduce_matrix (move_minNonzero_nc A k) k k * one_matrix t"
proof -
have "iterate_partRowReduce (partRowReduce_matrix (move_minNonzero_nc A k) k k) k = (partRowReduce_matrix (move_minNonzero_nc A k) k k)"
using True iterate_partRowReduce.simps[of "(partRowReduce_matrix (move_minNonzero_nc A k) k k)" k] by auto
also have "... = one_matrix z * partRowReduce_matrix (move_minNonzero_nc A k) k k"
by (metis cuadrada_P is_square_def nrows_mult one_matrix_mult_left order_trans part_move)
also have "... = one_matrix z * partRowReduce_matrix (move_minNonzero_nc A k) k k * one_matrix t"
by (metis `partRowReduce_matrix (move_minNonzero_nc A k) k k = one_matrix z * partRowReduce_matrix (move_minNonzero_nc A k) k k` cuadrada_Q
is_square_def ncols_mult one_matrix_mult_right part_move)

finally show ?thesis .
qed
show " is_invertible (one_matrix z)"
by (metis is_invertible_def one_matrix_inverse)
show "is_invertible (one_matrix t)"
by (metis is_invertible_def one_matrix_inverse)
show "is_square (one_matrix z) z"
by (metis is_square_def ncols_one_matrix nrows_one_matrix)
show "is_square (one_matrix t) t"
by (metis is_square_def ncols_one_matrix nrows_one_matrix)
qed
next
case False note False2=False
from this obtain j where j: "j>k" and nz_j: "Rep_matrix (partRowReduce_matrix(move_minNonzero_nc A k) k k) j k ≠ 0" by auto
show ?thesis
proof (rule h1, unfold f'_def, auto)
show move_notzero: "Rep_matrix (move_minNonzero_nc A k) k k = 0 ==> False"
proof -
assume eq_0: "Rep_matrix (move_minNonzero_nc A k) k k = 0"
have "Rep_matrix (move_minNonzero_nc A k) k k ≠ 0"
proof (rule move_minNonzero_not_zero, rule minNonzero_true)
from False1 obtain i where i: "i>k" and nz4: "Rep_matrix A i k ≠ 0" by auto
have j2: "i<nrows A" using nz4 nrows_notzero by auto
have k2: "k<ncols A"using nz4 ncols_notzero by auto
show "¬ (∀i j. i ∈ {k..<nrows A} ∧ j ∈ {k..<ncols A} --> Rep_matrix A i j = 0)" using nz4 i j2 k2 by force
qed
thus ?thesis using eq_0 by contradiction
qed
show "¦Rep_matrix (move_minNonzero_nc (partRowReduce_matrix (move_minNonzero_nc A k) k k) k) k k¦ < ¦Rep_matrix (move_minNonzero_nc A k) k k¦"
proof -
have "¦Rep_matrix (move_minNonzero_nc (partRowReduce_matrix (move_minNonzero_nc A k) k k) k) k k¦
< ¦Rep_matrix (partRowReduce_matrix (move_minNonzero_nc A k) k k) k k¦"

proof (rule Akk_strict_decreasing_move_minNonzero)
show "j>k" using j .
show " ¦Rep_matrix (partRowReduce_matrix (move_minNonzero_nc A k) k k) j k¦ < ¦Rep_matrix (partRowReduce_matrix (move_minNonzero_nc A k) k k) k k¦"
by (metis `Rep_matrix (move_minNonzero_nc A k) k k = 0 ==> False` j partRowReduce_abs partRowReduce_eq_abs)
show "Rep_matrix (partRowReduce_matrix (move_minNonzero_nc A k) k k) j k ≠ 0" using nz_j .
qed
also have "...= ¦Rep_matrix (move_minNonzero_nc A k) k k¦" using partRowReduce_eq_abs by auto
finally show ?thesis .
qed
show "k < nrows (partRowReduce_matrix (move_minNonzero_nc A k) k k)"
by (metis j nat_less_le nrows_notzero nz_j xt1(8))
show "k < ncols (partRowReduce_matrix (move_minNonzero_nc A k) k k)"
by (metis `!!thesis. (!!j. [|k < j; Rep_matrix (partRowReduce_matrix (move_minNonzero_nc A k) k k) j k ≠ 0|] ==> thesis) ==> thesis` ncols_notzero)
show "nrows (partRowReduce_matrix (move_minNonzero_nc A k) k k) ≤ z"
by (metis cuadrada_P is_square_def nrows_mult part_move xt1(6))
show "ncols (partRowReduce_matrix (move_minNonzero_nc A k) k k) ≤ t"
by (metis `∃P Q. is_invertible P ∧ is_invertible Q ∧ partRowReduce_matrix (move_minNonzero_nc A k) k k = P * A * Q ∧ is_square P z ∧ is_square Q t`
is_square_def ncols_mult)

qed
qed
from this obtain P' and Q' where invertible_P': "is_invertible P'"
and invertible_Q': "is_invertible Q'"
and iterate: "iterate_partRowReduce (partRowReduce_matrix(move_minNonzero_nc A k) k k) k
= P'*(partRowReduce_matrix(move_minNonzero_nc A k) k k)*Q'"
and is_squareP': "is_square P' z"
and is_squareQ': "is_square Q' t"
by auto
show ?thesis
proof (rule exI[of _ "P'*P"], rule exI[of _ "Q*Q'"], auto)
show "is_invertible (P' * P)"
by (metis cuadrada_P is_squareP' is_square_def invertible_P invertible_P' mult_matrix_inverse)
show "is_invertible (Q * Q')"
by (metis cuadrada_Q is_squareQ' is_square_def invertible_Q invertible_Q' mult_matrix_inverse)
show "iterate_partRowReduce A k = P' * P * A * (Q * Q')"
proof -
have "iterate_partRowReduce A k = iterate_partRowReduce (partRowReduce_matrix(move_minNonzero_nc A k) k k) k" using False iterate_partRowReduce.simps[of A k] by auto
also have "... = P'*(partRowReduce_matrix(move_minNonzero_nc A k) k k)*Q'" using iterate .
also have "... = P'*(P*A*Q)*Q'" unfolding part_move ..
also have "... = P'*(P*A*Q*Q')" using mult_assoc by metis
also have "... = P'*(P*A*(Q*Q'))" using mult_assoc by metis
also have "... = P' * P * A * (Q * Q')" using mult_assoc by metis
finally show ?thesis .
qed
show "is_square (P' * P) z"
by (metis (no_types) cuadrada_P is_squareP' is_square_def invertible_P invertible_P' mult_matrix_inverse)
show "is_square (Q * Q') t"
by (metis (no_types) cuadrada_Q is_squareQ' is_square_def invertible_Q invertible_Q' mult_matrix_inverse)
qed
qed
qed


lemma ncols_column_add_matrix_le:
assumes i: "i<ncols A"
shows "ncols (column_add_matrix A i j k) ≤ ncols A"

proof -
have "Rep_matrix (column_add_matrix A i j k) j (ncols A) = Rep_matrix A j (ncols A)"
by (metis Rep_matrix_column_add_matrix_preserve assms less_not_refl)
also have "... = 0"
by (metis less_not_refl ncols_notzero)
finally show ?thesis
by (metis Rep_matrix_column_add_matrix_preserve assms leD ncols ncols_le)
qed


partial_function (tailrec) column_add_matrix_tail_rec' :: "int matrix => nat => nat => nat => int matrix"
where
"column_add_matrix_tail_rec' A iterador i j=
(if ncols A ≤ (iterador) then A
else column_add_matrix_tail_rec'
(column_add_matrix A (iterador + 1) j (- (((Rep_matrix A) i (iterador + 1)) div ((Rep_matrix A) i j))))
(iterador + 1)
i
j)"




lemma aux_ncols:
"ncols (column_add_matrix A (ncols (A::int matrix)) j (- (Rep_matrix A i (ncols A) div Rep_matrix A i j))) ≤ ncols A"

proof -
have "(∀a b. (ncols A) ≤ b --> Rep_matrix (column_add_matrix A (ncols A) j (- (Rep_matrix A i (ncols A) div Rep_matrix A i j))) a b = 0)"
proof (rule, rule, rule)
fix a b
assume nr: "ncols A ≤ b"
have " Rep_matrix (column_add_matrix A (ncols A) j (- (Rep_matrix A i (ncols A) div Rep_matrix A i j))) a b
= Rep_matrix
(Abs_matrix (λa b. if b = ncols A then Rep_matrix A a (ncols A) + - (Rep_matrix A i (ncols A) div Rep_matrix A i j) * Rep_matrix A a j else Rep_matrix A a b)) a b"

unfolding column_add_matrix_def column_add_infmatrix_def ..
also have "... = (λa b. if b = ncols A then Rep_matrix A a (ncols A) + - (Rep_matrix A i (ncols A) div Rep_matrix A i j) * Rep_matrix A a j else Rep_matrix A a b) a b"
unfolding Rep_aux_column_add ..
also have "... = 0"
proof (cases "b=ncols A")
case True
have "(if b = ncols A then Rep_matrix A a (ncols A) + - (Rep_matrix A i (ncols A) div Rep_matrix A i j) * Rep_matrix A a j else Rep_matrix A a b)
= Rep_matrix A a (ncols A) + - (Rep_matrix A i (ncols A) div Rep_matrix A i j) * Rep_matrix A a j"
using True by simp
also have "... = 0"
by (metis True comm_semiring_1_class.normalizing_semiring_rules(10) comm_semiring_1_class.normalizing_semiring_rules(5)
comm_semiring_1_class.normalizing_semiring_rules(7) minus_mult_right ncols nr zdiv_zero)

finally show ?thesis by simp
next
case False
have " (if b = ncols A then Rep_matrix A a (ncols A) + - (Rep_matrix A i (ncols A) div Rep_matrix A i j) * Rep_matrix A a j else Rep_matrix A a b)= Rep_matrix A a b"
using False nr by auto
also have "... = 0" using False nr
by (metis ncols)
finally show ?thesis .
qed
finally show "Rep_matrix (column_add_matrix A (ncols A) j (- (Rep_matrix A i (ncols A) div Rep_matrix A i j))) a b = 0" .
qed
thus ?thesis using ncols_le by auto
qed

lemma Rep_matrix_column_add_matrix_preserve':
assumes "t≠i" shows "Rep_matrix (column_add_matrix A i j q) s t = Rep_matrix A s t"

using Rep_matrix_column_add_matrix_preserve assms by force

lemma column_add_infmatrix_preserve:
assumes "t≠i" shows "(column_add_infmatrix A i j q) s t = A s t"

unfolding column_add_infmatrix_def using assms by force

lemma aux_ncols':
"ncols (column_add_matrix A (ncols A) z (- (Rep_matrix A j (ncols A) div Rep_matrix A j z))) ≤ ncols (A::int matrix)"

using aux_ncols [of A z j] by simp



lemma
column_add_matrix_tail_rec_preserves'':
assumes s_le_i: "t ≤ i"
and z:"z≤i"
shows "Rep_matrix (column_add_matrix_tail_rec' (A::int matrix) i j z) s t = (Rep_matrix A) s t"

proof -
obtain p where "ncols A = p" and nr: "ncols A ≤ p" by simp
show ?thesis
using s_le_i nr z
proof (induct i arbitrary: A rule: measure_induct_rule [of "(λk. p - k)"])
case (less i)
show ?case
proof (cases "ncols A ≤ i")
case True
show ?thesis
unfolding column_add_matrix_tail_rec'.simps [of A i] by (simp add: True)
next
case False
hence i: "i < ncols A" by arith
show ?thesis
proof (cases "Suc i = ncols A")
case False
have rewrite_hyps: "Rep_matrix
(column_add_matrix_tail_rec' (column_add_matrix A (Suc i) z (- (Rep_matrix A j (Suc i) div Rep_matrix A j z))) (Suc i) j z)
s t =
Rep_matrix (column_add_matrix A (Suc i) z (- (Rep_matrix A j (Suc i) div Rep_matrix A j z)))
s t"

proof (rule less.hyps)
show "p - Suc i < p - i" using less.prems False i by auto
show "t≤ Suc i"
by (metis le_Suc_eq less(2))
show "z ≤ Suc i"by (metis le_Suc_eq less(4))
show "ncols (column_add_matrix A (Suc i) z (- (Rep_matrix A j (Suc i) div Rep_matrix A j z))) ≤ p"
by (metis False Suc_lessI i le_trans less(3) ncols_column_add_matrix_le)
qed
show ?thesis
unfolding column_add_matrix_tail_rec'.simps [of A i j z]
apply (simp add: False)
unfolding rewrite_hyps
by (metis Rep_matrix_column_add_matrix_preserve Suc_n_not_le_n less(2))
next
case True
have aux: "ncols (column_add_matrix A (Suc i) z (- (Rep_matrix A j (Suc i) div Rep_matrix A j z))) ≤ Suc i"
by (metis True aux_ncols)
have "column_add_matrix_tail_rec' A i j z = (column_add_matrix_tail_rec' (column_add_matrix A (Suc i) z (- (Rep_matrix A j (Suc i) div Rep_matrix A j z))) (Suc i) j z)"
using column_add_matrix_tail_rec'.simps[of A i j] True by auto
also have "... = column_add_matrix A (Suc i) z (- (Rep_matrix A j (Suc i) div Rep_matrix A j z))"
using column_add_matrix_tail_rec'.simps[of "column_add_matrix A (Suc i) z (- (Rep_matrix A j (Suc i) div Rep_matrix A j z))" "Suc i"] using aux by metis
finally show ?thesis
by (metis False Rep_matrix_column_add_matrix_preserve' True less(2))
qed
qed
qed
qed


lemma column_add_matrix_tail_rec_eq_2'':
assumes nr_le_k: "ncols (A::int matrix) ≤ l"
and z: "z≤i"
shows "Rep_matrix (column_add_matrix_tail_rec' A i j z) k l = 0"

proof -
obtain p where p:"ncols A = p" and nr: "ncols A ≤ p" by simp
show ?thesis
using nr nr_le_k z
proof (induct i arbitrary: A rule: measure_induct_rule [of "(λi. p- i)"])
case (less i)
show "Rep_matrix (column_add_matrix_tail_rec' A i j z) k l = 0"
proof (cases "i<ncols A")
case False
show ?thesis
unfolding column_add_matrix_tail_rec'.simps [of A i j]
by (metis (full_types) False le_neq_implies_less less(3) nat_le_linear ncols)
next
case True note i=True
have rewr: "Rep_matrix (column_add_matrix_tail_rec' (column_add_matrix A (Suc i) z (- (Rep_matrix A j (Suc i) div Rep_matrix A j z)))
(i + 1) j z)
k l = 0"

proof (cases "Suc i = ncols A")
case True
have aux: "ncols (column_add_matrix A (Suc i) z (- (Rep_matrix A j (Suc i) div Rep_matrix A j z))) ≤ Suc i" unfolding True using aux_ncols' by metis
have "column_add_matrix_tail_rec' A i j z = (column_add_matrix_tail_rec' (column_add_matrix A (Suc i) z (- (Rep_matrix A j (Suc i) div Rep_matrix A j z))) (Suc i) j z)"
using column_add_matrix_tail_rec'.simps[of A i j] True by auto
also have "... = column_add_matrix A (Suc i) z (- (Rep_matrix A j (Suc i) div Rep_matrix A j z))"
using column_add_matrix_tail_rec'.simps[of "(column_add_matrix A (Suc i) z (- (Rep_matrix A j (Suc i) div Rep_matrix A j z)))" "Suc i"] using aux by auto
finally show ?thesis
by (metis Suc_eq_plus1 True `column_add_matrix_tail_rec' A i j z = column_add_matrix_tail_rec' (column_add_matrix A (Suc i) z
(- (Rep_matrix A j (Suc i) div Rep_matrix A j z))) (Suc i) j z`
aux less(3) ncols_le)

next
case False
show ?thesis
proof (rule less.hyps)
show "p - (i + 1) < p - i"
by (metis True diff_less_mono2 leD less(2) less_add_one linorder_le_less_linear xt1(6))
show "ncols (column_add_matrix A (Suc i) z (- (Rep_matrix A j (Suc i) div Rep_matrix A j z))) ≤ p"
by (metis False Suc_lessI True le_trans less(2) ncols_column_add_matrix_le)
show "ncols (column_add_matrix A (Suc i) z (- (Rep_matrix A j (Suc i) div Rep_matrix A j z))) ≤ l"
by (metis `ncols (column_add_matrix A (Suc i) z (- (Rep_matrix A j (Suc i) div Rep_matrix A j z))) ≤ p` le_trans nr_le_k p)
show "z≤i+1" using less(4) by simp
qed
qed
thus ?thesis
by (metis Suc_eq_plus1 column_add_matrix_tail_rec'.simps less(3) ncols)
qed
qed
qed



lemma column_add_matrix_tail_rec_eq':
assumes i_l_k: "i<l"
and z_le_i: "z≤i"
shows "Rep_matrix (column_add_matrix_tail_rec' A i j z) k l =
column_add_infmatrix (Rep_matrix A) l z (- (Rep_matrix A j l div Rep_matrix A j z)) k l"

proof -
obtain p where "ncols A = p" and nr: "ncols A ≤ p" by simp
show ?thesis
using i_l_k nr z_le_i
proof (induct i arbitrary: A rule: measure_induct_rule [of "(λk. p - k)"])
case (less i)
show ?case
proof (cases "ncols A ≤ i")
case True
show ?thesis
unfolding column_add_matrix_tail_rec'.simps [of A i]
apply (simp add: True)
unfolding column_add_infmatrix_def using ncols [of A l] using less.prems (1) using True by simp
next
case False
hence i: "i<ncols A" by simp
show ?thesis
proof (cases "Suc i = ncols A")
case True
have "column_add_matrix_tail_rec' A i j z = column_add_matrix_tail_rec' (column_add_matrix A (i + 1) z (- (Rep_matrix A j (i + 1) div Rep_matrix A j z))) (i + 1) j z"
using column_add_matrix_tail_rec'.simps[of A i j z] using i by simp
also have "... = column_add_matrix A (i + 1) z (- (Rep_matrix A j (i + 1) div Rep_matrix A j z))"
using aux_ncols'[of A z j]
using True using column_add_matrix_tail_rec'.simps by auto
finally have 1: "column_add_matrix_tail_rec' A i j z = column_add_matrix A (i + 1) z (- (Rep_matrix A j (i + 1) div Rep_matrix A j z))" .
show ?thesis
proof (cases "Suc i = l")
case True thus ?thesis unfolding 1 unfolding column_add_matrix_def unfolding column_add_infmatrix_def
unfolding Rep_aux_column_add by simp
next
case False
have k: "Suc i < l" using False i less.prems by auto
have k_nrows: "ncols A ≤ l" by (metis True k order_less_imp_le)
have 2: "Rep_matrix A k l + - (Rep_matrix A j l div Rep_matrix A j z) * Rep_matrix A k z = 0"
by (metis comm_semiring_1_class.normalizing_semiring_rules(6) comm_semiring_1_class.normalizing_semiring_rules(7)
k_nrows mult_zero_right ncols zdiv_zero zminus_0)

show ?thesis unfolding 1 unfolding column_add_matrix_def unfolding column_add_infmatrix_def
unfolding Rep_aux_column_add using False unfolding 2 using ncols[OF k_nrows, of k] by auto
qed
next
case False
note Suc_i_not_nrows = False
def A'"(column_add_matrix A (i + 1) z (- (Rep_matrix A j (i + 1) div Rep_matrix A j z)))"
show ?thesis
proof (cases "i+1<l")
case True
have "Rep_matrix (column_add_matrix_tail_rec' A i j z) k l
= Rep_matrix (column_add_matrix_tail_rec' (column_add_matrix A (i + 1) z (- (Rep_matrix A j (i + 1) div Rep_matrix A j z))) (i + 1) j z) k l"

using column_add_matrix_tail_rec'.simps[of A i j z] using i by simp
also have "... = column_add_infmatrix (Rep_matrix A') l z (- (Rep_matrix A' j l div Rep_matrix A' j z)) k l"
proof (unfold A'_def, rule less.hyps)
show "p - (i + 1) < p - i"
by (metis Suc_eq_plus1 diff_less_mono2 i less(3) lessI order_less_le_trans)
show "ncols (column_add_matrix A (i + 1) z (- (Rep_matrix A j (i + 1) div Rep_matrix A j z))) ≤ p"
by (metis False Suc_eq_plus1 Suc_lessI i le_trans less(3) ncols_column_add_matrix_le)
show "i + 1 < l" using True .
show "z ≤ i + 1" using less (4) by simp
qed
also have "... = column_add_infmatrix (Rep_matrix A) l z (- (Rep_matrix A j l div Rep_matrix A j z)) k l"
unfolding A'_def unfolding column_add_matrix_def column_add_infmatrix_def
unfolding Rep_aux_column_add using True using True less(4) by simp
finally show ?thesis .
next
case False
have i_k: "i+1=l"
by (metis False Suc_eq_plus1 Suc_lessI less(2))
have "Rep_matrix (column_add_matrix_tail_rec' A i j z) k l
= Rep_matrix (column_add_matrix_tail_rec' (column_add_matrix A (i + 1) z (- (Rep_matrix A j (i + 1) div Rep_matrix A j z))) (i + 1) j z) k l"

using column_add_matrix_tail_rec'.simps[of A i j z] using i by simp
also have "... = Rep_matrix A' k l"
unfolding A'_def using column_add_matrix_tail_rec_preserves''
by (metis A'_def i_k le_Suc_eq less(2) less(4) less_eq_Suc_le order_eq_refl xt1(6))
finally show ?thesis unfolding A'_def unfolding column_add_matrix_def column_add_infmatrix_def unfolding Rep_aux_column_add using i_k by auto
qed
qed
qed
qed
qed



lemma
column_add_matrix_tail_rec_preserves':
assumes s_le_i: "t ≤ i"
shows "Rep_matrix (column_add_matrix_tail_rec' (A::int matrix) i j i) s t = (Rep_matrix A) s t"

proof -
obtain p where "ncols A = p" and nr: "ncols A ≤ p" by simp
def z == "i"
show ?thesis apply (subst (2) z_def [symmetric])
using s_le_i nr
proof (induct i arbitrary: A rule: measure_induct_rule [of "(λk. p - k)"])
case (less i)
show ?case
proof (cases "ncols A ≤ i")
case True
show ?thesis
unfolding column_add_matrix_tail_rec'.simps [of A i] by (simp add: True)
next
case False
hence i: "i < ncols A" by arith
show ?thesis
proof (cases "Suc i = ncols A")
case False
have rewrite_hyps: "Rep_matrix
(column_add_matrix_tail_rec' (column_add_matrix A (Suc i) z (- (Rep_matrix A j (Suc i) div Rep_matrix A j z))) (Suc i) j z)
s t =
Rep_matrix (column_add_matrix A (Suc i) z (- (Rep_matrix A j (Suc i) div Rep_matrix A j z)))
s t"

proof (rule less.hyps)
show "p - Suc i < p - i" using less.prems False i by auto
show "t ≤ Suc i" by (metis le_Suc_eq less(2))
show "ncols (column_add_matrix A (Suc i) z (- (Rep_matrix A j (Suc i) div Rep_matrix A j z))) ≤ p"
by (metis False Suc_lessI i le_trans less(3) ncols_column_add_matrix_le)
qed
show ?thesis
unfolding column_add_matrix_tail_rec'.simps [of A i j z]
apply (simp add: False)
unfolding rewrite_hyps
by (metis Rep_matrix_column_add_matrix_preserve Suc_n_not_le_n less(2))
next
case True

have aux: "ncols (column_add_matrix A (Suc i) z (- (Rep_matrix A j (Suc i) div Rep_matrix A j z))) ≤ Suc i" unfolding True using aux_ncols' by auto
have "column_add_matrix_tail_rec' A i j z = (column_add_matrix_tail_rec' (column_add_matrix A (Suc i) z (- (Rep_matrix A j (Suc i) div Rep_matrix A j z))) (Suc i) j z)"
using column_add_matrix_tail_rec'.simps[of A i j] True by auto
also have "... = column_add_matrix A (Suc i) z (- (Rep_matrix A j (Suc i) div Rep_matrix A j z))"
using column_add_matrix_tail_rec'.simps[of "(column_add_matrix A (Suc i) z (- (Rep_matrix A j (Suc i) div Rep_matrix A j z)))" "Suc i" j] using aux by auto
finally show ?thesis
by (metis False Rep_matrix_column_add_matrix_preserve' True less(2))
qed
qed
qed
qed


lemma Rep_matrix_column_add_matrix_tail_rec'_eq_partColumnReduce_matrix:
shows "Rep_matrix (column_add_matrix_tail_rec' A i j i) = Rep_matrix (partColumnReduce_matrix A j i)"

proof (rule, rule)
fix a b
show "Rep_matrix (column_add_matrix_tail_rec' A i j i) a b = Rep_matrix (partColumnReduce_matrix A j i) a b"
proof (cases "i<b")
case True
have "Rep_matrix (column_add_matrix_tail_rec' A i j i) a b=
column_add_infmatrix (Rep_matrix A) b i (- (Rep_matrix A j b div Rep_matrix A j i)) a b"

using column_add_matrix_tail_rec_eq'[of i b i A j a] True by auto
also have "... = Rep_matrix (partColumnReduce_matrix A j i) a b" unfolding partColumnReduce_matrix_def partColumnReduce_infmatrix_def
unfolding aux_partColumnReduce using True by simp
finally show ?thesis .
next
case False
have "Rep_matrix (column_add_matrix_tail_rec' A i j i) a b = Rep_matrix A a b"
by (metis False column_add_matrix_tail_rec_preserves' linorder_not_less)
also have "... = Rep_matrix (partColumnReduce_matrix A j i) a b"
by (metis False linorder_not_less partColumnReduce')
finally show ?thesis .
qed
qed

corollary column_add_matrix_tail_rec'_eq_partColumnReduce_matrix:
shows "column_add_matrix_tail_rec' A j i j = partColumnReduce_matrix A i j"

by (metis Rep_matrix_column_add_matrix_tail_rec'_eq_partColumnReduce_matrix Rep_matrix_inverse)

lemma column_add_new_AQ':
assumes n:"n<ncols (A::int matrix)"
and m:"m<ncols A"
and na: "ncols A ≤ z "
and n_not_m: "n≠m"
shows "∃Q. is_invertible Q ∧ column_add_matrix A n m k = A*Q ∧ is_square Q z"

proof (rule exI[of _ "Q_column_add'_new z n m k"], auto)
show "is_invertible (Q_column_add'_new z n m k)"
by (metis Q_column_add_new_inverse is_invertible_def le_antisym less_imp_le_nat linorder_less_linear m n n_not_m na order_less_trans)
show "column_add_matrix A n m k = A * Q_column_add'_new z n m k"
by (metis m n n_not_m na column_add_new_AQ)
show "is_square (Q_column_add'_new z n m k) z"
proof -
have "ncols (Q_column_add'_new z n m k) = z"
by (metis Q_column_add'_new_nrows Q_column_add_new_inverse inverse_matrix_def left_inverse_matrix_dim m n n_not_m na xt1(8))
moreover have "nrows (Q_column_add'_new z n m k) = z"
by (metis Q_column_add'_new_nrows m n n_not_m na xt1(8))
ultimately show ?thesis unfolding is_square_def by simp
qed
qed


lemma column_add_matrix_tail_rec'_AQ:
assumes z:"z≥ncols A"
and not_zero: "Rep_matrix A k k ≠ 0"
and i: "k≤i"
shows "(∃Q. is_invertible Q ∧ (column_add_matrix_tail_rec' A i k k) = A*Q ∧ is_square Q z)"

proof -
obtain p where "ncols A = p" and nr: "ncols A ≤ p" by simp
show ?thesis
using z not_zero i nr
proof (induct i arbitrary: A rule: measure_induct_rule [of "(λk. p - k)"])
case (less i)
show ?case
proof (cases "ncols A ≤ i")
case True
show ?thesis
unfolding column_add_matrix_tail_rec'.simps [of A i]
apply (simp add: True)
by (metis is_square_def is_invertible_def less(2) ncols_one_matrix nrows_one_matrix one_matrix_inverse one_matrix_mult_right)
next
case False
hence i: "i<ncols A" by simp
show ?thesis
proof (cases "Suc i = ncols A")
case True
have 1: "column_add_matrix_tail_rec' A i k k = A"
proof -
have "Rep_matrix (column_add_matrix_tail_rec' A i k k) = Rep_matrix (column_add_matrix A (i + 1) k (- (Rep_matrix A k (i + 1) div Rep_matrix A k k)))"
using aux_ncols'[of A k k]
using column_add_matrix_tail_rec'.simps[of A i k k] using i True
by (metis False Suc_eq_plus1 column_add_matrix_tail_rec'.simps)
also have "... = Rep_matrix A"
proof (rule, rule)
fix a b
show " Rep_matrix (column_add_matrix A (i + 1) k (- (Rep_matrix A k (i + 1) div Rep_matrix A k k))) a b = Rep_matrix A a b"
proof (cases "b = i+1")
case False
show ?thesis
by (metis False Rep_matrix_column_add_matrix_preserve')
next
case True
have "Rep_matrix (column_add_matrix A (i + 1) k (- (Rep_matrix A k (i + 1) div Rep_matrix A k k))) a b
= Rep_matrix A a (i + 1) + - (Rep_matrix A k (i + 1) div Rep_matrix A k k) * Rep_matrix A a k"

unfolding column_add_matrix_def column_add_infmatrix_def unfolding Rep_aux_column_add using True by auto
also have "... = 0"
by (metis Suc_eq_plus1 `Suc i = ncols A` div_0 le_refl minus_mult_commute mult_zero_right ncols zadd_0_right zmult_commute)
also have "... = Rep_matrix A a b"
by (metis Suc_eq_plus1 True `Suc i = ncols A` eq_imp_le ncols)
finally show ?thesis .
qed
qed
finally show ?thesis using Rep_matrix_inject by auto
qed

show ?thesis
unfolding 1
proof (rule exI[of _ "one_matrix z"], auto)
show "is_invertible (one_matrix z)"
by (metis is_invertible_def one_matrix_inverse)
show "A = A*one_matrix z"
by (metis less(2) one_matrix_mult_right)
show "is_square (one_matrix z) z"
by (metis is_square_def ncols_one_matrix nrows_one_matrix)
qed

next
case False
def A' "(column_add_matrix A (i + 1) k (- (Rep_matrix A k (i + 1) div Rep_matrix A k k)))"
have "column_add_matrix_tail_rec' A i k k = column_add_matrix_tail_rec' A' (i+1) k k"
unfolding A'_def using column_add_matrix_tail_rec'.simps using i by simp
have "∃Q. is_invertible Q ∧ column_add_matrix_tail_rec' A' (i + 1) k k = A'*Q ∧ is_square Q z"
proof (unfold A'_def, rule less.hyps)
show "p - (i + 1) < p - i"
by (metis Suc_eq_plus1 diff_less_mono2 i less(5) lessI linorder_not_less xt1(7))
show " ncols (column_add_matrix A (i + 1) k (- (Rep_matrix A k (i + 1) div Rep_matrix A k k))) ≤ z"
proof -
have "ncols (column_add_matrix A (i + 1) k (- (Rep_matrix A k (i + 1) div Rep_matrix A k k))) ≤ ncols A "
by (metis False Suc_eq_plus1 Suc_lessI i ncols_column_add_matrix_le)
also have "... ≤ z" using less(2) .
finally show ?thesis .
qed
show " Rep_matrix (column_add_matrix A (i + 1) k (- (Rep_matrix A k (i + 1) div Rep_matrix A k k))) k k ≠ 0"
by (metis Rep_matrix_column_add_matrix_preserve less(3) less(4) less_add_one linorder_not_less)
show "k ≤ i + 1"
by (metis less(4) trans_le_add1)
show "ncols (column_add_matrix A (i + 1) k (- (Rep_matrix A k (i + 1) div Rep_matrix A k k))) ≤ p"
proof -
have "ncols (column_add_matrix A (i + 1) k (- (Rep_matrix A k (i + 1) div Rep_matrix A k k))) ≤ ncols A "
by (metis False Suc_eq_plus1 Suc_lessI i ncols_column_add_matrix_le)
also have "... ≤ p" using less(5) .
finally show ?thesis .
qed
qed
from this obtain Q where inv_Q: "is_invertible Q" and eq: "column_add_matrix_tail_rec' A' (i + 1) k k = A'*Q" and cuadrada: "is_square Q z"
by blast

have "∃Q'. is_invertible Q' ∧ A' = A*Q' ∧ is_square Q' z" --"Hago primero el existencial con el teorema para no complicar el obtain"
proof (unfold A'_def, rule column_add_new_AQ')
show "i + 1 < ncols A" using False i by simp
show "k < ncols A"
by (metis less(3) ncols_notzero)
show "ncols A ≤ z" using less(2) .
show "i + 1 ≠ k"
by (metis Suc_eq_plus1 Suc_n_not_le_n less(4))
qed
from this obtain Q' where inv_Q': "is_invertible Q'" and eq_Q': "A'= A*Q'" and cuadradaQ': "is_square Q' z" by blast
show ?thesis
proof (rule exI[of _ "Q'*Q"], auto)
show " is_invertible (Q' * Q)"
by (metis cuadrada cuadradaQ' is_square_def inv_Q inv_Q' mult_matrix_inverse2)
show "is_square (Q'* Q) z"
by (metis (no_types) cuadrada cuadradaQ' is_square_def inv_Q inv_Q' mult_matrix_inverse)
have "column_add_matrix_tail_rec' A i k k = column_add_matrix_tail_rec' A' (i+1) k k"
unfolding A'_def using column_add_matrix_tail_rec'.simps using i by simp
also have "... = A'*Q" using eq .
also have "... = (A*Q')*Q" unfolding eq_Q' ..
also have "... = A*(Q'*Q)" using mult_assoc by metis
finally show "column_add_matrix_tail_rec' A i k k = A*(Q'*Q)" .
qed
qed
qed
qed
qed

lemma Q_partColumnReduce:
assumes "t≥ncols A"
and "Rep_matrix A k k ≠ 0"
shows "(∃Q. is_invertible Q ∧ (partColumnReduce_matrix A k k) = A*Q ∧ is_square Q t)"

by (metis assms(1) assms(2) column_add_matrix_tail_rec'_AQ column_add_matrix_tail_rec'_eq_partColumnReduce_matrix eq_imp_le)


lemma nrows_move_minNonzero_nc:
assumes not_all_zero: "fst_bis (minNonzero_nc A k) = True"
shows "nrows (move_minNonzero_nc A k) ≤ nrows A"

proof -
have "∀z t. nrows A ≤ z ∧ ncols A ≤ t --> (∃P Q. is_invertible P ∧ is_invertible Q ∧ (move_minNonzero_nc A k)
= P * A * Q ∧ is_square P z ∧ is_square Q t)"
using move_minNonzero_PAQ[OF not_all_zero] .
from this obtain P' Q' where inv_P': "is_invertible P'" and inv_Q': "is_invertible Q'" and move: "(move_minNonzero_nc A k) = P' * A * Q'"
and cuadradaP': " is_square P' (nrows A)"
and cuadradaQ':"is_square Q' (ncols A)"
by force
have "nrows (move_minNonzero_nc A k) = nrows (P' * A * Q')" unfolding move ..
also have " ... ≤ nrows (P' * A)" by (metis nrows_mult)
also have "... ≤ nrows P'" by (metis nrows_mult)
also have "... = nrows A" using cuadradaP' unfolding is_square_def by simp
finally show ?thesis .
qed


lemma ncols_move_minNonzero_nc:
assumes not_all_zero: "fst_bis (minNonzero_nc A k) = True"
shows "ncols (move_minNonzero_nc A k) ≤ ncols A"

proof -
have "∀z t. nrows A ≤ z ∧ ncols A ≤ t --> (∃P Q. is_invertible P ∧ is_invertible Q ∧ (move_minNonzero_nc A k)
= P * A * Q ∧ is_square P z ∧ is_square Q t)"
using move_minNonzero_PAQ[OF not_all_zero] .
from this obtain P' Q' where inv_P': "is_invertible P'" and inv_Q': "is_invertible Q'" and move: "(move_minNonzero_nc A k) = P' * A * Q'"
and cuadradaP': " is_square P' (nrows A)"
and cuadradaQ':"is_square Q' (ncols A)"
by force
have "ncols (move_minNonzero_nc A k) = ncols (P' * A * Q')" unfolding move ..
also have "... = ncols (P' * (A * Q'))" using mult_assoc by metis
also have " ... ≤ ncols (A*Q')" by (metis ncols_mult)
also have "... ≤ ncols Q'" by (metis ncols_mult)
also have "... = ncols A" using cuadradaQ' unfolding is_square_def by simp
finally show ?thesis .
qed


lemma PQ_partColumn_iterate_move:
assumes z:"z≥nrows A"
and t:"t≥ncols A"
and not_all_zero: "fst_bis (minNonzero_nc A k) = True"
shows "∃P Q. is_invertible P ∧ is_invertible Q ∧ (partColumnReduce_matrix(iterate_partRowReduce(move_minNonzero_nc A k) k) k k)
= P * A * Q ∧ is_square P z ∧ is_square Q t"

proof -
have "∃P Q. is_invertible P ∧ is_invertible Q ∧ (iterate_partRowReduce(move_minNonzero_nc A k) k)
= P * (move_minNonzero_nc A k) * Q ∧ is_square P z ∧ is_square Q t"

proof (rule PQ_iterate_partRowReduce)
show "k < nrows (move_minNonzero_nc A k)"
by (metis assms(3) move_minNonzero_not_zero nrows_notzero)
show " nrows (move_minNonzero_nc A k) ≤ z"
by (metis not_all_zero nrows_move_minNonzero_nc order_trans z)
show " k < ncols (move_minNonzero_nc A k)"
by (metis move_minNonzero_not_zero ncols_notzero not_all_zero)
show "ncols (move_minNonzero_nc A k) ≤ t"
by (metis ncols_move_minNonzero_nc not_all_zero order_trans t)
qed
from this obtain P Q where inv_P: "is_invertible P" and inv_Q: "is_invertible Q" and iter: "(iterate_partRowReduce(move_minNonzero_nc A k) k)
= P * (move_minNonzero_nc A k) * Q"
and cuadradaP: "is_square P z" and cuadradaQ: "is_square Q t"
by auto
have "∃P Q. is_invertible P ∧ is_invertible Q ∧ (move_minNonzero_nc A k)
= P * A * Q ∧ is_square P z ∧ is_square Q t"

by (metis move_minNonzero_PAQ not_all_zero t z)
from this obtain P' Q' where inv_P': "is_invertible P'" and inv_Q': "is_invertible Q'" and move: "(move_minNonzero_nc A k) = P' * A * Q'"
and cuadradaP': " is_square P' z"
and cuadradaQ':"is_square Q' t"
by auto
have "(∃Q. is_invertible Q ∧ (partColumnReduce_matrix (iterate_partRowReduce(move_minNonzero_nc A k) k) k k) = (iterate_partRowReduce(move_minNonzero_nc A k) k)*Q
∧ is_square Q t)"

proof (rule Q_partColumnReduce)
show " ncols (iterate_partRowReduce (move_minNonzero_nc A k) k) ≤ t"
by (metis cuadradaQ is_square_def iter ncols_mult)
show "Rep_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k ≠ 0"
by (metis abs_zero iterate_nonzero_kk move_minNonzero_not_zero not_all_zero)
qed
from this obtain Q2 where invertibleQ2 :"is_invertible Q2" and part: "(partColumnReduce_matrix (iterate_partRowReduce(move_minNonzero_nc A k) k) k k)
= (iterate_partRowReduce(move_minNonzero_nc A k) k)*Q2 "

and cuadradaQ2: "is_square Q2 t"
by auto
show ?thesis
proof (rule exI[of _ "P*P'"], rule exI[of _ "Q'*Q*Q2"], auto)
show "is_invertible (P * P')"
by (metis cuadradaP cuadradaP' is_square_def inv_P inv_P' mult_matrix_inverse)
show "is_invertible (Q' * Q * Q2)"
by (metis cuadradaQ cuadradaQ' cuadradaQ2 is_square_def inv_Q inv_Q' invertibleQ2 mult_matrix_inverse)
show "partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k = P * P' * A * (Q' * Q * Q2)"
proof -
have "(partColumnReduce_matrix (iterate_partRowReduce(move_minNonzero_nc A k) k) k k) =(iterate_partRowReduce(move_minNonzero_nc A k) k)*Q2 " using part .
also have "... = P*(move_minNonzero_nc A k)*Q*Q2" unfolding iter ..
also have "... = P* (P'*A *Q')*Q*Q2" unfolding move ..
also have "... = P* (P'*A *Q'*Q*Q2)" using mult_assoc by metis
also have "... = P* (P'*A *(Q'*Q*Q2))" using mult_assoc by metis
also have "...=P * P' * A * (Q' * Q * Q2)" using mult_assoc by metis
finally show ?thesis .
qed
show "is_square (P * P') z"
by (metis (no_types) cuadradaP cuadradaP' is_square_def inv_P inv_P' mult_matrix_inverse)
show "is_square (Q' * Q * Q2) t"
by (metis (no_types) cuadradaQ cuadradaQ' cuadradaQ2 is_square_def inv_Q inv_Q' invertibleQ2 mult_matrix_inverse)
qed
qed


lemma PQ_Diagonalize:
assumes z: "z≥nrows A"
and t: "t≥ncols A"
(*and k: "k<nrows A"
and k2: "k<ncols A"*)

shows "∃P Q. is_invertible P ∧ is_invertible Q ∧ Diagonalize A k = P*A*Q ∧ is_square P z ∧ is_square Q t"

using z t (*k k2*)
proof (induct "(A,k)" arbitrary: A k rule: measure_induct_rule [of f'])
fix A k assume (*h1:"!!Aa ka. [|f' (Aa, ka) < f' (A, k); nrows Aa ≤ z; ncols Aa ≤ t; ka < nrows Aa; ka < ncols Aa|]
==> ∃P Q. is_invertible P ∧ is_invertible Q ∧ Diagonalize Aa ka = P * Aa * Q ∧ is_square P z ∧ is_square Q t" *)

h1: "!!Aa ka. [|f' (Aa, ka) < f' (A, k); nrows Aa ≤ z; ncols Aa ≤ t|]
==> ∃P Q. is_invertible P ∧ is_invertible Q ∧ Diagonalize Aa ka = P * Aa * Q ∧ is_square P z ∧ is_square Q t"

and z: "nrows A ≤ z" and t:"ncols A ≤ t"
(*and k:"k < nrows A" and k2:"k < ncols A"*)
show "∃P Q. is_invertible P ∧ is_invertible Q ∧ Diagonalize A k = P*A*Q ∧ is_square P z ∧ is_square Q t"
proof (cases "(∀m>k. Rep_matrix A m k = 0) ∧ (∀n>k. Rep_matrix A k n = 0)")
case True
have "Diagonalize A k = A" using Diagonalize.simps[of A k] using True by auto
thus ?thesis
by (metis is_square_def is_invertible_def ncols_one_matrix nrows_one_matrix one_matrix_inverse one_matrix_mult_left one_matrix_mult_right t z)
next
case False
have not_all_zeroA: "fst_bis (minNonzero_nc A k) = True"
proof (rule minNonzero_true, cases "(∀m>k. Rep_matrix A m k = 0)")
case True from this obtain i where i: "i>k" and r:"(Rep_matrix A k i ≠ 0)" using False by auto
hence "i<ncols A" and "k<nrows A" using ncols_notzero[of A k i] nrows_notzero[of A k i] by auto
thus "¬ (∀i j. i ∈ {k..<nrows A} ∧ j ∈ {k..<ncols A} --> Rep_matrix A i j = 0)" using r i by auto
next
case False from this obtain i where i: "i>k" and r:"(Rep_matrix A i k ≠ 0)" using False by auto
hence "i<nrows A" and "k<ncols A" using ncols_notzero[of A i k] nrows_notzero[of A i k] by auto
thus "¬ (∀i j. i ∈ {k..<nrows A} ∧ j ∈ {k..<ncols A} --> Rep_matrix A i j = 0)" using r i by auto
qed
note False_1=False
have 1: "∃P Q. is_invertible P ∧ is_invertible Q ∧ (partColumnReduce_matrix(iterate_partRowReduce(move_minNonzero_nc A k) k) k k)
= P * A * Q ∧ is_square P z ∧ is_square Q t"
using PQ_partColumn_iterate_move[OF z t not_all_zeroA] .
from this obtain P Q where invP: "is_invertible P" and invQ: "is_invertible Q" and eq: "(partColumnReduce_matrix(iterate_partRowReduce(move_minNonzero_nc A k) k) k k)
= P * A * Q"
and cuadradaP: "is_square P z" and cuadradaQ: "is_square Q t"
by auto
show ?thesis
proof (cases "(∀m>k. Rep_matrix (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) m k= 0)
∧ (∀n>k. Rep_matrix (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) k n = 0)"
)

case True
have "Diagonalize A k = (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k)" using Diagonalize.simps[of A k] using False_1
using Diagonalize.simps[of "(partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k)" k] using True by simp
thus ?thesis using 1 by auto
next
case False
have not_all_row_zero: "¬ (∀n>k. Rep_matrix (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) k n = 0)"
proof (rule ccontr)
assume "¬ ¬ (∀n>k. Rep_matrix (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) k n = 0)"
hence "¬ (∀m>k. Rep_matrix (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) m k= 0)" using False by auto
from this obtain i where i:"i>k" and nz_i: "Rep_matrix (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) i k ≠ 0"
by blast
have "Rep_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) i k
= Rep_matrix (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) i k"

by (metis order_eq_refl partColumnReduce')
also have "...≠0" using nz_i by simp
finally show False using iterate_partRowReduce[of "(move_minNonzero_nc A k)" k] i
by (metis move_minNonzero_not_zero not_all_zeroA)
qed
from this obtain i where i: "i>k" and nz_i: "Rep_matrix (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) k i ≠ 0"
by blast
have "∃P Q. is_invertible P ∧ is_invertible Q ∧ Diagonalize (partColumnReduce_matrix(iterate_partRowReduce(move_minNonzero_nc A k) k) k k) k
= P * (partColumnReduce_matrix(iterate_partRowReduce(move_minNonzero_nc A k) k) k k) * Q ∧ is_square P z ∧ is_square Q t"

proof (rule h1, unfold f'_def, auto)
have "Rep_matrix (move_minNonzero_nc A k) k k ≠ 0"
by (metis move_minNonzero_not_zero not_all_zeroA)
thus "Rep_matrix (move_minNonzero_nc A k) k k = 0 ==> False" by contradiction
show "nrows (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) ≤ z"
by (metis cuadradaP eq is_square_def nrows_mult order_trans)
show "ncols (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) ≤ t"
by (metis cuadradaQ eq is_square_def ncols_mult)
show " ¦Rep_matrix (move_minNonzero_nc (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) k) k k¦
< ¦Rep_matrix (move_minNonzero_nc A k) k k¦"

proof -
have " ¦Rep_matrix (move_minNonzero_nc (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) k) k k¦
< ¦Rep_matrix (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) k k¦"

proof (rule Akk_strict_decreasing_move_minNonzero2)
show "k ≤ k" by simp
show "k < i" using i .
show "¦Rep_matrix (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) k i¦
< ¦Rep_matrix (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) k k¦"

by (metis `Rep_matrix (move_minNonzero_nc A k) k k = 0 ==> False` abs_zero iterate_nonzero_kk i partColumnReduce_abs partColumnReduce_eq_abs)
show "Rep_matrix (partColumnReduce_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k) k i ≠ 0" using nz_i .
qed
also have "...=¦Rep_matrix (iterate_partRowReduce (move_minNonzero_nc A k) k) k k¦" by (metis partColumnReduce_eq_abs)
also have "...≤ ¦Rep_matrix (move_minNonzero_nc A k) k k¦" by (metis `Rep_matrix (move_minNonzero_nc A k) k k = 0 ==> False` iterate_decrease_kk')
finally show ?thesis .
qed
qed
from this obtain P' Q' where inv_P': "is_invertible P'" and inv_Q': "is_invertible Q'"
and hermite:"Diagonalize (partColumnReduce_matrix(iterate_partRowReduce(move_minNonzero_nc A k) k) k k) k
= P' * (partColumnReduce_matrix(iterate_partRowReduce(move_minNonzero_nc A k) k) k k) * Q'"
and cuadradaP': "is_square P' z" and cuadradaQ': "is_square Q' t"

by auto
show ?thesis
proof (rule exI[of _ "P'*P"], rule exI[of _ "Q*Q'"], auto)
show "is_invertible (P' * P)"
by (metis cuadradaP cuadradaP' is_square_def invP inv_P' mult_matrix_inverse)
show "is_invertible (Q * Q')"
by (metis cuadradaQ cuadradaQ' is_square_def invQ inv_Q' mult_matrix_inverse)
show "Diagonalize A k = P' * P * A * (Q * Q')"
proof -
have "Diagonalize A k = Diagonalize (partColumnReduce_matrix(iterate_partRowReduce(move_minNonzero_nc A k) k) k k) k" using Diagonalize.simps[of A k] using False_1 by auto
also have "... = P' * (partColumnReduce_matrix(iterate_partRowReduce(move_minNonzero_nc A k) k) k k) * Q'" using hermite by auto
also have "... = P' * (P * A * Q) * Q'" unfolding eq ..
also have "... = P' * P * A * (Q * Q')" using mult_assoc by metis
finally show ?thesis .
qed
show "is_square (P' * P) z"
by (metis (no_types) cuadradaP cuadradaP' is_square_def invP inv_P' mult_matrix_inverse)
show "is_square (Q * Q') t"
by (metis (no_types) cuadradaQ cuadradaQ' is_square_def invQ inv_Q' mult_matrix_inverse)
qed
qed
qed
qed



lemma PQ_Diagonalize_up_to_k_not_null:
assumes z:"z≥nrows A"
and t:"t≥ncols A"
and nr:"nrows A > 0"
and nc: "ncols A > 0" --"Realmente me sobra, a partir de que nrows A > 0 ya puedo sacar que ncols A > 0"
shows "∃P Q. is_invertible P ∧ is_invertible Q ∧ Diagonalize_up_to_k (A::int matrix) n = P*A*Q ∧ is_square P z ∧ is_square Q t"

proof (induct n)
show "∃P Q. is_invertible P ∧ is_invertible Q ∧ Diagonalize_up_to_k A 0 = P * A * Q ∧ is_square P z ∧ is_square Q t"
proof -
have " ∃P Q. is_invertible P ∧ is_invertible Q ∧ Diagonalize A 0 = P * A * Q ∧ is_square P z ∧ is_square Q t" using PQ_Diagonalize[OF z t] .
thus ?thesis using Diagonalize_up_to_k.simps by simp
qed
next
fix n
assume hip: "∃P Q. is_invertible P ∧ is_invertible Q ∧ Diagonalize_up_to_k A n = P * A * Q ∧ is_square P z ∧ is_square Q t"
from this obtain P' Q' where invertible_P': "is_invertible P'" and invertible_Q': "is_invertible Q'" and h':"Diagonalize_up_to_k A n = P' * A * Q'"
and cuadradaP': "is_square P' z" and cuadradaQ': "is_square Q' t"
by auto
have "Diagonalize_up_to_k A (Suc n) = Diagonalize (Diagonalize_up_to_k A n) (Suc n)" using Diagonalize_up_to_k.simps by auto
have "∃P Q. is_invertible P ∧ is_invertible Q ∧ Diagonalize(Diagonalize_up_to_k A n) (Suc n) = P * (Diagonalize_up_to_k A n) * Q ∧ is_square P z ∧ is_square Q t"
thm Diagonalize.simps thm PQ_Diagonalize

proof (rule PQ_Diagonalize)
show " nrows (Diagonalize_up_to_k A n) ≤ z"
by (metis is_square_def hip nrows_mult xt1(6))
show "ncols (Diagonalize_up_to_k A n) ≤ t"
by (metis is_square_def hip ncols_mult)
qed
from this obtain P Q where invertible_P: "is_invertible P" and invertible_Q: "is_invertible Q" and h:"Diagonalize(Diagonalize_up_to_k A n) (Suc n) = P * (Diagonalize_up_to_k A n) * Q"
and cuadradaP: "is_square P z" and cuadradaQ: "is_square Q t"
by auto
show "∃P Q. is_invertible P ∧ is_invertible Q ∧ Diagonalize_up_to_k A (Suc n) = P * A * Q ∧ is_square P z ∧ is_square Q t"
proof (rule exI[of _ "P*P'"], rule exI[of _ "Q'*Q"], auto)
show "Diagonalize (Diagonalize_up_to_k A n) (Suc n) = P * P' * A * (Q' * Q)" unfolding h unfolding h' using mult_assoc by metis
show "is_invertible (P * P')"
by (metis cuadradaP cuadradaP' is_square_def invertible_P invertible_P' mult_matrix_inverse)
show "is_invertible (Q' * Q)" by (metis cuadradaQ cuadradaQ' is_square_def invertible_Q invertible_Q' mult_matrix_inverse)
show "is_square (P * P') z"
by (metis (no_types) cuadradaP cuadradaP' is_square_def invertible_P invertible_P' mult_matrix_inverse)
show "is_square (Q' * Q) t"
by (metis (no_types) cuadradaQ cuadradaQ' is_square_def invertible_Q invertible_Q' mult_matrix_inverse)
qed
qed


lemma Diagonalize_theorem:
shows "∃P Q B. is_invertible P ∧ is_invertible Q ∧ B = P*A*Q ∧ is_square P (nrows (A::int matrix)) ∧ is_square Q (ncols A) ∧ Diagonalize_p B (max (nrows A) (ncols A))"

proof (cases "nrows A = 0")
case True
have nc: "ncols A = 0" using True by (metis Rep_matrix_zero_imp_mult_zero le0 nrows nrows_transpose one_matrix_mult_left transpose_matrix_zero)
hence 1: "(max (nrows A) (ncols A)) = 0" using True by auto
have 2: "Diagonalize_p A (max (nrows A) (ncols A))" unfolding 1
by (metis Diagonalize_up_to_k.simps(1) Diagonalize_p_Diagonalize_n PQ_Diagonalize Rep_matrix_zero_imp_mult_zero True le0 nrows one_matrix_mult_left zero_matrix_def_ncols)
show ?thesis
proof (rule exI[of _ "(one_matrix 0)::int matrix"], rule exI[of _ "(one_matrix 0)::int matrix"], rule exI[of _ A], auto)
show "is_invertible (one_matrix 0)" using one_matrix_inverse unfolding is_invertible_def by auto
show "A = one_matrix (0) * A * one_matrix (0)"
by (metis True `ncols A = 0` less_eq_nat.simps(1) one_matrix_mult_left one_matrix_mult_right)
show "is_square (one_matrix 0) (nrows A)" using True nc unfolding is_square_def by auto
show "is_square (one_matrix 0) (ncols A)" using True nc unfolding is_square_def by auto
show "Diagonalize_p A (max (nrows A) (ncols A))" using 2 .
qed
next
case False
have r:"nrows A > 0" using False by auto
have c:"ncols A >0"
proof -
obtain j k where "Rep_matrix A j k ≠ 0" and "j<nrows A" using r
by (metis less_nrows order_refl)
thus ?thesis
by (metis less_eq_nat.simps(1) less_ncols)
qed
show ?thesis using PQ_Diagonalize_up_to_k_not_null[of A "nrows A" "ncols A" "(max (nrows A) (ncols A))"] using r c
by (metis Diagonalize_p_Diagonalize_n order_refl)
qed

end