| author | wenzelm | 
| Wed, 11 Mar 2009 20:36:20 +0100 | |
| changeset 30458 | 804de935c328 | 
| parent 30267 | 171b3bd93c90 | 
| child 30489 | 5d7d0add1741 | 
| permissions | -rw-r--r-- | 
| 29846 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1 | (* Title: Determinants | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 2 | Author: Amine Chaieb, University of Cambridge | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 3 | *) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 4 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 5 | header {* Traces, Determinant of square matrices and some properties *}
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 6 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 7 | theory Determinants | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 8 | imports Euclidean_Space Permutations | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 9 | begin | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 10 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 11 | subsection{* First some facts about products*}
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 12 | lemma setprod_insert_eq: "finite A \<Longrightarrow> setprod f (insert a A) = (if a \<in> A then setprod f A else f a * setprod f A)" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 13 | apply clarsimp | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 14 | by(subgoal_tac "insert a A = A", auto) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 15 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 16 | lemma setprod_add_split: | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 17 | assumes mn: "(m::nat) <= n + 1" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 18 |   shows "setprod f {m.. n+p} = setprod f {m .. n} * setprod f {n+1..n+p}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 19 | proof- | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 20 |   let ?A = "{m .. n+p}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 21 |   let ?B = "{m .. n}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 22 |   let ?C = "{n+1..n+p}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 23 | from mn have un: "?B \<union> ?C = ?A" by auto | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 24 |   from mn have dj: "?B \<inter> ?C = {}" by auto
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 25 | have f: "finite ?B" "finite ?C" by simp_all | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 26 | from setprod_Un_disjoint[OF f dj, of f, unfolded un] show ?thesis . | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 27 | qed | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 28 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 29 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 30 | lemma setprod_offset: "setprod f {(m::nat) + p .. n + p} = setprod (\<lambda>i. f (i + p)) {m..n}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 31 | apply (rule setprod_reindex_cong[where f="op + p"]) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 32 | apply (auto simp add: image_iff Bex_def inj_on_def) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 33 | apply arith | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 34 | apply (rule ext) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 35 | apply (simp add: add_commute) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 36 | done | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 37 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 38 | lemma setprod_singleton: "setprod f {x} = f x" by simp
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 39 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 40 | lemma setprod_singleton_nat_seg: "setprod f {n..n} = f (n::'a::order)" by simp
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 41 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 42 | lemma setprod_numseg: "setprod f {m..0} = (if m=0 then f 0 else 1)"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 43 |   "setprod f {m .. Suc n} = (if m \<le> Suc n then f (Suc n) * setprod f {m..n} 
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 44 |                              else setprod f {m..n})"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 45 | by (auto simp add: atLeastAtMostSuc_conv) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 46 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 47 | lemma setprod_le: assumes fS: "finite S" and fg: "\<forall>x\<in>S. f x \<ge> 0 \<and> f x \<le> (g x :: 'a::ordered_idom)" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 48 | shows "setprod f S \<le> setprod g S" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 49 | using fS fg | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 50 | apply(induct S) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 51 | apply simp | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 52 | apply auto | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 53 | apply (rule mult_mono) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 54 | apply (auto intro: setprod_nonneg) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 55 | done | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 56 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 57 | (* FIXME: In Finite_Set there is a useless further assumption *) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 58 | lemma setprod_inversef: "finite A ==> setprod (inverse \<circ> f) A = (inverse (setprod f A) :: 'a:: {division_by_zero, field})"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 59 | apply (erule finite_induct) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 60 | apply (simp) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 61 | apply simp | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 62 | done | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 63 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 64 | lemma setprod_le_1: assumes fS: "finite S" and f: "\<forall>x\<in>S. f x \<ge> 0 \<and> f x \<le> (1::'a::ordered_idom)" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 65 | shows "setprod f S \<le> 1" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 66 | using setprod_le[OF fS f] unfolding setprod_1 . | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 67 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 68 | subsection{* Trace *}
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 69 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 70 | definition trace :: "'a::semiring_1^'n^'n \<Rightarrow> 'a" where | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 71 |   "trace A = setsum (\<lambda>i. ((A$i)$i)) {1..dimindex(UNIV::'n set)}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 72 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 73 | lemma trace_0: "trace(mat 0) = 0" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 74 | by (simp add: trace_def mat_def Cart_lambda_beta setsum_0) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 75 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 76 | lemma trace_I: "trace(mat 1 :: 'a::semiring_1^'n^'n) = of_nat(dimindex(UNIV::'n set))" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 77 | by (simp add: trace_def mat_def Cart_lambda_beta) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 78 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 79 | lemma trace_add: "trace ((A::'a::comm_semiring_1^'n^'n) + B) = trace A + trace B" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 80 | by (simp add: trace_def setsum_addf Cart_lambda_beta vector_component) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 81 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 82 | lemma trace_sub: "trace ((A::'a::comm_ring_1^'n^'n) - B) = trace A - trace B" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 83 | by (simp add: trace_def setsum_subtractf Cart_lambda_beta vector_component) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 84 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 85 | lemma trace_mul_sym:"trace ((A::'a::comm_semiring_1^'n^'n) ** B) = trace (B**A)" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 86 | apply (simp add: trace_def matrix_matrix_mult_def Cart_lambda_beta) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 87 | apply (subst setsum_commute) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 88 | by (simp add: mult_commute) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 89 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 90 | (* ------------------------------------------------------------------------- *) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 91 | (* Definition of determinant. *) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 92 | (* ------------------------------------------------------------------------- *) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 93 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 94 | definition det:: "'a::comm_ring_1^'n^'n \<Rightarrow> 'a" where | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 95 |   "det A = setsum (\<lambda>p. of_int (sign p) * setprod (\<lambda>i. A$i$p i) {1 .. dimindex(UNIV :: 'n set)}) {p. p permutes {1 .. dimindex(UNIV :: 'n set)}}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 96 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 97 | (* ------------------------------------------------------------------------- *) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 98 | (* A few general lemmas we need below. *) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 99 | (* ------------------------------------------------------------------------- *) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 100 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 101 | lemma Cart_lambda_beta_perm: assumes p: "p permutes {1..dimindex(UNIV::'n set)}" 
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 102 |   and i: "i \<in> {1..dimindex(UNIV::'n set)}" 
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 103 | shows "Cart_nth (Cart_lambda g ::'a^'n) (p i) = g(p i)" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 104 | using permutes_in_image[OF p] i | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 105 | by (simp add: Cart_lambda_beta permutes_in_image[OF p]) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 106 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 107 | lemma setprod_permute: | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 108 | assumes p: "p permutes S" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 109 | shows "setprod f S = setprod (f o p) S" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 110 | proof- | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 111 |   {assume "\<not> finite S" hence ?thesis by simp}
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 112 | moreover | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 113 |   {assume fS: "finite S"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 114 | then have ?thesis | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 115 | apply (simp add: setprod_def) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 116 | apply (rule ab_semigroup_mult.fold_image_permute) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 117 | apply (auto simp add: p) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 118 | apply unfold_locales | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 119 | done} | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 120 | ultimately show ?thesis by blast | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 121 | qed | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 122 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 123 | lemma setproduct_permute_nat_interval: "p permutes {m::nat .. n} ==> setprod f {m..n} = setprod (f o p) {m..n}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 124 | by (auto intro: setprod_permute) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 125 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 126 | (* ------------------------------------------------------------------------- *) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 127 | (* Basic determinant properties. *) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 128 | (* ------------------------------------------------------------------------- *) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 129 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 130 | lemma det_transp: "det (transp A) = det (A::'a::comm_ring_1 ^'n^'n)" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 131 | proof- | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 132 | let ?di = "\<lambda>A i j. A$i$j" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 133 |   let ?U = "{1 .. dimindex (UNIV :: 'n set)}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 134 | have fU: "finite ?U" by blast | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 135 |   {fix p assume p: "p \<in> {p. p permutes ?U}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 136 | from p have pU: "p permutes ?U" by blast | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 137 | have sth: "sign (inv p) = sign p" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 138 | by (metis sign_inverse fU p mem_def Collect_def permutation_permutes) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 139 | from permutes_inj[OF pU] | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 140 | have pi: "inj_on p ?U" by (blast intro: subset_inj_on) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 141 | from permutes_image[OF pU] | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 142 | have "setprod (\<lambda>i. ?di (transp A) i (inv p i)) ?U = setprod (\<lambda>i. ?di (transp A) i (inv p i)) (p ` ?U)" by simp | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 143 | also have "\<dots> = setprod ((\<lambda>i. ?di (transp A) i (inv p i)) o p) ?U" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 144 | unfolding setprod_reindex[OF pi] .. | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 145 | also have "\<dots> = setprod (\<lambda>i. ?di A i (p i)) ?U" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 146 | proof- | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 147 |       {fix i assume i: "i \<in> ?U"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 148 | from i permutes_inv_o[OF pU] permutes_in_image[OF pU] | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 149 | have "((\<lambda>i. ?di (transp A) i (inv p i)) o p) i = ?di A i (p i)" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 150 | unfolding transp_def by (simp add: Cart_lambda_beta expand_fun_eq)} | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 151 | then show "setprod ((\<lambda>i. ?di (transp A) i (inv p i)) o p) ?U = setprod (\<lambda>i. ?di A i (p i)) ?U" by (auto intro: setprod_cong) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 152 | qed | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 153 | finally have "of_int (sign (inv p)) * (setprod (\<lambda>i. ?di (transp A) i (inv p i)) ?U) = of_int (sign p) * (setprod (\<lambda>i. ?di A i (p i)) ?U)" using sth | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 154 | by simp} | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 155 | then show ?thesis unfolding det_def apply (subst setsum_permutations_inverse) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 156 | apply (rule setsum_cong2) by blast | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 157 | qed | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 158 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 159 | lemma det_lowerdiagonal: | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 160 | fixes A :: "'a::comm_ring_1^'n^'n" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 161 |   assumes ld: "\<And>i j. i \<in> {1 .. dimindex (UNIV:: 'n set)} \<Longrightarrow> j \<in> {1 .. dimindex(UNIV:: 'n set)} \<Longrightarrow> i < j \<Longrightarrow> A$i$j = 0"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 162 |   shows "det A = setprod (\<lambda>i. A$i$i) {1..dimindex(UNIV:: 'n set)}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 163 | proof- | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 164 |   let ?U = "{1..dimindex(UNIV:: 'n set)}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 165 |   let ?PU = "{p. p permutes ?U}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 166 |   let ?pp = "\<lambda>p. of_int (sign p) * setprod (\<lambda>i. A$i$p i) {1 .. dimindex(UNIV :: 'n set)}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 167 | have fU: "finite ?U" by blast | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 168 | from finite_permutations[OF fU] have fPU: "finite ?PU" . | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 169 |   have id0: "{id} \<subseteq> ?PU" by (auto simp add: permutes_id)
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 170 |   {fix p assume p: "p \<in> ?PU -{id}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 171 | from p have pU: "p permutes ?U" and pid: "p \<noteq> id" by blast+ | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 172 | from permutes_natset_le[OF pU] pid obtain i where | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 173 | i: "i \<in> ?U" "p i > i" by (metis not_le) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 174 | from permutes_in_image[OF pU] i(1) have piU: "p i \<in> ?U" by blast | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 175 | from ld[OF i(1) piU i(2)] i(1) have ex:"\<exists>i \<in> ?U. A$i$p i = 0" by blast | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 176 | from setprod_zero[OF fU ex] have "?pp p = 0" by simp} | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 177 |   then have p0: "\<forall>p \<in> ?PU -{id}. ?pp p = 0"  by blast
 | 
| 30259 | 178 | from setsum_mono_zero_cong_left[OF fPU id0 p0] show ?thesis | 
| 29846 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 179 | unfolding det_def by (simp add: sign_id) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 180 | qed | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 181 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 182 | lemma det_upperdiagonal: | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 183 | fixes A :: "'a::comm_ring_1^'n^'n" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 184 |   assumes ld: "\<And>i j. i \<in> {1 .. dimindex (UNIV:: 'n set)} \<Longrightarrow> j \<in> {1 .. dimindex(UNIV:: 'n set)} \<Longrightarrow> i > j \<Longrightarrow> A$i$j = 0"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 185 |   shows "det A = setprod (\<lambda>i. A$i$i) {1..dimindex(UNIV:: 'n set)}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 186 | proof- | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 187 |   let ?U = "{1..dimindex(UNIV:: 'n set)}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 188 |   let ?PU = "{p. p permutes ?U}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 189 |   let ?pp = "(\<lambda>p. of_int (sign p) * setprod (\<lambda>i. A$i$p i) {1 .. dimindex(UNIV :: 'n set)})"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 190 | have fU: "finite ?U" by blast | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 191 | from finite_permutations[OF fU] have fPU: "finite ?PU" . | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 192 |   have id0: "{id} \<subseteq> ?PU" by (auto simp add: permutes_id)
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 193 |   {fix p assume p: "p \<in> ?PU -{id}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 194 | from p have pU: "p permutes ?U" and pid: "p \<noteq> id" by blast+ | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 195 | from permutes_natset_ge[OF pU] pid obtain i where | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 196 | i: "i \<in> ?U" "p i < i" by (metis not_le) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 197 | from permutes_in_image[OF pU] i(1) have piU: "p i \<in> ?U" by blast | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 198 | from ld[OF i(1) piU i(2)] i(1) have ex:"\<exists>i \<in> ?U. A$i$p i = 0" by blast | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 199 | from setprod_zero[OF fU ex] have "?pp p = 0" by simp} | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 200 |   then have p0: "\<forall>p \<in> ?PU -{id}. ?pp p = 0"  by blast
 | 
| 30259 | 201 | from setsum_mono_zero_cong_left[OF fPU id0 p0] show ?thesis | 
| 29846 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 202 | unfolding det_def by (simp add: sign_id) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 203 | qed | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 204 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 205 | lemma det_I: "det (mat 1 :: 'a::comm_ring_1^'n^'n) = 1" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 206 | proof- | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 207 | let ?A = "mat 1 :: 'a::comm_ring_1^'n^'n" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 208 |   let ?U = "{1 .. dimindex (UNIV :: 'n set)}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 209 | let ?f = "\<lambda>i j. ?A$i$j" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 210 |   {fix i assume i: "i \<in> ?U"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 211 | have "?f i i = 1" using i by (vector mat_def)} | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 212 | hence th: "setprod (\<lambda>i. ?f i i) ?U = setprod (\<lambda>x. 1) ?U" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 213 | by (auto intro: setprod_cong) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 214 |   {fix i j assume i: "i \<in> ?U" and j: "j \<in> ?U" and ij: "i < j"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 215 | have "?f i j = 0" using i j ij by (vector mat_def) } | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 216 | then have "det ?A = setprod (\<lambda>i. ?f i i) ?U" using det_lowerdiagonal | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 217 | by blast | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 218 | also have "\<dots> = 1" unfolding th setprod_1 .. | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 219 | finally show ?thesis . | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 220 | qed | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 221 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 222 | lemma det_0: "det (mat 0 :: 'a::comm_ring_1^'n^'n) = 0" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 223 | proof- | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 224 | let ?A = "mat 0 :: 'a::comm_ring_1^'n^'n" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 225 |   let ?U = "{1 .. dimindex (UNIV :: 'n set)}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 226 | let ?f = "\<lambda>i j. ?A$i$j" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 227 | have th:"setprod (\<lambda>i. ?f i i) ?U = 0" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 228 | apply (rule setprod_zero) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 229 | apply simp | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 230 | apply (rule bexI[where x=1]) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 231 | using dimindex_ge_1[of "UNIV :: 'n set"] | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 232 | by (simp_all add: mat_def Cart_lambda_beta) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 233 |   {fix i j assume i: "i \<in> ?U" and j: "j \<in> ?U" and ij: "i < j"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 234 | have "?f i j = 0" using i j ij by (vector mat_def) } | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 235 | then have "det ?A = setprod (\<lambda>i. ?f i i) ?U" using det_lowerdiagonal | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 236 | by blast | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 237 | also have "\<dots> = 0" unfolding th .. | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 238 | finally show ?thesis . | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 239 | qed | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 240 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 241 | lemma det_permute_rows: | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 242 | fixes A :: "'a::comm_ring_1^'n^'n" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 243 |   assumes p: "p permutes {1 .. dimindex (UNIV :: 'n set)}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 244 | shows "det(\<chi> i. A$p i :: 'a^'n^'n) = of_int (sign p) * det A" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 245 | apply (simp add: det_def setsum_right_distrib mult_assoc[symmetric] del: One_nat_def) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 246 | apply (subst sum_permutations_compose_right[OF p]) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 247 | proof(rule setsum_cong2) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 248 |   let ?U = "{1 .. dimindex (UNIV :: 'n set)}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 249 |   let ?PU = "{p. p permutes ?U}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 250 | let ?Ap = "(\<chi> i. A$p i :: 'a^'n^'n)" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 251 | fix q assume qPU: "q \<in> ?PU" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 252 | have fU: "finite ?U" by blast | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 253 | from qPU have q: "q permutes ?U" by blast | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 254 | from p q have pp: "permutation p" and qp: "permutation q" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 255 | by (metis fU permutation_permutes)+ | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 256 | from permutes_inv[OF p] have ip: "inv p permutes ?U" . | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 257 |     {fix i assume i: "i \<in> ?U"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 258 | from Cart_lambda_beta[rule_format, OF i, of "\<lambda>i. A$ p i"] | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 259 | have "?Ap$i$ (q o p) i = A $ p i $ (q o p) i " by simp} | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 260 | hence "setprod (\<lambda>i. ?Ap$i$ (q o p) i) ?U = setprod (\<lambda>i. A$p i$(q o p) i) ?U" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 261 | by (auto intro: setprod_cong) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 262 | also have "\<dots> = setprod ((\<lambda>i. A$p i$(q o p) i) o inv p) ?U" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 263 | by (simp only: setprod_permute[OF ip, symmetric]) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 264 | also have "\<dots> = setprod (\<lambda>i. A $ (p o inv p) i $ (q o (p o inv p)) i) ?U" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 265 | by (simp only: o_def) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 266 | also have "\<dots> = setprod (\<lambda>i. A$i$q i) ?U" by (simp only: o_def permutes_inverses[OF p]) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 267 | finally have thp: "setprod (\<lambda>i. ?Ap$i$ (q o p) i) ?U = setprod (\<lambda>i. A$i$q i) ?U" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 268 | by blast | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 269 | show "of_int (sign (q o p)) * setprod (\<lambda>i. ?Ap$i$ (q o p) i) ?U = of_int (sign p) * of_int (sign q) * setprod (\<lambda>i. A$i$q i) ?U" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 270 | by (simp only: thp sign_compose[OF qp pp] mult_commute of_int_mult) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 271 | qed | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 272 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 273 | lemma det_permute_columns: | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 274 | fixes A :: "'a::comm_ring_1^'n^'n" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 275 |   assumes p: "p permutes {1 .. dimindex (UNIV :: 'n set)}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 276 | shows "det(\<chi> i j. A$i$ p j :: 'a^'n^'n) = of_int (sign p) * det A" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 277 | proof- | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 278 | let ?Ap = "\<chi> i j. A$i$ p j :: 'a^'n^'n" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 279 | let ?At = "transp A" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 280 | have "of_int (sign p) * det A = det (transp (\<chi> i. transp A $ p i))" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 281 | unfolding det_permute_rows[OF p, of ?At] det_transp .. | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 282 | moreover | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 283 | have "?Ap = transp (\<chi> i. transp A $ p i)" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 284 | by (simp add: transp_def Cart_eq Cart_lambda_beta Cart_lambda_beta_perm[OF p]) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 285 | ultimately show ?thesis by simp | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 286 | qed | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 287 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 288 | lemma det_identical_rows: | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 289 | fixes A :: "'a::ordered_idom^'n^'n" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 290 |   assumes i: "i\<in>{1 .. dimindex (UNIV :: 'n set)}" 
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 291 |   and j: "j\<in>{1 .. dimindex (UNIV :: 'n set)}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 292 | and ij: "i \<noteq> j" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 293 | and r: "row i A = row j A" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 294 | shows "det A = 0" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 295 | proof- | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 296 | have tha: "\<And>(a::'a) b. a = b ==> b = - a ==> a = 0" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 297 | by simp | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 298 | have th1: "of_int (-1) = - 1" by (metis of_int_1 of_int_minus number_of_Min) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 299 | let ?p = "Fun.swap i j id" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 300 | let ?A = "\<chi> i. A $ ?p i" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 301 | from r have "A = ?A" by (simp add: Cart_eq Cart_lambda_beta Cart_lambda_beta_perm[OF permutes_swap_id[OF i j]] row_def swap_def) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 302 | hence "det A = det ?A" by simp | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 303 | moreover have "det A = - det ?A" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 304 | by (simp add: det_permute_rows[OF permutes_swap_id[OF i j]] sign_swap_id ij th1) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 305 | ultimately show "det A = 0" by (metis tha) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 306 | qed | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 307 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 308 | lemma det_identical_columns: | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 309 | fixes A :: "'a::ordered_idom^'n^'n" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 310 |   assumes i: "i\<in>{1 .. dimindex (UNIV :: 'n set)}" 
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 311 |   and j: "j\<in>{1 .. dimindex (UNIV :: 'n set)}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 312 | and ij: "i \<noteq> j" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 313 | and r: "column i A = column j A" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 314 | shows "det A = 0" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 315 | apply (subst det_transp[symmetric]) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 316 | apply (rule det_identical_rows[OF i j ij]) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 317 | by (metis row_transp i j r) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 318 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 319 | lemma det_zero_row: | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 320 |   fixes A :: "'a::{idom, ring_char_0}^'n^'n"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 321 |   assumes i: "i\<in>{1 .. dimindex (UNIV :: 'n set)}" 
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 322 | and r: "row i A = 0" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 323 | shows "det A = 0" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 324 | using i r | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 325 | apply (simp add: row_def det_def Cart_lambda_beta Cart_eq vector_component del: One_nat_def) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 326 | apply (rule setsum_0') | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 327 | apply (clarsimp simp add: sign_nz simp del: One_nat_def) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 328 | apply (rule setprod_zero) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 329 | apply simp | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 330 | apply (rule bexI[where x=i]) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 331 | apply (erule_tac x="a i" in ballE) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 332 | apply (subgoal_tac "(0\<Colon>'a ^ 'n) $ a i = 0") | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 333 | apply simp | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 334 | apply (rule zero_index) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 335 | apply (drule permutes_in_image[of _ _ i]) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 336 | apply simp | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 337 | apply (drule permutes_in_image[of _ _ i]) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 338 | apply simp | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 339 | apply simp | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 340 | done | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 341 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 342 | lemma det_zero_column: | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 343 |   fixes A :: "'a::{idom,ring_char_0}^'n^'n"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 344 |   assumes i: "i\<in>{1 .. dimindex (UNIV :: 'n set)}" 
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 345 | and r: "column i A = 0" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 346 | shows "det A = 0" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 347 | apply (subst det_transp[symmetric]) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 348 | apply (rule det_zero_row[OF i]) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 349 | by (metis row_transp r i) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 350 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 351 | lemma setsum_lambda_beta[simp]: "setsum (\<lambda>i. ((\<chi> i. g i) :: 'a::{comm_monoid_add}^'n) $ i ) {1 .. dimindex (UNIV :: 'n set)} = setsum g {1 .. dimindex (UNIV :: 'n set)}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 352 | by (simp add: Cart_lambda_beta) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 353 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 354 | lemma setprod_lambda_beta[simp]: "setprod (\<lambda>i. ((\<chi> i. g i) :: 'a::{comm_monoid_mult}^'n) $ i ) {1 .. dimindex (UNIV :: 'n set)} = setprod g {1 .. dimindex (UNIV :: 'n set)}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 355 | apply (rule setprod_cong) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 356 | apply simp | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 357 | apply (simp add: Cart_lambda_beta') | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 358 | done | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 359 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 360 | lemma setprod_lambda_beta2[simp]: "setprod (\<lambda>i. ((\<chi> i. g i) :: 'a::{comm_monoid_mult}^'n^'n) $ i$ f i ) {1 .. dimindex (UNIV :: 'n set)} = setprod (\<lambda>i. g i $ f i) {1 .. dimindex (UNIV :: 'n set)}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 361 | proof(rule setprod_cong[OF refl]) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 362 |   let ?U = "{1 .. dimindex (UNIV :: 'n set)}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 363 | fix i assume i: "i \<in> ?U" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 364 | from Cart_lambda_beta'[OF i, of g] have | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 365 | "((\<chi> i. g i) :: 'a^'n^'n) $ i = g i" . | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 366 | hence "((\<chi> i. g i) :: 'a^'n^'n) $ i $ f i = g i $ f i" by simp | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 367 | then | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 368 | show "((\<chi> i. g i):: 'a^'n^'n) $ i $ f i = g i $ f i" . | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 369 | qed | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 370 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 371 | lemma det_row_add: | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 372 |   assumes k: "k \<in> {1 .. dimindex (UNIV :: 'n set)}" 
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 373 | shows "det((\<chi> i. if i = k then a i + b i else c i)::'a::comm_ring_1^'n^'n) = | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 374 | det((\<chi> i. if i = k then a i else c i)::'a::comm_ring_1^'n^'n) + | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 375 | det((\<chi> i. if i = k then b i else c i)::'a::comm_ring_1^'n^'n)" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 376 | unfolding det_def setprod_lambda_beta2 setsum_addf[symmetric] | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 377 | proof (rule setsum_cong2) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 378 |   let ?U = "{1 .. dimindex (UNIV :: 'n set)}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 379 |   let ?pU = "{p. p permutes ?U}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 380 | let ?f = "(\<lambda>i. if i = k then a i + b i else c i)::nat \<Rightarrow> 'a::comm_ring_1^'n" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 381 | let ?g = "(\<lambda> i. if i = k then a i else c i)::nat \<Rightarrow> 'a::comm_ring_1^'n" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 382 | let ?h = "(\<lambda> i. if i = k then b i else c i)::nat \<Rightarrow> 'a::comm_ring_1^'n" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 383 | fix p assume p: "p \<in> ?pU" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 384 |   let ?Uk = "?U - {k}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 385 | from p have pU: "p permutes ?U" by blast | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 386 | from k have pkU: "p k \<in> ?U" by (simp only: permutes_in_image[OF pU]) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 387 | note pin[simp] = permutes_in_image[OF pU] | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 388 | have kU: "?U = insert k ?Uk" using k by blast | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 389 |   {fix j assume j: "j \<in> ?Uk"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 390 | from j have "?f j $ p j = ?g j $ p j" and "?f j $ p j= ?h j $ p j" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 391 | by simp_all} | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 392 | then have th1: "setprod (\<lambda>i. ?f i $ p i) ?Uk = setprod (\<lambda>i. ?g i $ p i) ?Uk" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 393 | and th2: "setprod (\<lambda>i. ?f i $ p i) ?Uk = setprod (\<lambda>i. ?h i $ p i) ?Uk" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 394 | apply - | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 395 | apply (rule setprod_cong, simp_all)+ | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 396 | done | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 397 | have th3: "finite ?Uk" "k \<notin> ?Uk" using k by auto | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 398 | have "setprod (\<lambda>i. ?f i $ p i) ?U = setprod (\<lambda>i. ?f i $ p i) (insert k ?Uk)" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 399 | unfolding kU[symmetric] .. | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 400 | also have "\<dots> = ?f k $ p k * setprod (\<lambda>i. ?f i $ p i) ?Uk" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 401 | apply (rule setprod_insert) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 402 | apply simp | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 403 | using k by blast | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 404 | also have "\<dots> = (a k $ p k * setprod (\<lambda>i. ?f i $ p i) ?Uk) + (b k$ p k * setprod (\<lambda>i. ?f i $ p i) ?Uk)" using pkU by (simp add: ring_simps vector_component) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 405 | also have "\<dots> = (a k $ p k * setprod (\<lambda>i. ?g i $ p i) ?Uk) + (b k$ p k * setprod (\<lambda>i. ?h i $ p i) ?Uk)" by (metis th1 th2) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 406 | also have "\<dots> = setprod (\<lambda>i. ?g i $ p i) (insert k ?Uk) + setprod (\<lambda>i. ?h i $ p i) (insert k ?Uk)" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 407 | unfolding setprod_insert[OF th3] by simp | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 408 | finally have "setprod (\<lambda>i. ?f i $ p i) ?U = setprod (\<lambda>i. ?g i $ p i) ?U + setprod (\<lambda>i. ?h i $ p i) ?U" unfolding kU[symmetric] . | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 409 | then show "of_int (sign p) * setprod (\<lambda>i. ?f i $ p i) ?U = of_int (sign p) * setprod (\<lambda>i. ?g i $ p i) ?U + of_int (sign p) * setprod (\<lambda>i. ?h i $ p i) ?U" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 410 | by (simp add: ring_simps) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 411 | qed | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 412 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 413 | lemma det_row_mul: | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 414 |   assumes k: "k \<in> {1 .. dimindex (UNIV :: 'n set)}" 
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 415 | shows "det((\<chi> i. if i = k then c *s a i else b i)::'a::comm_ring_1^'n^'n) = | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 416 | c* det((\<chi> i. if i = k then a i else b i)::'a::comm_ring_1^'n^'n)" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 417 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 418 | unfolding det_def setprod_lambda_beta2 setsum_right_distrib | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 419 | proof (rule setsum_cong2) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 420 |   let ?U = "{1 .. dimindex (UNIV :: 'n set)}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 421 |   let ?pU = "{p. p permutes ?U}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 422 | let ?f = "(\<lambda>i. if i = k then c*s a i else b i)::nat \<Rightarrow> 'a::comm_ring_1^'n" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 423 | let ?g = "(\<lambda> i. if i = k then a i else b i)::nat \<Rightarrow> 'a::comm_ring_1^'n" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 424 | fix p assume p: "p \<in> ?pU" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 425 |   let ?Uk = "?U - {k}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 426 | from p have pU: "p permutes ?U" by blast | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 427 | from k have pkU: "p k \<in> ?U" by (simp only: permutes_in_image[OF pU]) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 428 | note pin[simp] = permutes_in_image[OF pU] | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 429 | have kU: "?U = insert k ?Uk" using k by blast | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 430 |   {fix j assume j: "j \<in> ?Uk"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 431 | from j have "?f j $ p j = ?g j $ p j" by simp} | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 432 | then have th1: "setprod (\<lambda>i. ?f i $ p i) ?Uk = setprod (\<lambda>i. ?g i $ p i) ?Uk" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 433 | apply - | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 434 | apply (rule setprod_cong, simp_all) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 435 | done | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 436 | have th3: "finite ?Uk" "k \<notin> ?Uk" using k by auto | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 437 | have "setprod (\<lambda>i. ?f i $ p i) ?U = setprod (\<lambda>i. ?f i $ p i) (insert k ?Uk)" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 438 | unfolding kU[symmetric] .. | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 439 | also have "\<dots> = ?f k $ p k * setprod (\<lambda>i. ?f i $ p i) ?Uk" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 440 | apply (rule setprod_insert) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 441 | apply simp | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 442 | using k by blast | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 443 | also have "\<dots> = (c*s a k) $ p k * setprod (\<lambda>i. ?f i $ p i) ?Uk" using pkU by (simp add: ring_simps vector_component) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 444 | also have "\<dots> = c* (a k $ p k * setprod (\<lambda>i. ?g i $ p i) ?Uk)" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 445 | unfolding th1 using pkU by (simp add: vector_component mult_ac) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 446 | also have "\<dots> = c* (setprod (\<lambda>i. ?g i $ p i) (insert k ?Uk))" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 447 | unfolding setprod_insert[OF th3] by simp | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 448 | finally have "setprod (\<lambda>i. ?f i $ p i) ?U = c* (setprod (\<lambda>i. ?g i $ p i) ?U)" unfolding kU[symmetric] . | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 449 | then show "of_int (sign p) * setprod (\<lambda>i. ?f i $ p i) ?U = c * (of_int (sign p) * setprod (\<lambda>i. ?g i $ p i) ?U)" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 450 | by (simp add: ring_simps) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 451 | qed | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 452 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 453 | lemma det_row_0: | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 454 |   assumes k: "k \<in> {1 .. dimindex (UNIV :: 'n set)}" 
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 455 | shows "det((\<chi> i. if i = k then 0 else b i)::'a::comm_ring_1^'n^'n) = 0" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 456 | using det_row_mul[OF k, of 0 "\<lambda>i. 1" b] | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 457 | apply (simp) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 458 | unfolding vector_smult_lzero . | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 459 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 460 | lemma det_row_operation: | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 461 | fixes A :: "'a::ordered_idom^'n^'n" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 462 |   assumes i: "i \<in> {1 .. dimindex(UNIV :: 'n set)}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 463 |   and j: "j \<in> {1 .. dimindex(UNIV :: 'n set)}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 464 | and ij: "i \<noteq> j" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 465 | shows "det (\<chi> k. if k = i then row i A + c *s row j A else row k A) = det A" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 466 | proof- | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 467 | let ?Z = "(\<chi> k. if k = i then row j A else row k A) :: 'a ^'n^'n" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 468 | have th: "row i ?Z = row j ?Z" using i j by (vector row_def) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 469 | have th2: "((\<chi> k. if k = i then row i A else row k A) :: 'a^'n^'n) = A" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 470 | using i j by (vector row_def) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 471 | show ?thesis | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 472 | unfolding det_row_add [OF i] det_row_mul[OF i] det_identical_rows[OF i j ij th] th2 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 473 | by simp | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 474 | qed | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 475 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 476 | lemma det_row_span: | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 477 | fixes A :: "'a:: ordered_idom^'n^'n" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 478 |   assumes i: "i \<in> {1 .. dimindex(UNIV :: 'n set)}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 479 |   and x: "x \<in> span {row j A |j. j\<in> {1 .. dimindex(UNIV :: 'n set)} \<and> j\<noteq> i}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 480 | shows "det (\<chi> k. if k = i then row i A + x else row k A) = det A" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 481 | proof- | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 482 |   let ?U = "{1 .. dimindex(UNIV :: 'n set)}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 483 |   let ?S = "{row j A |j. j\<in> ?U \<and> j\<noteq> i}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 484 | let ?d = "\<lambda>x. det (\<chi> k. if k = i then x else row k A)" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 485 | let ?P = "\<lambda>x. ?d (row i A + x) = det A" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 486 |   {fix k 
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 487 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 488 | have "(if k = i then row i A + 0 else row k A) = row k A" by simp} | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 489 | then have P0: "?P 0" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 490 | apply - | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 491 | apply (rule cong[of det, OF refl]) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 492 | using i by (vector row_def) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 493 | moreover | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 494 |   {fix c z y assume zS: "z \<in> ?S" and Py: "?P y"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 495 | from zS obtain j where j: "z = row j A" "j \<in> ?U" "i \<noteq> j" by blast | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 496 | let ?w = "row i A + y" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 497 | have th0: "row i A + (c*s z + y) = ?w + c*s z" by vector | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 498 | have thz: "?d z = 0" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 499 | apply (rule det_identical_rows[OF i j(2,3)]) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 500 | using i j by (vector row_def) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 501 | have "?d (row i A + (c*s z + y)) = ?d (?w + c*s z)" unfolding th0 .. | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 502 | then have "?P (c*s z + y)" unfolding thz Py det_row_mul[OF i] det_row_add[OF i] | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 503 | by simp } | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 504 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 505 | ultimately show ?thesis | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 506 | apply - | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 507 | apply (rule span_induct_alt[of ?P ?S, OF P0]) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 508 | apply blast | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 509 | apply (rule x) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 510 | done | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 511 | qed | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 512 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 513 | (* ------------------------------------------------------------------------- *) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 514 | (* May as well do this, though it's a bit unsatisfactory since it ignores *) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 515 | (* exact duplicates by considering the rows/columns as a set. *) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 516 | (* ------------------------------------------------------------------------- *) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 517 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 518 | lemma det_dependent_rows: | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 519 | fixes A:: "'a::ordered_idom^'n^'n" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 520 | assumes d: "dependent (rows A)" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 521 | shows "det A = 0" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 522 | proof- | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 523 |   let ?U = "{1 .. dimindex (UNIV :: 'n set)}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 524 |   from d obtain i where i: "i \<in> ?U" "row i A \<in> span (rows A - {row i A})"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 525 | unfolding dependent_def rows_def by blast | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 526 |   {fix j k assume j: "j \<in>?U" and k: "k \<in> ?U" and jk: "j \<noteq> k"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 527 | and c: "row j A = row k A" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 528 | from det_identical_rows[OF j k jk c] have ?thesis .} | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 529 | moreover | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 530 |   {assume H: "\<And> i j. i\<in> ?U \<Longrightarrow> j \<in> ?U \<Longrightarrow> i \<noteq> j \<Longrightarrow> row i A \<noteq> row j A"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 531 |     have th0: "- row i A \<in> span {row j A|j. j \<in> ?U \<and> j \<noteq> i}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 532 | apply (rule span_neg) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 533 | apply (rule set_rev_mp) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 534 | apply (rule i(2)) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 535 | apply (rule span_mono) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 536 | using H i by (auto simp add: rows_def) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 537 | from det_row_span[OF i(1) th0] | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 538 | have "det A = det (\<chi> k. if k = i then 0 *s 1 else row k A)" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 539 | unfolding right_minus vector_smult_lzero .. | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 540 | with det_row_mul[OF i(1), of "0::'a" "\<lambda>i. 1"] | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 541 | have "det A = 0" by simp} | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 542 | ultimately show ?thesis by blast | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 543 | qed | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 544 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 545 | lemma det_dependent_columns: assumes d: "dependent(columns (A::'a::ordered_idom^'n^'n))" shows "det A = 0" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 546 | by (metis d det_dependent_rows rows_transp det_transp) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 547 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 548 | (* ------------------------------------------------------------------------- *) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 549 | (* Multilinearity and the multiplication formula. *) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 550 | (* ------------------------------------------------------------------------- *) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 551 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 552 | lemma Cart_lambda_cong: "(\<And>x. x \<in> {1 .. dimindex (UNIV :: 'n set)} \<Longrightarrow> f x = g x) \<Longrightarrow> (Cart_lambda f::'a^'n) = (Cart_lambda g :: 'a^'n)"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 553 | apply (rule iffD1[OF Cart_lambda_unique]) by vector | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 554 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 555 | lemma det_linear_row_setsum: | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 556 |   assumes fS: "finite S" and k: "k \<in> {1 .. dimindex (UNIV :: 'n set)}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 557 | shows "det ((\<chi> i. if i = k then setsum (a i) S else c i)::'a::comm_ring_1^'n^'n) = setsum (\<lambda>j. det ((\<chi> i. if i = k then a i j else c i)::'a^'n^'n)) S" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 558 | using k | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 559 | proof(induct rule: finite_induct[OF fS]) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 560 | case 1 thus ?case apply simp unfolding setsum_empty det_row_0[OF k] .. | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 561 | next | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 562 | case (2 x F) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 563 | then show ?case by (simp add: det_row_add cong del: if_weak_cong) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 564 | qed | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 565 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 566 | lemma finite_bounded_functions: | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 567 | assumes fS: "finite S" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 568 |   shows "finite {f. (\<forall>i \<in> {1.. (k::nat)}. f i \<in> S) \<and> (\<forall>i. i \<notin> {1 .. k} \<longrightarrow> f i = i)}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 569 | proof(induct k) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 570 | case 0 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 571 |   have th: "{f. \<forall>i. f i = i} = {id}" by (auto intro: ext)
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 572 | show ?case by (auto simp add: th) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 573 | next | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 574 | case (Suc k) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 575 | let ?f = "\<lambda>(y::nat,g) i. if i = Suc k then y else g i" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 576 |   let ?S = "?f ` (S \<times> {f. (\<forall>i\<in>{1..k}. f i \<in> S) \<and> (\<forall>i. i \<notin> {1..k} \<longrightarrow> f i = i)})"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 577 |   have "?S = {f. (\<forall>i\<in>{1.. Suc k}. f i \<in> S) \<and> (\<forall>i. i \<notin> {1.. Suc k} \<longrightarrow> f i = i)}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 578 | apply (auto simp add: image_iff) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 579 | apply (rule_tac x="x (Suc k)" in bexI) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 580 | apply (rule_tac x = "\<lambda>i. if i = Suc k then i else x i" in exI) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 581 | apply (auto intro: ext) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 582 | done | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 583 | with finite_imageI[OF finite_cartesian_product[OF fS Suc.hyps(1)], of ?f] | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 584 | show ?case by metis | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 585 | qed | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 586 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 587 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 588 | lemma eq_id_iff[simp]: "(\<forall>x. f x = x) = (f = id)" by (auto intro: ext) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 589 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 590 | lemma det_linear_rows_setsum_lemma: | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 591 | assumes fS: "finite S" and k: "k \<le> dimindex (UNIV :: 'n set)" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 592 | shows "det((\<chi> i. if i <= k then setsum (a i) S else c i):: 'a::comm_ring_1^'n^'n) = | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 593 | setsum (\<lambda>f. det((\<chi> i. if i <= k then a i (f i) else c i)::'a^'n^'n)) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 594 |                  {f. (\<forall>i \<in> {1 .. k}. f i \<in> S) \<and> (\<forall>i. i \<notin> {1..k} \<longrightarrow> f i = i)}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 595 | using k | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 596 | proof(induct k arbitrary: a c) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 597 | case 0 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 598 | have th0: "\<And>x y. (\<chi> i. if i <= 0 then x i else y i) = (\<chi> i. y i)" by vector | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 599 | from "0.prems" show ?case unfolding th0 by simp | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 600 | next | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 601 | case (Suc k a c) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 602 |   let ?F = "\<lambda>k. {f. (\<forall>i \<in> {1 .. k}. f i \<in> S) \<and> (\<forall>i. i \<notin> {1..k} \<longrightarrow> f i = i)}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 603 | let ?h = "\<lambda>(y::nat,g) i. if i = Suc k then y else g i" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 604 | let ?k = "\<lambda>h. (h(Suc k),(\<lambda>i. if i = Suc k then i else h i))" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 605 | let ?s = "\<lambda> k a c f. det((\<chi> i. if i <= k then a i (f i) else c i)::'a^'n^'n)" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 606 | let ?c = "\<lambda>i. if i = Suc k then a i j else c i" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 607 |   from Suc.prems have Sk: "Suc k \<in> {1 .. dimindex (UNIV :: 'n set)}" by simp
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 608 | from Suc.prems have k': "k \<le> dimindex (UNIV :: 'n set)" by arith | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 609 | have thif: "\<And>a b c d. (if b \<or> a then c else d) = (if a then c else if b then c else d)" by simp | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 610 | have thif2: "\<And>a b c d e. (if a then b else if c then d else e) = | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 611 | (if c then (if a then b else d) else (if a then b else e))" by simp | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 612 | have "det (\<chi> i. if i \<le> Suc k then setsum (a i) S else c i) = | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 613 | det (\<chi> i. if i = Suc k then setsum (a i) S | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 614 | else if i \<le> k then setsum (a i) S else c i)" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 615 | unfolding le_Suc_eq thif .. | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 616 | also have "\<dots> = (\<Sum>j\<in>S. det (\<chi> i. if i \<le> k then setsum (a i) S | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 617 | else if i = Suc k then a i j else c i))" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 618 | unfolding det_linear_row_setsum[OF fS Sk] | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 619 | apply (subst thif2) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 620 | by (simp cong del: if_weak_cong cong add: if_cong) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 621 | finally have tha: | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 622 | "det (\<chi> i. if i \<le> Suc k then setsum (a i) S else c i) = | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 623 | (\<Sum>(j, f)\<in>S \<times> ?F k. det (\<chi> i. if i \<le> k then a i (f i) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 624 | else if i = Suc k then a i j | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 625 | else c i))" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 626 | unfolding Suc.hyps[OF k'] unfolding setsum_cartesian_product by blast | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 627 | show ?case unfolding tha | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 628 | apply(rule setsum_eq_general_reverses[where h= "?h" and k= "?k"], | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 629 | blast intro: finite_cartesian_product fS finite_bounded_functions[OF fS], | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 630 | blast intro: finite_cartesian_product fS finite_bounded_functions[OF fS], auto intro: ext) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 631 | apply (rule cong[OF refl[of det]]) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 632 | by vector | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 633 | qed | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 634 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 635 | lemma det_linear_rows_setsum: | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 636 | assumes fS: "finite S" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 637 |   shows "det (\<chi> i. setsum (a i) S) = setsum (\<lambda>f. det (\<chi> i. a i (f i) :: 'a::comm_ring_1 ^ 'n^'n)) {f. (\<forall>i \<in> {1 .. dimindex (UNIV :: 'n set)}. f i \<in> S) \<and> (\<forall>i. i \<notin> {1.. dimindex (UNIV :: 'n set)} \<longrightarrow> f i = i)}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 638 | proof- | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 639 | have th0: "\<And>x y. ((\<chi> i. if i <= dimindex(UNIV:: 'n set) then x i else y i) :: 'a^'n^'n) = (\<chi> i. x i)" by vector | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 640 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 641 | from det_linear_rows_setsum_lemma[OF fS, of "dimindex (UNIV :: 'n set)" a, unfolded th0, OF order_refl] show ?thesis by blast | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 642 | qed | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 643 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 644 | lemma matrix_mul_setsum_alt: | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 645 | fixes A B :: "'a::comm_ring_1^'n^'n" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 646 |   shows "A ** B = (\<chi> i. setsum (\<lambda>k. A$i$k *s B $ k) {1 .. dimindex (UNIV :: 'n set)})"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 647 | by (vector matrix_matrix_mult_def setsum_component) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 648 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 649 | lemma det_rows_mul: | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 650 | "det((\<chi> i. c i *s a i)::'a::comm_ring_1^'n^'n) = | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 651 |   setprod (\<lambda>i. c i) {1..dimindex(UNIV:: 'n set)} * det((\<chi> i. a i)::'a^'n^'n)"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 652 | proof (simp add: det_def Cart_lambda_beta' setsum_right_distrib vector_component cong add: setprod_cong del: One_nat_def, rule setsum_cong2) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 653 |   let ?U = "{1 .. dimindex(UNIV :: 'n set)}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 654 |   let ?PU = "{p. p permutes ?U}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 655 | fix p assume pU: "p \<in> ?PU" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 656 | let ?s = "of_int (sign p)" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 657 | from pU have p: "p permutes ?U" by blast | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 658 | have "setprod (\<lambda>i. (c i *s a i) $ p i) ?U = setprod (\<lambda>i. c i * a i $ p i) ?U" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 659 | apply (rule setprod_cong, blast) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 660 | by (auto simp only: permutes_in_image[OF p] intro: vector_smult_component) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 661 | also have "\<dots> = setprod c ?U * setprod (\<lambda>i. a i $ p i) ?U" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 662 | unfolding setprod_timesf .. | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 663 | finally show "?s * (\<Prod>xa\<in>?U. (c xa *s a xa) $ p xa) = | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 664 | setprod c ?U * (?s* (\<Prod>xa\<in>?U. a xa $ p xa))" by (simp add: ring_simps) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 665 | qed | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 666 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 667 | lemma det_mul: | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 668 | fixes A B :: "'a::ordered_idom^'n^'n" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 669 | shows "det (A ** B) = det A * det B" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 670 | proof- | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 671 |   let ?U = "{1 .. dimindex (UNIV :: 'n set)}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 672 |   let ?F = "{f. (\<forall>i\<in> ?U. f i \<in> ?U) \<and> (\<forall>i. i \<notin> ?U \<longrightarrow> f i = i)}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 673 |   let ?PU = "{p. p permutes ?U}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 674 | have fU: "finite ?U" by simp | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 675 | have fF: "finite ?F" using finite_bounded_functions[OF fU] . | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 676 |   {fix p assume p: "p permutes ?U"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 677 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 678 | have "p \<in> ?F" unfolding mem_Collect_eq permutes_in_image[OF p] | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 679 | using p[unfolded permutes_def] by simp} | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 680 | then have PUF: "?PU \<subseteq> ?F" by blast | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 681 |   {fix f assume fPU: "f \<in> ?F - ?PU"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 682 | have fUU: "f ` ?U \<subseteq> ?U" using fPU by auto | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 683 | from fPU have f: "\<forall>i \<in> ?U. f i \<in> ?U" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 684 | "\<forall>i. i \<notin> ?U \<longrightarrow> f i = i" "\<not>(\<forall>y. \<exists>!x. f x = y)" unfolding permutes_def | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 685 | by auto | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 686 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 687 | let ?A = "(\<chi> i. A$i$f i *s B$f i) :: 'a^'n^'n" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 688 | let ?B = "(\<chi> i. B$f i) :: 'a^'n^'n" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 689 |     {assume fni: "\<not> inj_on f ?U"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 690 | then obtain i j where ij: "i \<in> ?U" "j \<in> ?U" "f i = f j" "i \<noteq> j" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 691 | unfolding inj_on_def by blast | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 692 | from ij | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 693 | have rth: "row i ?B = row j ?B" by (vector row_def) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 694 | from det_identical_rows[OF ij(1,2,4) rth] | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 695 | have "det (\<chi> i. A$i$f i *s B$f i) = 0" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 696 | unfolding det_rows_mul by simp} | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 697 | moreover | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 698 |     {assume fi: "inj_on f ?U"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 699 | from f fi have fith: "\<And>i j. f i = f j \<Longrightarrow> i = j" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 700 | unfolding inj_on_def | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 701 | apply (case_tac "i \<in> ?U") | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 702 | apply (case_tac "j \<in> ?U") by metis+ | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 703 | note fs = fi[unfolded surjective_iff_injective_gen[OF fU fU refl fUU, symmetric]] | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 704 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 705 |       {fix y
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 706 | from fs f have "\<exists>x. f x = y" by (cases "y \<in> ?U") blast+ | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 707 | then obtain x where x: "f x = y" by blast | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 708 | 	{fix z assume z: "f z = y" from fith x z have "z = x" by metis}
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 709 | with x have "\<exists>!x. f x = y" by blast} | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 710 | with f(3) have "det (\<chi> i. A$i$f i *s B$f i) = 0" by blast} | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 711 | ultimately have "det (\<chi> i. A$i$f i *s B$f i) = 0" by blast} | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 712 | hence zth: "\<forall> f\<in> ?F - ?PU. det (\<chi> i. A$i$f i *s B$f i) = 0" by simp | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 713 |   {fix p assume pU: "p \<in> ?PU"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 714 | from pU have p: "p permutes ?U" by blast | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 715 | let ?s = "\<lambda>p. of_int (sign p)" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 716 | let ?f = "\<lambda>q. ?s p * (\<Prod>i\<in> ?U. A $ i $ p i) * | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 717 | (?s q * (\<Prod>i\<in> ?U. B $ i $ q i))" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 718 | have "(setsum (\<lambda>q. ?s q * | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 719 | (\<Prod>i\<in> ?U. (\<chi> i. A $ i $ p i *s B $ p i :: 'a^'n^'n) $ i $ q i)) ?PU) = | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 720 | (setsum (\<lambda>q. ?s p * (\<Prod>i\<in> ?U. A $ i $ p i) * | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 721 | (?s q * (\<Prod>i\<in> ?U. B $ i $ q i))) ?PU)" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 722 | unfolding sum_permutations_compose_right[OF permutes_inv[OF p], of ?f] | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 723 | proof(rule setsum_cong2) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 724 | fix q assume qU: "q \<in> ?PU" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 725 | hence q: "q permutes ?U" by blast | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 726 | from p q have pp: "permutation p" and pq: "permutation q" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 727 | unfolding permutation_permutes by auto | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 728 | have th00: "of_int (sign p) * of_int (sign p) = (1::'a)" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 729 | "\<And>a. of_int (sign p) * (of_int (sign p) * a) = a" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 730 | unfolding mult_assoc[symmetric] unfolding of_int_mult[symmetric] | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 731 | by (simp_all add: sign_idempotent) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 732 | have ths: "?s q = ?s p * ?s (q o inv p)" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 733 | using pp pq permutation_inverse[OF pp] sign_inverse[OF pp] | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 734 | by (simp add: th00 mult_ac sign_idempotent sign_compose) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 735 | have th001: "setprod (\<lambda>i. B$i$ q (inv p i)) ?U = setprod ((\<lambda>i. B$i$ q (inv p i)) o p) ?U" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 736 | by (rule setprod_permute[OF p]) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 737 | have thp: "setprod (\<lambda>i. (\<chi> i. A$i$p i *s B$p i :: 'a^'n^'n) $i $ q i) ?U = setprod (\<lambda>i. A$i$p i) ?U * setprod (\<lambda>i. B$i$ q (inv p i)) ?U" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 738 | unfolding th001 setprod_timesf[symmetric] o_def permutes_inverses[OF p] | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 739 | apply (rule setprod_cong[OF refl]) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 740 | using permutes_in_image[OF q] by vector | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 741 | show "?s q * setprod (\<lambda>i. (((\<chi> i. A$i$p i *s B$p i) :: 'a^'n^'n)$i$q i)) ?U = ?s p * (setprod (\<lambda>i. A$i$p i) ?U) * (?s (q o inv p) * setprod (\<lambda>i. B$i$(q o inv p) i) ?U)" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 742 | using ths thp pp pq permutation_inverse[OF pp] sign_inverse[OF pp] | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 743 | by (simp add: sign_nz th00 ring_simps sign_idempotent sign_compose) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 744 | qed | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 745 | } | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 746 | then have th2: "setsum (\<lambda>f. det (\<chi> i. A$i$f i *s B$f i)) ?PU = det A * det B" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 747 | unfolding det_def setsum_product | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 748 | by (rule setsum_cong2) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 749 | have "det (A**B) = setsum (\<lambda>f. det (\<chi> i. A $ i $ f i *s B $ f i)) ?F" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 750 | unfolding matrix_mul_setsum_alt det_linear_rows_setsum[OF fU] .. | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 751 | also have "\<dots> = setsum (\<lambda>f. det (\<chi> i. A$i$f i *s B$f i)) ?PU" | 
| 30259 | 752 | using setsum_mono_zero_cong_left[OF fF PUF zth, symmetric] | 
| 753 | unfolding det_rows_mul by auto | |
| 29846 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 754 | finally show ?thesis unfolding th2 . | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 755 | qed | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 756 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 757 | (* ------------------------------------------------------------------------- *) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 758 | (* Relation to invertibility. *) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 759 | (* ------------------------------------------------------------------------- *) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 760 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 761 | lemma invertible_left_inverse: | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 762 | fixes A :: "real^'n^'n" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 763 | shows "invertible A \<longleftrightarrow> (\<exists>(B::real^'n^'n). B** A = mat 1)" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 764 | by (metis invertible_def matrix_left_right_inverse) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 765 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 766 | lemma invertible_righ_inverse: | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 767 | fixes A :: "real^'n^'n" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 768 | shows "invertible A \<longleftrightarrow> (\<exists>(B::real^'n^'n). A** B = mat 1)" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 769 | by (metis invertible_def matrix_left_right_inverse) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 770 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 771 | lemma invertible_det_nz: | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 772 | fixes A::"real ^'n^'n" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 773 | shows "invertible A \<longleftrightarrow> det A \<noteq> 0" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 774 | proof- | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 775 |   {assume "invertible A"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 776 | then obtain B :: "real ^'n^'n" where B: "A ** B = mat 1" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 777 | unfolding invertible_righ_inverse by blast | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 778 | hence "det (A ** B) = det (mat 1 :: real ^'n^'n)" by simp | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 779 | hence "det A \<noteq> 0" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 780 | apply (simp add: det_mul det_I) by algebra } | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 781 | moreover | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 782 |   {assume H: "\<not> invertible A"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 783 |     let ?U = "{1 .. dimindex(UNIV :: 'n set)}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 784 | have fU: "finite ?U" by simp | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 785 | from H obtain c i where c: "setsum (\<lambda>i. c i *s row i A) ?U = 0" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 786 | and iU: "i \<in> ?U" and ci: "c i \<noteq> 0" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 787 | unfolding invertible_righ_inverse | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 788 | unfolding matrix_right_invertible_independent_rows by blast | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 789 | have stupid: "\<And>(a::real^'n) b. a + b = 0 \<Longrightarrow> -a = b" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 790 | apply (drule_tac f="op + (- a)" in cong[OF refl]) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 791 | apply (simp only: ab_left_minus add_assoc[symmetric]) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 792 | apply simp | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 793 | done | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 794 | from c ci | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 795 |     have thr0: "- row i A = setsum (\<lambda>j. (1/ c i) *s c j *s row j A) (?U - {i})"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 796 | unfolding setsum_diff1'[OF fU iU] setsum_cmul | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 797 | apply (simp add: field_simps) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 798 | apply (rule vector_mul_lcancel_imp[OF ci]) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 799 | apply (auto simp add: vector_smult_assoc vector_smult_rneg field_simps) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 800 | unfolding stupid .. | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 801 |     have thr: "- row i A \<in> span {row j A| j. j\<in> ?U \<and> j \<noteq> i}" 
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 802 | unfolding thr0 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 803 | apply (rule span_setsum) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 804 | apply simp | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 805 | apply (rule ballI) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 806 | apply (rule span_mul)+ | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 807 | apply (rule span_superset) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 808 | apply auto | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 809 | done | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 810 | let ?B = "(\<chi> k. if k = i then 0 else row k A) :: real ^'n^'n" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 811 | have thrb: "row i ?B = 0" using iU by (vector row_def) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 812 | have "det A = 0" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 813 | unfolding det_row_span[OF iU thr, symmetric] right_minus | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 814 | unfolding det_zero_row[OF iU thrb] ..} | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 815 | ultimately show ?thesis by blast | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 816 | qed | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 817 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 818 | (* ------------------------------------------------------------------------- *) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 819 | (* Cramer's rule. *) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 820 | (* ------------------------------------------------------------------------- *) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 821 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 822 | lemma cramer_lemma_transp: | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 823 | fixes A:: "'a::ordered_idom^'n^'n" and x :: "'a ^'n" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 824 |   assumes k: "k \<in> {1 .. dimindex(UNIV ::'n set)}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 825 |   shows "det ((\<chi> i. if i = k then setsum (\<lambda>i. x$i *s row i A) {1 .. dimindex(UNIV::'n set)}
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 826 | else row i A)::'a^'n^'n) = x$k * det A" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 827 | (is "?lhs = ?rhs") | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 828 | proof- | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 829 |   let ?U = "{1 .. dimindex (UNIV :: 'n set)}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 830 |   let ?Uk = "?U - {k}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 831 | have U: "?U = insert k ?Uk" using k by blast | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 832 | have fUk: "finite ?Uk" by simp | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 833 | have kUk: "k \<notin> ?Uk" by simp | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 834 | have th00: "\<And>k s. x$k *s row k A + s = (x$k - 1) *s row k A + row k A + s" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 835 | by (vector ring_simps) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 836 | have th001: "\<And>f k . (\<lambda>x. if x = k then f k else f x) = f" by (auto intro: ext) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 837 | have "(\<chi> i. row i A) = A" by (vector row_def) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 838 | then have thd1: "det (\<chi> i. row i A) = det A" by simp | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 839 | have thd0: "det (\<chi> i. if i = k then row k A + (\<Sum>i \<in> ?Uk. x $ i *s row i A) else row i A) = det A" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 840 | apply (rule det_row_span[OF k]) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 841 | apply (rule span_setsum[OF fUk]) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 842 | apply (rule ballI) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 843 | apply (rule span_mul) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 844 | apply (rule span_superset) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 845 | apply auto | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 846 | done | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 847 | show "?lhs = x$k * det A" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 848 | apply (subst U) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 849 | unfolding setsum_insert[OF fUk kUk] | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 850 | apply (subst th00) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 851 | unfolding add_assoc | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 852 | apply (subst det_row_add[OF k]) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 853 | unfolding thd0 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 854 | unfolding det_row_mul[OF k] | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 855 | unfolding th001[of k "\<lambda>i. row i A"] | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 856 | unfolding thd1 by (simp add: ring_simps) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 857 | qed | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 858 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 859 | lemma cramer_lemma: | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 860 | fixes A :: "'a::ordered_idom ^'n^'n" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 861 |   assumes k: "k \<in> {1 .. dimindex (UNIV :: 'n set)}" (is " _ \<in> ?U")
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 862 | shows "det((\<chi> i j. if j = k then (A *v x)$i else A$i$j):: 'a^'n^'n) = x$k * det A" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 863 | proof- | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 864 | have stupid: "\<And>c. setsum (\<lambda>i. c i *s row i (transp A)) ?U = setsum (\<lambda>i. c i *s column i A) ?U" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 865 | by (auto simp add: row_transp intro: setsum_cong2) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 866 | show ?thesis | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 867 | unfolding matrix_mult_vsum | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 868 | unfolding cramer_lemma_transp[OF k, of x "transp A", unfolded det_transp, symmetric] | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 869 | unfolding stupid[of "\<lambda>i. x$i"] | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 870 | apply (subst det_transp[symmetric]) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 871 | apply (rule cong[OF refl[of det]]) by (vector transp_def column_def row_def) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 872 | qed | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 873 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 874 | lemma cramer: | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 875 | fixes A ::"real^'n^'n" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 876 | assumes d0: "det A \<noteq> 0" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 877 | shows "A *v x = b \<longleftrightarrow> x = (\<chi> k. det(\<chi> i j. if j=k then b$i else A$i$j :: real^'n^'n) / det A)" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 878 | proof- | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 879 | from d0 obtain B where B: "A ** B = mat 1" "B ** A = mat 1" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 880 | unfolding invertible_det_nz[symmetric] invertible_def by blast | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 881 | have "(A ** B) *v b = b" by (simp add: B matrix_vector_mul_lid) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 882 | hence "A *v (B *v b) = b" by (simp add: matrix_vector_mul_assoc) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 883 | then have xe: "\<exists>x. A*v x = b" by blast | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 884 |   {fix x assume x: "A *v x = b"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 885 | have "x = (\<chi> k. det(\<chi> i j. if j=k then b$i else A$i$j :: real^'n^'n) / det A)" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 886 | unfolding x[symmetric] | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 887 | using d0 by (simp add: Cart_eq Cart_lambda_beta' cramer_lemma field_simps)} | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 888 | with xe show ?thesis by auto | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 889 | qed | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 890 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 891 | (* ------------------------------------------------------------------------- *) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 892 | (* Orthogonality of a transformation and matrix. *) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 893 | (* ------------------------------------------------------------------------- *) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 894 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 895 | definition "orthogonal_transformation f \<longleftrightarrow> linear f \<and> (\<forall>v w. f v \<bullet> f w = v \<bullet> w)" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 896 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 897 | lemma orthogonal_transformation: "orthogonal_transformation f \<longleftrightarrow> linear f \<and> (\<forall>(v::real ^'n). norm (f v) = norm v)" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 898 | unfolding orthogonal_transformation_def | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 899 | apply auto | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 900 | apply (erule_tac x=v in allE)+ | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 901 | apply (simp add: real_vector_norm_def) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 902 | by (simp add: dot_norm linear_add[symmetric]) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 903 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 904 | definition "orthogonal_matrix (Q::'a::semiring_1^'n^'n) \<longleftrightarrow> transp Q ** Q = mat 1 \<and> Q ** transp Q = mat 1" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 905 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 906 | lemma orthogonal_matrix: "orthogonal_matrix (Q:: real ^'n^'n) \<longleftrightarrow> transp Q ** Q = mat 1" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 907 | by (metis matrix_left_right_inverse orthogonal_matrix_def) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 908 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 909 | lemma orthogonal_matrix_id: "orthogonal_matrix (mat 1)" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 910 | by (simp add: orthogonal_matrix_def transp_mat matrix_mul_lid) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 911 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 912 | lemma orthogonal_matrix_mul: | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 913 | fixes A :: "real ^'n^'n" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 914 | assumes oA : "orthogonal_matrix A" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 915 | and oB: "orthogonal_matrix B" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 916 | shows "orthogonal_matrix(A ** B)" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 917 | using oA oB | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 918 | unfolding orthogonal_matrix matrix_transp_mul | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 919 | apply (subst matrix_mul_assoc) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 920 | apply (subst matrix_mul_assoc[symmetric]) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 921 | by (simp add: matrix_mul_rid) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 922 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 923 | lemma orthogonal_transformation_matrix: | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 924 | fixes f:: "real^'n \<Rightarrow> real^'n" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 925 | shows "orthogonal_transformation f \<longleftrightarrow> linear f \<and> orthogonal_matrix(matrix f)" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 926 | (is "?lhs \<longleftrightarrow> ?rhs") | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 927 | proof- | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 928 | let ?mf = "matrix f" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 929 | let ?ot = "orthogonal_transformation f" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 930 |   let ?U = "{1 .. dimindex (UNIV :: 'n set)}"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 931 | have fU: "finite ?U" by simp | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 932 | let ?m1 = "mat 1 :: real ^'n^'n" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 933 |   {assume ot: ?ot
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 934 | from ot have lf: "linear f" and fd: "\<forall>v w. f v \<bullet> f w = v \<bullet> w" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 935 | unfolding orthogonal_transformation_def orthogonal_matrix by blast+ | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 936 |     {fix i j assume i: "i \<in> ?U" and j: "j \<in> ?U"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 937 | let ?A = "transp ?mf ** ?mf" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 938 | have th0: "\<And>b (x::'a::comm_ring_1). (if b then 1 else 0)*x = (if b then x else 0)" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 939 | "\<And>b (x::'a::comm_ring_1). x*(if b then 1 else 0) = (if b then x else 0)" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 940 | by simp_all | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 941 | from fd[rule_format, of "basis i" "basis j", unfolded matrix_works[OF lf, symmetric] dot_matrix_vector_mul] i j | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 942 | have "?A$i$j = ?m1 $ i $ j" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 943 | by (simp add: Cart_lambda_beta' dot_def matrix_matrix_mult_def columnvector_def rowvector_def basis_def th0 setsum_delta[OF fU] mat_def del: One_nat_def)} | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 944 | hence "orthogonal_matrix ?mf" unfolding orthogonal_matrix by vector | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 945 | with lf have ?rhs by blast} | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 946 | moreover | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 947 |   {assume lf: "linear f" and om: "orthogonal_matrix ?mf"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 948 | from lf om have ?lhs | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 949 | unfolding orthogonal_matrix_def norm_eq orthogonal_transformation | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 950 | unfolding matrix_works[OF lf, symmetric] | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 951 | apply (subst dot_matrix_vector_mul) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 952 | by (simp add: dot_matrix_product matrix_mul_lid del: One_nat_def)} | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 953 | ultimately show ?thesis by blast | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 954 | qed | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 955 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 956 | lemma det_orthogonal_matrix: | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 957 | fixes Q:: "'a::ordered_idom^'n^'n" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 958 | assumes oQ: "orthogonal_matrix Q" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 959 | shows "det Q = 1 \<or> det Q = - 1" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 960 | proof- | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 961 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 962 | have th: "\<And>x::'a. x = 1 \<or> x = - 1 \<longleftrightarrow> x*x = 1" (is "\<And>x::'a. ?ths x") | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 963 | proof- | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 964 | fix x:: 'a | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 965 | have th0: "x*x - 1 = (x - 1)*(x + 1)" by (simp add: ring_simps) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 966 | have th1: "\<And>(x::'a) y. x = - y \<longleftrightarrow> x + y = 0" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 967 | apply (subst eq_iff_diff_eq_0) by simp | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 968 | have "x*x = 1 \<longleftrightarrow> x*x - 1 = 0" by simp | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 969 | also have "\<dots> \<longleftrightarrow> x = 1 \<or> x = - 1" unfolding th0 th1 by simp | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 970 | finally show "?ths x" .. | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 971 | qed | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 972 | from oQ have "Q ** transp Q = mat 1" by (metis orthogonal_matrix_def) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 973 | hence "det (Q ** transp Q) = det (mat 1:: 'a^'n^'n)" by simp | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 974 | hence "det Q * det Q = 1" by (simp add: det_mul det_I det_transp) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 975 | then show ?thesis unfolding th . | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 976 | qed | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 977 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 978 | (* ------------------------------------------------------------------------- *) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 979 | (* Linearity of scaling, and hence isometry, that preserves origin. *) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 980 | (* ------------------------------------------------------------------------- *) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 981 | lemma scaling_linear: | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 982 | fixes f :: "real ^'n \<Rightarrow> real ^'n" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 983 | assumes f0: "f 0 = 0" and fd: "\<forall>x y. dist (f x) (f y) = c * dist x y" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 984 | shows "linear f" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 985 | proof- | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 986 |   {fix v w 
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 987 |     {fix x note fd[rule_format, of x 0, unfolded dist_def f0 diff_0_right] }
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 988 | note th0 = this | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 989 | have "f v \<bullet> f w = c^2 * (v \<bullet> w)" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 990 | unfolding dot_norm_neg dist_def[symmetric] | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 991 | unfolding th0 fd[rule_format] by (simp add: power2_eq_square field_simps)} | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 992 | note fc = this | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 993 | show ?thesis unfolding linear_def vector_eq | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 994 | by (simp add: dot_lmult dot_ladd dot_rmult dot_radd fc ring_simps) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 995 | qed | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 996 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 997 | lemma isometry_linear: | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 998 | "f (0:: real^'n) = (0:: real^'n) \<Longrightarrow> \<forall>x y. dist(f x) (f y) = dist x y | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 999 | \<Longrightarrow> linear f" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1000 | by (rule scaling_linear[where c=1]) simp_all | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1001 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1002 | (* ------------------------------------------------------------------------- *) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1003 | (* Hence another formulation of orthogonal transformation. *) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1004 | (* ------------------------------------------------------------------------- *) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1005 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1006 | lemma orthogonal_transformation_isometry: | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1007 | "orthogonal_transformation f \<longleftrightarrow> f(0::real^'n) = (0::real^'n) \<and> (\<forall>x y. dist(f x) (f y) = dist x y)" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1008 | unfolding orthogonal_transformation | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1009 | apply (rule iffI) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1010 | apply clarify | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1011 | apply (clarsimp simp add: linear_0 linear_sub[symmetric] dist_def) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1012 | apply (rule conjI) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1013 | apply (rule isometry_linear) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1014 | apply simp | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1015 | apply simp | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1016 | apply clarify | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1017 | apply (erule_tac x=v in allE) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1018 | apply (erule_tac x=0 in allE) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1019 | by (simp add: dist_def) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1020 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1021 | (* ------------------------------------------------------------------------- *) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1022 | (* Can extend an isometry from unit sphere. *) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1023 | (* ------------------------------------------------------------------------- *) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1024 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1025 | lemma isometry_sphere_extend: | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1026 | fixes f:: "real ^'n \<Rightarrow> real ^'n" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1027 | assumes f1: "\<forall>x. norm x = 1 \<longrightarrow> norm (f x) = 1" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1028 | and fd1: "\<forall> x y. norm x = 1 \<longrightarrow> norm y = 1 \<longrightarrow> dist (f x) (f y) = dist x y" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1029 | shows "\<exists>g. orthogonal_transformation g \<and> (\<forall>x. norm x = 1 \<longrightarrow> g x = f x)" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1030 | proof- | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1031 |   {fix x y x' y' x0 y0 x0' y0' :: "real ^'n" 
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1032 | assume H: "x = norm x *s x0" "y = norm y *s y0" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1033 | "x' = norm x *s x0'" "y' = norm y *s y0'" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1034 | "norm x0 = 1" "norm x0' = 1" "norm y0 = 1" "norm y0' = 1" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1035 | "norm(x0' - y0') = norm(x0 - y0)" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1036 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1037 | have "norm(x' - y') = norm(x - y)" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1038 | apply (subst H(1)) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1039 | apply (subst H(2)) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1040 | apply (subst H(3)) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1041 | apply (subst H(4)) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1042 | using H(5-9) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1043 | apply (simp add: norm_eq norm_eq_1) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1044 | apply (simp add: dot_lsub dot_rsub dot_lmult dot_rmult) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1045 | apply (simp add: ring_simps) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1046 | by (simp only: right_distrib[symmetric])} | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1047 | note th0 = this | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1048 | let ?g = "\<lambda>x. if x = 0 then 0 else norm x *s f (inverse (norm x) *s x)" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1049 |   {fix x:: "real ^'n" assume nx: "norm x = 1"
 | 
| 30041 | 1050 | have "?g x = f x" using nx by auto} | 
| 29846 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1051 | hence thfg: "\<forall>x. norm x = 1 \<longrightarrow> ?g x = f x" by blast | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1052 | have g0: "?g 0 = 0" by simp | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1053 |   {fix x y :: "real ^'n"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1054 |     {assume "x = 0" "y = 0"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1055 | then have "dist (?g x) (?g y) = dist x y" by simp } | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1056 | moreover | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1057 |     {assume "x = 0" "y \<noteq> 0"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1058 | then have "dist (?g x) (?g y) = dist x y" | 
| 30041 | 1059 | apply (simp add: dist_def norm_mul) | 
| 29846 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1060 | apply (rule f1[rule_format]) | 
| 30041 | 1061 | by(simp add: norm_mul field_simps)} | 
| 29846 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1062 | moreover | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1063 |     {assume "x \<noteq> 0" "y = 0"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1064 | then have "dist (?g x) (?g y) = dist x y" | 
| 30041 | 1065 | apply (simp add: dist_def norm_mul) | 
| 29846 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1066 | apply (rule f1[rule_format]) | 
| 30041 | 1067 | by(simp add: norm_mul field_simps)} | 
| 29846 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1068 | moreover | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1069 |     {assume z: "x \<noteq> 0" "y \<noteq> 0"
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1070 | have th00: "x = norm x *s inverse (norm x) *s x" "y = norm y *s inverse (norm y) *s y" "norm x *s f (inverse (norm x) *s x) = norm x *s f (inverse (norm x) *s x)" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1071 | "norm y *s f (inverse (norm y) *s y) = norm y *s f (inverse (norm y) *s y)" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1072 | "norm (inverse (norm x) *s x) = 1" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1073 | "norm (f (inverse (norm x) *s x)) = 1" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1074 | "norm (inverse (norm y) *s y) = 1" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1075 | "norm (f (inverse (norm y) *s y)) = 1" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1076 | "norm (f (inverse (norm x) *s x) - f (inverse (norm y) *s y)) = | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1077 | norm (inverse (norm x) *s x - inverse (norm y) *s y)" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1078 | using z | 
| 30041 | 1079 | by (auto simp add: vector_smult_assoc field_simps norm_mul intro: f1[rule_format] fd1[rule_format, unfolded dist_def]) | 
| 29846 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1080 | from z th0[OF th00] have "dist (?g x) (?g y) = dist x y" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1081 | by (simp add: dist_def)} | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1082 | ultimately have "dist (?g x) (?g y) = dist x y" by blast} | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1083 | note thd = this | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1084 | show ?thesis | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1085 | apply (rule exI[where x= ?g]) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1086 | unfolding orthogonal_transformation_isometry | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1087 | using g0 thfg thd by metis | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1088 | qed | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1089 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1090 | (* ------------------------------------------------------------------------- *) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1091 | (* Rotation, reflection, rotoinversion. *) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1092 | (* ------------------------------------------------------------------------- *) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1093 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1094 | definition "rotation_matrix Q \<longleftrightarrow> orthogonal_matrix Q \<and> det Q = 1" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1095 | definition "rotoinversion_matrix Q \<longleftrightarrow> orthogonal_matrix Q \<and> det Q = - 1" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1096 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1097 | lemma orthogonal_rotation_or_rotoinversion: | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1098 | fixes Q :: "'a::ordered_idom^'n^'n" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1099 | shows " orthogonal_matrix Q \<longleftrightarrow> rotation_matrix Q \<or> rotoinversion_matrix Q" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1100 | by (metis rotoinversion_matrix_def rotation_matrix_def det_orthogonal_matrix) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1101 | (* ------------------------------------------------------------------------- *) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1102 | (* Explicit formulas for low dimensions. *) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1103 | (* ------------------------------------------------------------------------- *) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1104 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1105 | lemma setprod_1: "setprod f {(1::nat)..1} = f 1" by simp
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1106 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1107 | lemma setprod_2: "setprod f {(1::nat)..2} = f 1 * f 2" 
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1108 | by (simp add: nat_number setprod_numseg mult_commute) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1109 | lemma setprod_3: "setprod f {(1::nat)..3} = f 1 * f 2 * f 3" 
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1110 | by (simp add: nat_number setprod_numseg mult_commute) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1111 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1112 | lemma det_1: "det (A::'a::comm_ring_1^1^1) = A$1$1" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1113 | by (simp add: det_def dimindex_def permutes_sing sign_id del: One_nat_def) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1114 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1115 | lemma det_2: "det (A::'a::comm_ring_1^2^2) = A$1$1 * A$2$2 - A$1$2 * A$2$1" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1116 | proof- | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1117 |   have f12: "finite {2::nat}" "1 \<notin> {2::nat}" by auto
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1118 |   have th12: "{1 .. 2} = insert (1::nat) {2}" by auto
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1119 | show ?thesis | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1120 | apply (simp add: det_def dimindex_def th12 del: One_nat_def) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1121 | unfolding setsum_over_permutations_insert[OF f12] | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1122 | unfolding permutes_sing | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1123 | apply (simp add: sign_swap_id sign_id swap_id_eq del: One_nat_def) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1124 | by (simp add: arith_simps(31)[symmetric] of_int_minus of_int_1 del: arith_simps(31)) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1125 | qed | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1126 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1127 | lemma det_3: "det (A::'a::comm_ring_1^3^3) = | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1128 | A$1$1 * A$2$2 * A$3$3 + | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1129 | A$1$2 * A$2$3 * A$3$1 + | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1130 | A$1$3 * A$2$1 * A$3$2 - | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1131 | A$1$1 * A$2$3 * A$3$2 - | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1132 | A$1$2 * A$2$1 * A$3$3 - | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1133 | A$1$3 * A$2$2 * A$3$1" | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1134 | proof- | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1135 |   have f123: "finite {(2::nat), 3}" "1 \<notin> {(2::nat), 3}" by auto
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1136 |   have f23: "finite {(3::nat)}" "2 \<notin> {(3::nat)}" by auto
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1137 |   have th12: "{1 .. 3} = insert (1::nat) (insert 2 {3})" by auto
 | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1138 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1139 | show ?thesis | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1140 | apply (simp add: det_def dimindex_def th12 del: One_nat_def) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1141 | unfolding setsum_over_permutations_insert[OF f123] | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1142 | unfolding setsum_over_permutations_insert[OF f23] | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1143 | |
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1144 | unfolding permutes_sing | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1145 | apply (simp add: sign_swap_id permutation_swap_id sign_compose sign_id swap_id_eq del: One_nat_def) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1146 | apply (simp add: arith_simps(31)[symmetric] of_int_minus of_int_1 del: arith_simps(31) One_nat_def) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1147 | by (simp add: ring_simps) | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1148 | qed | 
| 
57dccccc37b3
Traces, Determinant of square matrices and some properties
 chaieb parents: diff
changeset | 1149 | |
| 30041 | 1150 | end |