| author | wenzelm | 
| Sun, 02 Nov 2014 17:06:05 +0100 | |
| changeset 58876 | 1888e3cb8048 | 
| parent 57514 | bdc2c6b40bf2 | 
| child 58877 | 262572d90bc6 | 
| permissions | -rw-r--r-- | 
| 41959 | 1 | (* Title: HOL/Multivariate_Analysis/Determinants.thy | 
| 2 | Author: Amine Chaieb, University of Cambridge | |
| 33175 | 3 | *) | 
| 4 | ||
| 5 | header {* Traces, Determinant of square matrices and some properties *}
 | |
| 6 | ||
| 7 | theory Determinants | |
| 44228 
5f974bead436
get Multivariate_Analysis/Determinants.thy compiled and working again
 huffman parents: 
41959diff
changeset | 8 | imports | 
| 
5f974bead436
get Multivariate_Analysis/Determinants.thy compiled and working again
 huffman parents: 
41959diff
changeset | 9 | Cartesian_Euclidean_Space | 
| 
5f974bead436
get Multivariate_Analysis/Determinants.thy compiled and working again
 huffman parents: 
41959diff
changeset | 10 | "~~/src/HOL/Library/Permutations" | 
| 33175 | 11 | begin | 
| 12 | ||
| 13 | subsection{* First some facts about products*}
 | |
| 53253 | 14 | |
| 33175 | 15 | lemma setprod_add_split: | 
| 53854 | 16 | fixes m n :: nat | 
| 17 | assumes mn: "m \<le> n + 1" | |
| 18 |   shows "setprod f {m..n+p} = setprod f {m .. n} * setprod f {n+1..n+p}"
 | |
| 53253 | 19 | proof - | 
| 53854 | 20 |   let ?A = "{m..n+p}"
 | 
| 21 |   let ?B = "{m..n}"
 | |
| 33175 | 22 |   let ?C = "{n+1..n+p}"
 | 
| 53854 | 23 | from mn have un: "?B \<union> ?C = ?A" | 
| 24 | by auto | |
| 25 |   from mn have dj: "?B \<inter> ?C = {}"
 | |
| 26 | by auto | |
| 27 | have f: "finite ?B" "finite ?C" | |
| 28 | by simp_all | |
| 57418 | 29 | from setprod.union_disjoint[OF f dj, of f, unfolded un] show ?thesis . | 
| 33175 | 30 | qed | 
| 31 | ||
| 32 | ||
| 53854 | 33 | lemma setprod_offset: | 
| 34 | fixes m n :: nat | |
| 35 |   shows "setprod f {m + p .. n + p} = setprod (\<lambda>i. f (i + p)) {m..n}"
 | |
| 57129 
7edb7550663e
introduce more powerful reindexing rules for big operators
 hoelzl parents: 
56545diff
changeset | 36 | by (rule setprod.reindex_bij_witness[where i="op + p" and j="\<lambda>i. i - p"]) auto | 
| 33175 | 37 | |
| 53253 | 38 | lemma setprod_singleton: "setprod f {x} = f x"
 | 
| 39 | by simp | |
| 33175 | 40 | |
| 53854 | 41 | lemma setprod_singleton_nat_seg: | 
| 42 | fixes n :: "'a::order" | |
| 43 |   shows "setprod f {n..n} = f n"
 | |
| 53253 | 44 | by simp | 
| 45 | ||
| 46 | lemma setprod_numseg: | |
| 47 |   "setprod f {m..0} = (if m = 0 then f 0 else 1)"
 | |
| 48 |   "setprod f {m .. Suc n} =
 | |
| 49 |     (if m \<le> Suc n then f (Suc n) * setprod f {m..n} else setprod f {m..n})"
 | |
| 33175 | 50 | by (auto simp add: atLeastAtMostSuc_conv) | 
| 51 | ||
| 53253 | 52 | lemma setprod_le: | 
| 53854 | 53 | fixes f g :: "'b \<Rightarrow> 'a::linordered_idom" | 
| 53253 | 54 | assumes fS: "finite S" | 
| 53854 | 55 | and fg: "\<forall>x\<in>S. f x \<ge> 0 \<and> f x \<le> g x" | 
| 33175 | 56 | shows "setprod f S \<le> setprod g S" | 
| 53253 | 57 | using fS fg | 
| 58 | apply (induct S) | |
| 59 | apply simp | |
| 60 | apply auto | |
| 61 | apply (rule mult_mono) | |
| 62 | apply (auto intro: setprod_nonneg) | |
| 63 | done | |
| 33175 | 64 | |
| 53854 | 65 | (* FIXME: In Finite_Set there is a useless further assumption *) | 
| 53253 | 66 | lemma setprod_inversef: | 
| 67 | "finite A \<Longrightarrow> setprod (inverse \<circ> f) A = (inverse (setprod f A) :: 'a:: field_inverse_zero)" | |
| 33175 | 68 | apply (erule finite_induct) | 
| 69 | apply (simp) | |
| 70 | apply simp | |
| 71 | done | |
| 72 | ||
| 53253 | 73 | lemma setprod_le_1: | 
| 53854 | 74 | fixes f :: "'b \<Rightarrow> 'a::linordered_idom" | 
| 53253 | 75 | assumes fS: "finite S" | 
| 53854 | 76 | and f: "\<forall>x\<in>S. f x \<ge> 0 \<and> f x \<le> 1" | 
| 33175 | 77 | shows "setprod f S \<le> 1" | 
| 57418 | 78 | using setprod_le[OF fS f] unfolding setprod.neutral_const . | 
| 33175 | 79 | |
| 53253 | 80 | |
| 81 | subsection {* Trace *}
 | |
| 33175 | 82 | |
| 53253 | 83 | definition trace :: "'a::semiring_1^'n^'n \<Rightarrow> 'a" | 
| 84 | where "trace A = setsum (\<lambda>i. ((A$i)$i)) (UNIV::'n set)" | |
| 33175 | 85 | |
| 53854 | 86 | lemma trace_0: "trace (mat 0) = 0" | 
| 33175 | 87 | by (simp add: trace_def mat_def) | 
| 88 | ||
| 53854 | 89 | lemma trace_I: "trace (mat 1 :: 'a::semiring_1^'n^'n) = of_nat(CARD('n))"
 | 
| 33175 | 90 | by (simp add: trace_def mat_def) | 
| 91 | ||
| 34291 
4e896680897e
finite annotation on cartesian product is now implicit.
 hoelzl parents: 
34289diff
changeset | 92 | lemma trace_add: "trace ((A::'a::comm_semiring_1^'n^'n) + B) = trace A + trace B" | 
| 57418 | 93 | by (simp add: trace_def setsum.distrib) | 
| 33175 | 94 | |
| 34291 
4e896680897e
finite annotation on cartesian product is now implicit.
 hoelzl parents: 
34289diff
changeset | 95 | lemma trace_sub: "trace ((A::'a::comm_ring_1^'n^'n) - B) = trace A - trace B" | 
| 33175 | 96 | by (simp add: trace_def setsum_subtractf) | 
| 97 | ||
| 53854 | 98 | lemma trace_mul_sym: "trace ((A::'a::comm_semiring_1^'n^'m) ** B) = trace (B**A)" | 
| 33175 | 99 | apply (simp add: trace_def matrix_matrix_mult_def) | 
| 57418 | 100 | apply (subst setsum.commute) | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57418diff
changeset | 101 | apply (simp add: mult.commute) | 
| 53253 | 102 | done | 
| 33175 | 103 | |
| 53854 | 104 | text {* Definition of determinant. *}
 | 
| 33175 | 105 | |
| 34291 
4e896680897e
finite annotation on cartesian product is now implicit.
 hoelzl parents: 
34289diff
changeset | 106 | definition det:: "'a::comm_ring_1^'n^'n \<Rightarrow> 'a" where | 
| 53253 | 107 | "det A = | 
| 108 | setsum (\<lambda>p. of_int (sign p) * setprod (\<lambda>i. A$i$p i) (UNIV :: 'n set)) | |
| 109 |       {p. p permutes (UNIV :: 'n set)}"
 | |
| 33175 | 110 | |
| 53854 | 111 | text {* A few general lemmas we need below. *}
 | 
| 33175 | 112 | |
| 113 | lemma setprod_permute: | |
| 114 | assumes p: "p permutes S" | |
| 53854 | 115 | shows "setprod f S = setprod (f \<circ> p) S" | 
| 51489 | 116 | using assms by (fact setprod.permute) | 
| 33175 | 117 | |
| 53253 | 118 | lemma setproduct_permute_nat_interval: | 
| 53854 | 119 | fixes m n :: nat | 
| 120 |   shows "p permutes {m..n} \<Longrightarrow> setprod f {m..n} = setprod (f \<circ> p) {m..n}"
 | |
| 33175 | 121 | by (blast intro!: setprod_permute) | 
| 122 | ||
| 53854 | 123 | text {* Basic determinant properties. *}
 | 
| 33175 | 124 | |
| 35150 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 hoelzl parents: 
35028diff
changeset | 125 | lemma det_transpose: "det (transpose A) = det (A::'a::comm_ring_1 ^'n^'n)" | 
| 53253 | 126 | proof - | 
| 33175 | 127 | let ?di = "\<lambda>A i j. A$i$j" | 
| 128 | let ?U = "(UNIV :: 'n set)" | |
| 129 | have fU: "finite ?U" by simp | |
| 53253 | 130 |   {
 | 
| 131 | fix p | |
| 132 |     assume p: "p \<in> {p. p permutes ?U}"
 | |
| 53854 | 133 | from p have pU: "p permutes ?U" | 
| 134 | by blast | |
| 33175 | 135 | have sth: "sign (inv p) = sign p" | 
| 44260 
7784fa3232ce
Determinants.thy: avoid using mem_def/Collect_def
 huffman parents: 
44228diff
changeset | 136 | by (metis sign_inverse fU p mem_Collect_eq permutation_permutes) | 
| 33175 | 137 | from permutes_inj[OF pU] | 
| 53854 | 138 | have pi: "inj_on p ?U" | 
| 139 | by (blast intro: subset_inj_on) | |
| 33175 | 140 | from permutes_image[OF pU] | 
| 53253 | 141 | have "setprod (\<lambda>i. ?di (transpose A) i (inv p i)) ?U = | 
| 53854 | 142 | setprod (\<lambda>i. ?di (transpose A) i (inv p i)) (p ` ?U)" | 
| 143 | by simp | |
| 144 | also have "\<dots> = setprod ((\<lambda>i. ?di (transpose A) i (inv p i)) \<circ> p) ?U" | |
| 57418 | 145 | unfolding setprod.reindex[OF pi] .. | 
| 33175 | 146 | also have "\<dots> = setprod (\<lambda>i. ?di A i (p i)) ?U" | 
| 53253 | 147 | proof - | 
| 148 |       {
 | |
| 149 | fix i | |
| 150 | assume i: "i \<in> ?U" | |
| 33175 | 151 | from i permutes_inv_o[OF pU] permutes_in_image[OF pU] | 
| 53854 | 152 | have "((\<lambda>i. ?di (transpose A) i (inv p i)) \<circ> p) i = ?di A i (p i)" | 
| 53253 | 153 | unfolding transpose_def by (simp add: fun_eq_iff) | 
| 154 | } | |
| 53854 | 155 | then show "setprod ((\<lambda>i. ?di (transpose A) i (inv p i)) \<circ> p) ?U = | 
| 156 | setprod (\<lambda>i. ?di A i (p i)) ?U" | |
| 57418 | 157 | by (auto intro: setprod.cong) | 
| 33175 | 158 | qed | 
| 53253 | 159 | finally have "of_int (sign (inv p)) * (setprod (\<lambda>i. ?di (transpose A) i (inv p i)) ?U) = | 
| 53854 | 160 | of_int (sign p) * (setprod (\<lambda>i. ?di A i (p i)) ?U)" | 
| 161 | using sth by simp | |
| 53253 | 162 | } | 
| 163 | then show ?thesis | |
| 164 | unfolding det_def | |
| 165 | apply (subst setsum_permutations_inverse) | |
| 57418 | 166 | apply (rule setsum.cong) | 
| 167 | apply (rule refl) | |
| 53253 | 168 | apply blast | 
| 169 | done | |
| 33175 | 170 | qed | 
| 171 | ||
| 172 | lemma det_lowerdiagonal: | |
| 34291 
4e896680897e
finite annotation on cartesian product is now implicit.
 hoelzl parents: 
34289diff
changeset | 173 |   fixes A :: "'a::comm_ring_1^('n::{finite,wellorder})^('n::{finite,wellorder})"
 | 
| 33175 | 174 | assumes ld: "\<And>i j. i < j \<Longrightarrow> A$i$j = 0" | 
| 175 | shows "det A = setprod (\<lambda>i. A$i$i) (UNIV:: 'n set)" | |
| 53253 | 176 | proof - | 
| 33175 | 177 | let ?U = "UNIV:: 'n set" | 
| 178 |   let ?PU = "{p. p permutes ?U}"
 | |
| 179 | let ?pp = "\<lambda>p. of_int (sign p) * setprod (\<lambda>i. A$i$p i) (UNIV :: 'n set)" | |
| 53854 | 180 | have fU: "finite ?U" | 
| 181 | by simp | |
| 33175 | 182 | from finite_permutations[OF fU] have fPU: "finite ?PU" . | 
| 53854 | 183 |   have id0: "{id} \<subseteq> ?PU"
 | 
| 184 | by (auto simp add: permutes_id) | |
| 53253 | 185 |   {
 | 
| 186 | fix p | |
| 53854 | 187 |     assume p: "p \<in> ?PU - {id}"
 | 
| 53253 | 188 | from p have pU: "p permutes ?U" and pid: "p \<noteq> id" | 
| 189 | by blast+ | |
| 190 | from permutes_natset_le[OF pU] pid obtain i where i: "p i > i" | |
| 191 | by (metis not_le) | |
| 192 | from ld[OF i] have ex:"\<exists>i \<in> ?U. A$i$p i = 0" | |
| 193 | by blast | |
| 194 | from setprod_zero[OF fU ex] have "?pp p = 0" | |
| 195 | by simp | |
| 196 | } | |
| 53854 | 197 |   then have p0: "\<forall>p \<in> ?PU - {id}. ?pp p = 0"
 | 
| 53253 | 198 | by blast | 
| 57418 | 199 | from setsum.mono_neutral_cong_left[OF fPU id0 p0] show ?thesis | 
| 33175 | 200 | unfolding det_def by (simp add: sign_id) | 
| 201 | qed | |
| 202 | ||
| 203 | lemma det_upperdiagonal: | |
| 34291 
4e896680897e
finite annotation on cartesian product is now implicit.
 hoelzl parents: 
34289diff
changeset | 204 |   fixes A :: "'a::comm_ring_1^'n::{finite,wellorder}^'n::{finite,wellorder}"
 | 
| 33175 | 205 | assumes ld: "\<And>i j. i > j \<Longrightarrow> A$i$j = 0" | 
| 206 | shows "det A = setprod (\<lambda>i. A$i$i) (UNIV:: 'n set)" | |
| 53253 | 207 | proof - | 
| 33175 | 208 | let ?U = "UNIV:: 'n set" | 
| 209 |   let ?PU = "{p. p permutes ?U}"
 | |
| 210 | let ?pp = "(\<lambda>p. of_int (sign p) * setprod (\<lambda>i. A$i$p i) (UNIV :: 'n set))" | |
| 53854 | 211 | have fU: "finite ?U" | 
| 212 | by simp | |
| 33175 | 213 | from finite_permutations[OF fU] have fPU: "finite ?PU" . | 
| 53854 | 214 |   have id0: "{id} \<subseteq> ?PU"
 | 
| 215 | by (auto simp add: permutes_id) | |
| 53253 | 216 |   {
 | 
| 217 | fix p | |
| 53854 | 218 |     assume p: "p \<in> ?PU - {id}"
 | 
| 53253 | 219 | from p have pU: "p permutes ?U" and pid: "p \<noteq> id" | 
| 220 | by blast+ | |
| 221 | from permutes_natset_ge[OF pU] pid obtain i where i: "p i < i" | |
| 222 | by (metis not_le) | |
| 53854 | 223 | from ld[OF i] have ex:"\<exists>i \<in> ?U. A$i$p i = 0" | 
| 224 | by blast | |
| 225 | from setprod_zero[OF fU ex] have "?pp p = 0" | |
| 226 | by simp | |
| 53253 | 227 | } | 
| 228 |   then have p0: "\<forall>p \<in> ?PU -{id}. ?pp p = 0"
 | |
| 229 | by blast | |
| 57418 | 230 | from setsum.mono_neutral_cong_left[OF fPU id0 p0] show ?thesis | 
| 33175 | 231 | unfolding det_def by (simp add: sign_id) | 
| 232 | qed | |
| 233 | ||
| 234 | lemma det_diagonal: | |
| 34291 
4e896680897e
finite annotation on cartesian product is now implicit.
 hoelzl parents: 
34289diff
changeset | 235 | fixes A :: "'a::comm_ring_1^'n^'n" | 
| 33175 | 236 | assumes ld: "\<And>i j. i \<noteq> j \<Longrightarrow> A$i$j = 0" | 
| 237 | shows "det A = setprod (\<lambda>i. A$i$i) (UNIV::'n set)" | |
| 53253 | 238 | proof - | 
| 33175 | 239 | let ?U = "UNIV:: 'n set" | 
| 240 |   let ?PU = "{p. p permutes ?U}"
 | |
| 241 | let ?pp = "\<lambda>p. of_int (sign p) * setprod (\<lambda>i. A$i$p i) (UNIV :: 'n set)" | |
| 242 | have fU: "finite ?U" by simp | |
| 243 | from finite_permutations[OF fU] have fPU: "finite ?PU" . | |
| 53854 | 244 |   have id0: "{id} \<subseteq> ?PU"
 | 
| 245 | by (auto simp add: permutes_id) | |
| 53253 | 246 |   {
 | 
| 247 | fix p | |
| 248 |     assume p: "p \<in> ?PU - {id}"
 | |
| 53854 | 249 | then have "p \<noteq> id" | 
| 250 | by simp | |
| 251 | then obtain i where i: "p i \<noteq> i" | |
| 252 | unfolding fun_eq_iff by auto | |
| 253 | from ld [OF i [symmetric]] have ex:"\<exists>i \<in> ?U. A$i$p i = 0" | |
| 254 | by blast | |
| 255 | from setprod_zero [OF fU ex] have "?pp p = 0" | |
| 256 | by simp | |
| 257 | } | |
| 258 |   then have p0: "\<forall>p \<in> ?PU - {id}. ?pp p = 0"
 | |
| 259 | by blast | |
| 57418 | 260 | from setsum.mono_neutral_cong_left[OF fPU id0 p0] show ?thesis | 
| 33175 | 261 | unfolding det_def by (simp add: sign_id) | 
| 262 | qed | |
| 263 | ||
| 34291 
4e896680897e
finite annotation on cartesian product is now implicit.
 hoelzl parents: 
34289diff
changeset | 264 | lemma det_I: "det (mat 1 :: 'a::comm_ring_1^'n^'n) = 1" | 
| 53253 | 265 | proof - | 
| 33175 | 266 | let ?A = "mat 1 :: 'a::comm_ring_1^'n^'n" | 
| 267 | let ?U = "UNIV :: 'n set" | |
| 268 | let ?f = "\<lambda>i j. ?A$i$j" | |
| 53253 | 269 |   {
 | 
| 270 | fix i | |
| 271 | assume i: "i \<in> ?U" | |
| 53854 | 272 | have "?f i i = 1" | 
| 273 | using i by (vector mat_def) | |
| 53253 | 274 | } | 
| 275 | then have th: "setprod (\<lambda>i. ?f i i) ?U = setprod (\<lambda>x. 1) ?U" | |
| 57418 | 276 | by (auto intro: setprod.cong) | 
| 53253 | 277 |   {
 | 
| 278 | fix i j | |
| 279 | assume i: "i \<in> ?U" and j: "j \<in> ?U" and ij: "i \<noteq> j" | |
| 53854 | 280 | have "?f i j = 0" using i j ij | 
| 281 | by (vector mat_def) | |
| 53253 | 282 | } | 
| 53854 | 283 | then have "det ?A = setprod (\<lambda>i. ?f i i) ?U" | 
| 284 | using det_diagonal by blast | |
| 285 | also have "\<dots> = 1" | |
| 57418 | 286 | unfolding th setprod.neutral_const .. | 
| 33175 | 287 | finally show ?thesis . | 
| 288 | qed | |
| 289 | ||
| 34291 
4e896680897e
finite annotation on cartesian product is now implicit.
 hoelzl parents: 
34289diff
changeset | 290 | lemma det_0: "det (mat 0 :: 'a::comm_ring_1^'n^'n) = 0" | 
| 33175 | 291 | by (simp add: det_def setprod_zero) | 
| 292 | ||
| 293 | lemma det_permute_rows: | |
| 34291 
4e896680897e
finite annotation on cartesian product is now implicit.
 hoelzl parents: 
34289diff
changeset | 294 | fixes A :: "'a::comm_ring_1^'n^'n" | 
| 33175 | 295 | assumes p: "p permutes (UNIV :: 'n::finite set)" | 
| 53854 | 296 | shows "det (\<chi> i. A$p i :: 'a^'n^'n) = of_int (sign p) * det A" | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57418diff
changeset | 297 | apply (simp add: det_def setsum_right_distrib mult.assoc[symmetric]) | 
| 33175 | 298 | apply (subst sum_permutations_compose_right[OF p]) | 
| 57418 | 299 | proof (rule setsum.cong) | 
| 33175 | 300 | let ?U = "UNIV :: 'n set" | 
| 301 |   let ?PU = "{p. p permutes ?U}"
 | |
| 53253 | 302 | fix q | 
| 303 | assume qPU: "q \<in> ?PU" | |
| 53854 | 304 | have fU: "finite ?U" | 
| 305 | by simp | |
| 53253 | 306 | from qPU have q: "q permutes ?U" | 
| 307 | by blast | |
| 33175 | 308 | from p q have pp: "permutation p" and qp: "permutation q" | 
| 309 | by (metis fU permutation_permutes)+ | |
| 310 | from permutes_inv[OF p] have ip: "inv p permutes ?U" . | |
| 53854 | 311 | have "setprod (\<lambda>i. A$p i$ (q \<circ> p) i) ?U = setprod ((\<lambda>i. A$p i$(q \<circ> p) i) \<circ> inv p) ?U" | 
| 53253 | 312 | by (simp only: setprod_permute[OF ip, symmetric]) | 
| 53854 | 313 | also have "\<dots> = setprod (\<lambda>i. A $ (p \<circ> inv p) i $ (q \<circ> (p \<circ> inv p)) i) ?U" | 
| 53253 | 314 | by (simp only: o_def) | 
| 315 | also have "\<dots> = setprod (\<lambda>i. A$i$q i) ?U" | |
| 316 | by (simp only: o_def permutes_inverses[OF p]) | |
| 53854 | 317 | finally have thp: "setprod (\<lambda>i. A$p i$ (q \<circ> p) i) ?U = setprod (\<lambda>i. A$i$q i) ?U" | 
| 53253 | 318 | by blast | 
| 53854 | 319 | show "of_int (sign (q \<circ> p)) * setprod (\<lambda>i. A$ p i$ (q \<circ> p) i) ?U = | 
| 53253 | 320 | of_int (sign p) * of_int (sign q) * setprod (\<lambda>i. A$i$q i) ?U" | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57418diff
changeset | 321 | by (simp only: thp sign_compose[OF qp pp] mult.commute of_int_mult) | 
| 57418 | 322 | qed rule | 
| 33175 | 323 | |
| 324 | lemma det_permute_columns: | |
| 34291 
4e896680897e
finite annotation on cartesian product is now implicit.
 hoelzl parents: 
34289diff
changeset | 325 | fixes A :: "'a::comm_ring_1^'n^'n" | 
| 33175 | 326 | assumes p: "p permutes (UNIV :: 'n set)" | 
| 327 | shows "det(\<chi> i j. A$i$ p j :: 'a^'n^'n) = of_int (sign p) * det A" | |
| 53253 | 328 | proof - | 
| 33175 | 329 | let ?Ap = "\<chi> i j. A$i$ p j :: 'a^'n^'n" | 
| 35150 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 hoelzl parents: 
35028diff
changeset | 330 | let ?At = "transpose A" | 
| 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 hoelzl parents: 
35028diff
changeset | 331 | have "of_int (sign p) * det A = det (transpose (\<chi> i. transpose A $ p i))" | 
| 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 hoelzl parents: 
35028diff
changeset | 332 | unfolding det_permute_rows[OF p, of ?At] det_transpose .. | 
| 33175 | 333 | moreover | 
| 35150 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 hoelzl parents: 
35028diff
changeset | 334 | have "?Ap = transpose (\<chi> i. transpose A $ p i)" | 
| 44228 
5f974bead436
get Multivariate_Analysis/Determinants.thy compiled and working again
 huffman parents: 
41959diff
changeset | 335 | by (simp add: transpose_def vec_eq_iff) | 
| 53854 | 336 | ultimately show ?thesis | 
| 337 | by simp | |
| 33175 | 338 | qed | 
| 339 | ||
| 340 | lemma det_identical_rows: | |
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
34291diff
changeset | 341 | fixes A :: "'a::linordered_idom^'n^'n" | 
| 33175 | 342 | assumes ij: "i \<noteq> j" | 
| 53253 | 343 | and r: "row i A = row j A" | 
| 33175 | 344 | shows "det A = 0" | 
| 345 | proof- | |
| 53253 | 346 | have tha: "\<And>(a::'a) b. a = b \<Longrightarrow> b = - a \<Longrightarrow> a = 0" | 
| 33175 | 347 | by simp | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
44457diff
changeset | 348 | have th1: "of_int (-1) = - 1" by simp | 
| 33175 | 349 | let ?p = "Fun.swap i j id" | 
| 350 | let ?A = "\<chi> i. A $ ?p i" | |
| 56545 | 351 | from r have "A = ?A" by (simp add: vec_eq_iff row_def Fun.swap_def) | 
| 53253 | 352 | then have "det A = det ?A" by simp | 
| 33175 | 353 | moreover have "det A = - det ?A" | 
| 354 | by (simp add: det_permute_rows[OF permutes_swap_id] sign_swap_id ij th1) | |
| 355 | ultimately show "det A = 0" by (metis tha) | |
| 356 | qed | |
| 357 | ||
| 358 | lemma det_identical_columns: | |
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
34291diff
changeset | 359 | fixes A :: "'a::linordered_idom^'n^'n" | 
| 33175 | 360 | assumes ij: "i \<noteq> j" | 
| 53253 | 361 | and r: "column i A = column j A" | 
| 33175 | 362 | shows "det A = 0" | 
| 53253 | 363 | apply (subst det_transpose[symmetric]) | 
| 364 | apply (rule det_identical_rows[OF ij]) | |
| 365 | apply (metis row_transpose r) | |
| 366 | done | |
| 33175 | 367 | |
| 368 | lemma det_zero_row: | |
| 34291 
4e896680897e
finite annotation on cartesian product is now implicit.
 hoelzl parents: 
34289diff
changeset | 369 |   fixes A :: "'a::{idom, ring_char_0}^'n^'n"
 | 
| 33175 | 370 | assumes r: "row i A = 0" | 
| 371 | shows "det A = 0" | |
| 53253 | 372 | using r | 
| 373 | apply (simp add: row_def det_def vec_eq_iff) | |
| 57418 | 374 | apply (rule setsum.neutral) | 
| 53253 | 375 | apply (auto simp: sign_nz) | 
| 376 | done | |
| 33175 | 377 | |
| 378 | lemma det_zero_column: | |
| 34291 
4e896680897e
finite annotation on cartesian product is now implicit.
 hoelzl parents: 
34289diff
changeset | 379 |   fixes A :: "'a::{idom,ring_char_0}^'n^'n"
 | 
| 33175 | 380 | assumes r: "column i A = 0" | 
| 381 | shows "det A = 0" | |
| 35150 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 hoelzl parents: 
35028diff
changeset | 382 | apply (subst det_transpose[symmetric]) | 
| 33175 | 383 | apply (rule det_zero_row [of i]) | 
| 53253 | 384 | apply (metis row_transpose r) | 
| 385 | done | |
| 33175 | 386 | |
| 387 | lemma det_row_add: | |
| 388 | fixes a b c :: "'n::finite \<Rightarrow> _ ^ 'n" | |
| 389 | shows "det((\<chi> i. if i = k then a i + b i else c i)::'a::comm_ring_1^'n^'n) = | |
| 53253 | 390 | det((\<chi> i. if i = k then a i else c i)::'a::comm_ring_1^'n^'n) + | 
| 391 | det((\<chi> i. if i = k then b i else c i)::'a::comm_ring_1^'n^'n)" | |
| 57418 | 392 | unfolding det_def vec_lambda_beta setsum.distrib[symmetric] | 
| 393 | proof (rule setsum.cong) | |
| 33175 | 394 | let ?U = "UNIV :: 'n set" | 
| 395 |   let ?pU = "{p. p permutes ?U}"
 | |
| 396 | let ?f = "(\<lambda>i. if i = k then a i + b i else c i)::'n \<Rightarrow> 'a::comm_ring_1^'n" | |
| 397 | let ?g = "(\<lambda> i. if i = k then a i else c i)::'n \<Rightarrow> 'a::comm_ring_1^'n" | |
| 398 | let ?h = "(\<lambda> i. if i = k then b i else c i)::'n \<Rightarrow> 'a::comm_ring_1^'n" | |
| 53253 | 399 | fix p | 
| 400 | assume p: "p \<in> ?pU" | |
| 33175 | 401 |   let ?Uk = "?U - {k}"
 | 
| 53854 | 402 | from p have pU: "p permutes ?U" | 
| 403 | by blast | |
| 404 | have kU: "?U = insert k ?Uk" | |
| 405 | by blast | |
| 53253 | 406 |   {
 | 
| 407 | fix j | |
| 408 | assume j: "j \<in> ?Uk" | |
| 33175 | 409 | from j have "?f j $ p j = ?g j $ p j" and "?f j $ p j= ?h j $ p j" | 
| 53253 | 410 | by simp_all | 
| 411 | } | |
| 33175 | 412 | then have th1: "setprod (\<lambda>i. ?f i $ p i) ?Uk = setprod (\<lambda>i. ?g i $ p i) ?Uk" | 
| 413 | and th2: "setprod (\<lambda>i. ?f i $ p i) ?Uk = setprod (\<lambda>i. ?h i $ p i) ?Uk" | |
| 414 | apply - | |
| 57418 | 415 | apply (rule setprod.cong, simp_all)+ | 
| 33175 | 416 | done | 
| 53854 | 417 | have th3: "finite ?Uk" "k \<notin> ?Uk" | 
| 418 | by auto | |
| 33175 | 419 | have "setprod (\<lambda>i. ?f i $ p i) ?U = setprod (\<lambda>i. ?f i $ p i) (insert k ?Uk)" | 
| 420 | unfolding kU[symmetric] .. | |
| 53854 | 421 | also have "\<dots> = ?f k $ p k * setprod (\<lambda>i. ?f i $ p i) ?Uk" | 
| 57418 | 422 | apply (rule setprod.insert) | 
| 33175 | 423 | apply simp | 
| 53253 | 424 | apply blast | 
| 425 | done | |
| 426 | 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)" | |
| 427 | by (simp add: field_simps) | |
| 428 | 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)" | |
| 429 | by (metis th1 th2) | |
| 33175 | 430 | also have "\<dots> = setprod (\<lambda>i. ?g i $ p i) (insert k ?Uk) + setprod (\<lambda>i. ?h i $ p i) (insert k ?Uk)" | 
| 57418 | 431 | unfolding setprod.insert[OF th3] by simp | 
| 53854 | 432 | 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" | 
| 433 | unfolding kU[symmetric] . | |
| 53253 | 434 | then show "of_int (sign p) * setprod (\<lambda>i. ?f i $ p i) ?U = | 
| 435 | of_int (sign p) * setprod (\<lambda>i. ?g i $ p i) ?U + of_int (sign p) * setprod (\<lambda>i. ?h i $ p i) ?U" | |
| 36350 | 436 | by (simp add: field_simps) | 
| 57418 | 437 | qed rule | 
| 33175 | 438 | |
| 439 | lemma det_row_mul: | |
| 440 | fixes a b :: "'n::finite \<Rightarrow> _ ^ 'n" | |
| 441 | shows "det((\<chi> i. if i = k then c *s a i else b i)::'a::comm_ring_1^'n^'n) = | |
| 53253 | 442 | c * det((\<chi> i. if i = k then a i else b i)::'a::comm_ring_1^'n^'n)" | 
| 443 | unfolding det_def vec_lambda_beta setsum_right_distrib | |
| 57418 | 444 | proof (rule setsum.cong) | 
| 33175 | 445 | let ?U = "UNIV :: 'n set" | 
| 446 |   let ?pU = "{p. p permutes ?U}"
 | |
| 447 | let ?f = "(\<lambda>i. if i = k then c*s a i else b i)::'n \<Rightarrow> 'a::comm_ring_1^'n" | |
| 448 | let ?g = "(\<lambda> i. if i = k then a i else b i)::'n \<Rightarrow> 'a::comm_ring_1^'n" | |
| 53253 | 449 | fix p | 
| 450 | assume p: "p \<in> ?pU" | |
| 33175 | 451 |   let ?Uk = "?U - {k}"
 | 
| 53854 | 452 | from p have pU: "p permutes ?U" | 
| 453 | by blast | |
| 454 | have kU: "?U = insert k ?Uk" | |
| 455 | by blast | |
| 53253 | 456 |   {
 | 
| 457 | fix j | |
| 458 | assume j: "j \<in> ?Uk" | |
| 53854 | 459 | from j have "?f j $ p j = ?g j $ p j" | 
| 460 | by simp | |
| 53253 | 461 | } | 
| 33175 | 462 | then have th1: "setprod (\<lambda>i. ?f i $ p i) ?Uk = setprod (\<lambda>i. ?g i $ p i) ?Uk" | 
| 463 | apply - | |
| 57418 | 464 | apply (rule setprod.cong) | 
| 53253 | 465 | apply simp_all | 
| 33175 | 466 | done | 
| 53854 | 467 | have th3: "finite ?Uk" "k \<notin> ?Uk" | 
| 468 | by auto | |
| 33175 | 469 | have "setprod (\<lambda>i. ?f i $ p i) ?U = setprod (\<lambda>i. ?f i $ p i) (insert k ?Uk)" | 
| 470 | unfolding kU[symmetric] .. | |
| 471 | also have "\<dots> = ?f k $ p k * setprod (\<lambda>i. ?f i $ p i) ?Uk" | |
| 57418 | 472 | apply (rule setprod.insert) | 
| 33175 | 473 | apply simp | 
| 53253 | 474 | apply blast | 
| 475 | done | |
| 476 | also have "\<dots> = (c*s a k) $ p k * setprod (\<lambda>i. ?f i $ p i) ?Uk" | |
| 477 | by (simp add: field_simps) | |
| 33175 | 478 | also have "\<dots> = c* (a k $ p k * setprod (\<lambda>i. ?g i $ p i) ?Uk)" | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 479 | unfolding th1 by (simp add: ac_simps) | 
| 33175 | 480 | also have "\<dots> = c* (setprod (\<lambda>i. ?g i $ p i) (insert k ?Uk))" | 
| 57418 | 481 | unfolding setprod.insert[OF th3] by simp | 
| 53253 | 482 | finally have "setprod (\<lambda>i. ?f i $ p i) ?U = c* (setprod (\<lambda>i. ?g i $ p i) ?U)" | 
| 483 | unfolding kU[symmetric] . | |
| 484 | then show "of_int (sign p) * setprod (\<lambda>i. ?f i $ p i) ?U = | |
| 485 | c * (of_int (sign p) * setprod (\<lambda>i. ?g i $ p i) ?U)" | |
| 36350 | 486 | by (simp add: field_simps) | 
| 57418 | 487 | qed rule | 
| 33175 | 488 | |
| 489 | lemma det_row_0: | |
| 490 | fixes b :: "'n::finite \<Rightarrow> _ ^ 'n" | |
| 491 | shows "det((\<chi> i. if i = k then 0 else b i)::'a::comm_ring_1^'n^'n) = 0" | |
| 53253 | 492 | using det_row_mul[of k 0 "\<lambda>i. 1" b] | 
| 493 | apply simp | |
| 494 | apply (simp only: vector_smult_lzero) | |
| 495 | done | |
| 33175 | 496 | |
| 497 | lemma det_row_operation: | |
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
34291diff
changeset | 498 | fixes A :: "'a::linordered_idom^'n^'n" | 
| 33175 | 499 | assumes ij: "i \<noteq> j" | 
| 500 | shows "det (\<chi> k. if k = i then row i A + c *s row j A else row k A) = det A" | |
| 53253 | 501 | proof - | 
| 33175 | 502 | let ?Z = "(\<chi> k. if k = i then row j A else row k A) :: 'a ^'n^'n" | 
| 503 | have th: "row i ?Z = row j ?Z" by (vector row_def) | |
| 504 | have th2: "((\<chi> k. if k = i then row i A else row k A) :: 'a^'n^'n) = A" | |
| 505 | by (vector row_def) | |
| 506 | show ?thesis | |
| 507 | unfolding det_row_add [of i] det_row_mul[of i] det_identical_rows[OF ij th] th2 | |
| 508 | by simp | |
| 509 | qed | |
| 510 | ||
| 511 | lemma det_row_span: | |
| 36593 
fb69c8cd27bd
define linear algebra concepts using scaleR instead of (op *s); generalized many lemmas, though a few theorems that used to work on type int^'n are a bit less general
 huffman parents: 
36585diff
changeset | 512 | fixes A :: "real^'n^'n" | 
| 33175 | 513 |   assumes x: "x \<in> span {row j A |j. j \<noteq> i}"
 | 
| 514 | shows "det (\<chi> k. if k = i then row i A + x else row k A) = det A" | |
| 53253 | 515 | proof - | 
| 33175 | 516 | let ?U = "UNIV :: 'n set" | 
| 517 |   let ?S = "{row j A |j. j \<noteq> i}"
 | |
| 518 | let ?d = "\<lambda>x. det (\<chi> k. if k = i then x else row k A)" | |
| 519 | let ?P = "\<lambda>x. ?d (row i A + x) = det A" | |
| 53253 | 520 |   {
 | 
| 521 | fix k | |
| 53854 | 522 | have "(if k = i then row i A + 0 else row k A) = row k A" | 
| 523 | by simp | |
| 53253 | 524 | } | 
| 33175 | 525 | then have P0: "?P 0" | 
| 526 | apply - | |
| 527 | apply (rule cong[of det, OF refl]) | |
| 53253 | 528 | apply (vector row_def) | 
| 529 | done | |
| 33175 | 530 | moreover | 
| 53253 | 531 |   {
 | 
| 532 | fix c z y | |
| 533 | assume zS: "z \<in> ?S" and Py: "?P y" | |
| 53854 | 534 | from zS obtain j where j: "z = row j A" "i \<noteq> j" | 
| 535 | by blast | |
| 33175 | 536 | let ?w = "row i A + y" | 
| 53854 | 537 | have th0: "row i A + (c*s z + y) = ?w + c*s z" | 
| 538 | by vector | |
| 33175 | 539 | have thz: "?d z = 0" | 
| 540 | apply (rule det_identical_rows[OF j(2)]) | |
| 53253 | 541 | using j | 
| 542 | apply (vector row_def) | |
| 543 | done | |
| 544 | have "?d (row i A + (c*s z + y)) = ?d (?w + c*s z)" | |
| 545 | unfolding th0 .. | |
| 546 | then have "?P (c*s z + y)" | |
| 547 | unfolding thz Py det_row_mul[of i] det_row_add[of i] | |
| 548 | by simp | |
| 549 | } | |
| 33175 | 550 | ultimately show ?thesis | 
| 551 | apply - | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
47108diff
changeset | 552 | apply (rule span_induct_alt[of ?P ?S, OF P0, folded scalar_mult_eq_scaleR]) | 
| 33175 | 553 | apply blast | 
| 554 | apply (rule x) | |
| 555 | done | |
| 556 | qed | |
| 557 | ||
| 53854 | 558 | text {*
 | 
| 559 | May as well do this, though it's a bit unsatisfactory since it ignores | |
| 560 | exact duplicates by considering the rows/columns as a set. | |
| 561 | *} | |
| 33175 | 562 | |
| 563 | lemma det_dependent_rows: | |
| 36593 
fb69c8cd27bd
define linear algebra concepts using scaleR instead of (op *s); generalized many lemmas, though a few theorems that used to work on type int^'n are a bit less general
 huffman parents: 
36585diff
changeset | 564 | fixes A:: "real^'n^'n" | 
| 33175 | 565 | assumes d: "dependent (rows A)" | 
| 566 | shows "det A = 0" | |
| 53253 | 567 | proof - | 
| 33175 | 568 | let ?U = "UNIV :: 'n set" | 
| 569 |   from d obtain i where i: "row i A \<in> span (rows A - {row i A})"
 | |
| 570 | unfolding dependent_def rows_def by blast | |
| 53253 | 571 |   {
 | 
| 572 | fix j k | |
| 573 | assume jk: "j \<noteq> k" and c: "row j A = row k A" | |
| 574 | from det_identical_rows[OF jk c] have ?thesis . | |
| 575 | } | |
| 33175 | 576 | moreover | 
| 53253 | 577 |   {
 | 
| 578 | assume H: "\<And> i j. i \<noteq> j \<Longrightarrow> row i A \<noteq> row j A" | |
| 33175 | 579 |     have th0: "- row i A \<in> span {row j A|j. j \<noteq> i}"
 | 
| 580 | apply (rule span_neg) | |
| 581 | apply (rule set_rev_mp) | |
| 582 | apply (rule i) | |
| 583 | apply (rule span_mono) | |
| 53253 | 584 | using H i | 
| 585 | apply (auto simp add: rows_def) | |
| 586 | done | |
| 33175 | 587 | from det_row_span[OF th0] | 
| 588 | have "det A = det (\<chi> k. if k = i then 0 *s 1 else row k A)" | |
| 589 | unfolding right_minus vector_smult_lzero .. | |
| 36593 
fb69c8cd27bd
define linear algebra concepts using scaleR instead of (op *s); generalized many lemmas, though a few theorems that used to work on type int^'n are a bit less general
 huffman parents: 
36585diff
changeset | 590 | with det_row_mul[of i "0::real" "\<lambda>i. 1"] | 
| 53253 | 591 | have "det A = 0" by simp | 
| 592 | } | |
| 33175 | 593 | ultimately show ?thesis by blast | 
| 594 | qed | |
| 595 | ||
| 53253 | 596 | lemma det_dependent_columns: | 
| 597 | assumes d: "dependent (columns (A::real^'n^'n))" | |
| 598 | shows "det A = 0" | |
| 599 | by (metis d det_dependent_rows rows_transpose det_transpose) | |
| 33175 | 600 | |
| 53854 | 601 | text {* Multilinearity and the multiplication formula. *}
 | 
| 33175 | 602 | |
| 44228 
5f974bead436
get Multivariate_Analysis/Determinants.thy compiled and working again
 huffman parents: 
41959diff
changeset | 603 | lemma Cart_lambda_cong: "(\<And>x. f x = g x) \<Longrightarrow> (vec_lambda f::'a^'n) = (vec_lambda g :: 'a^'n)" | 
| 53253 | 604 | by (rule iffD1[OF vec_lambda_unique]) vector | 
| 33175 | 605 | |
| 606 | lemma det_linear_row_setsum: | |
| 607 | assumes fS: "finite S" | |
| 53253 | 608 | shows "det ((\<chi> i. if i = k then setsum (a i) S else c i)::'a::comm_ring_1^'n^'n) = | 
| 609 | setsum (\<lambda>j. det ((\<chi> i. if i = k then a i j else c i)::'a^'n^'n)) S" | |
| 610 | proof (induct rule: finite_induct[OF fS]) | |
| 611 | case 1 | |
| 612 | then show ?case | |
| 613 | apply simp | |
| 57418 | 614 | unfolding setsum.empty det_row_0[of k] | 
| 53253 | 615 | apply rule | 
| 616 | done | |
| 33175 | 617 | next | 
| 618 | case (2 x F) | |
| 53253 | 619 | then show ?case | 
| 620 | by (simp add: det_row_add cong del: if_weak_cong) | |
| 33175 | 621 | qed | 
| 622 | ||
| 623 | lemma finite_bounded_functions: | |
| 624 | assumes fS: "finite S" | |
| 625 |   shows "finite {f. (\<forall>i \<in> {1.. (k::nat)}. f i \<in> S) \<and> (\<forall>i. i \<notin> {1 .. k} \<longrightarrow> f i = i)}"
 | |
| 53253 | 626 | proof (induct k) | 
| 33175 | 627 | case 0 | 
| 53854 | 628 |   have th: "{f. \<forall>i. f i = i} = {id}"
 | 
| 629 | by auto | |
| 630 | show ?case | |
| 631 | by (auto simp add: th) | |
| 33175 | 632 | next | 
| 633 | case (Suc k) | |
| 634 | let ?f = "\<lambda>(y::nat,g) i. if i = Suc k then y else g i" | |
| 635 |   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)})"
 | |
| 636 |   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)}"
 | |
| 637 | apply (auto simp add: image_iff) | |
| 638 | apply (rule_tac x="x (Suc k)" in bexI) | |
| 639 | apply (rule_tac x = "\<lambda>i. if i = Suc k then i else x i" in exI) | |
| 44457 
d366fa5551ef
declare euclidean_simps [simp] at the point they are proved;
 huffman parents: 
44260diff
changeset | 640 | apply auto | 
| 33175 | 641 | done | 
| 642 | with finite_imageI[OF finite_cartesian_product[OF fS Suc.hyps(1)], of ?f] | |
| 53854 | 643 | show ?case | 
| 644 | by metis | |
| 33175 | 645 | qed | 
| 646 | ||
| 647 | ||
| 53854 | 648 | lemma eq_id_iff[simp]: "(\<forall>x. f x = x) \<longleftrightarrow> f = id" | 
| 649 | by auto | |
| 33175 | 650 | |
| 651 | lemma det_linear_rows_setsum_lemma: | |
| 53854 | 652 | assumes fS: "finite S" | 
| 653 | and fT: "finite T" | |
| 654 | shows "det ((\<chi> i. if i \<in> T then setsum (a i) S else c i):: 'a::comm_ring_1^'n^'n) = | |
| 53253 | 655 | setsum (\<lambda>f. det((\<chi> i. if i \<in> T then a i (f i) else c i)::'a^'n^'n)) | 
| 656 |       {f. (\<forall>i \<in> T. f i \<in> S) \<and> (\<forall>i. i \<notin> T \<longrightarrow> f i = i)}"
 | |
| 657 | using fT | |
| 658 | proof (induct T arbitrary: a c set: finite) | |
| 33175 | 659 | case empty | 
| 53253 | 660 |   have th0: "\<And>x y. (\<chi> i. if i \<in> {} then x i else y i) = (\<chi> i. y i)"
 | 
| 661 | by vector | |
| 53854 | 662 | from empty.prems show ?case | 
| 663 | unfolding th0 by simp | |
| 33175 | 664 | next | 
| 665 | case (insert z T a c) | |
| 666 |   let ?F = "\<lambda>T. {f. (\<forall>i \<in> T. f i \<in> S) \<and> (\<forall>i. i \<notin> T \<longrightarrow> f i = i)}"
 | |
| 667 | let ?h = "\<lambda>(y,g) i. if i = z then y else g i" | |
| 668 | let ?k = "\<lambda>h. (h(z),(\<lambda>i. if i = z then i else h i))" | |
| 669 | let ?s = "\<lambda> k a c f. det((\<chi> i. if i \<in> T then a i (f i) else c i)::'a^'n^'n)" | |
| 57129 
7edb7550663e
introduce more powerful reindexing rules for big operators
 hoelzl parents: 
56545diff
changeset | 670 | let ?c = "\<lambda>j i. if i = z then a i j else c i" | 
| 53253 | 671 | have thif: "\<And>a b c d. (if a \<or> b then c else d) = (if a then c else if b then c else d)" | 
| 672 | by simp | |
| 33175 | 673 | have thif2: "\<And>a b c d e. (if a then b else if c then d else e) = | 
| 53253 | 674 | (if c then (if a then b else d) else (if a then b else e))" | 
| 675 | by simp | |
| 676 | from `z \<notin> T` have nz: "\<And>i. i \<in> T \<Longrightarrow> i = z \<longleftrightarrow> False" | |
| 677 | by auto | |
| 33175 | 678 | have "det (\<chi> i. if i \<in> insert z T then setsum (a i) S else c i) = | 
| 53253 | 679 | det (\<chi> i. if i = z then setsum (a i) S else if i \<in> T then setsum (a i) S else c i)" | 
| 33175 | 680 | unfolding insert_iff thif .. | 
| 53253 | 681 | also have "\<dots> = (\<Sum>j\<in>S. det (\<chi> i. if i \<in> T then setsum (a i) S else if i = z then a i j else c i))" | 
| 33175 | 682 | unfolding det_linear_row_setsum[OF fS] | 
| 683 | apply (subst thif2) | |
| 53253 | 684 | using nz | 
| 685 | apply (simp cong del: if_weak_cong cong add: if_cong) | |
| 686 | done | |
| 33175 | 687 | finally have tha: | 
| 688 | "det (\<chi> i. if i \<in> insert z T then setsum (a i) S else c i) = | |
| 689 | (\<Sum>(j, f)\<in>S \<times> ?F T. det (\<chi> i. if i \<in> T then a i (f i) | |
| 690 | else if i = z then a i j | |
| 691 | else c i))" | |
| 57418 | 692 | unfolding insert.hyps unfolding setsum.cartesian_product by blast | 
| 33175 | 693 | show ?case unfolding tha | 
| 694 | using `z \<notin> T` | |
| 57129 
7edb7550663e
introduce more powerful reindexing rules for big operators
 hoelzl parents: 
56545diff
changeset | 695 | by (intro setsum.reindex_bij_witness[where i="?k" and j="?h"]) | 
| 
7edb7550663e
introduce more powerful reindexing rules for big operators
 hoelzl parents: 
56545diff
changeset | 696 | (auto intro!: cong[OF refl[of det]] simp: vec_eq_iff) | 
| 33175 | 697 | qed | 
| 698 | ||
| 699 | lemma det_linear_rows_setsum: | |
| 53854 | 700 | fixes S :: "'n::finite set" | 
| 701 | assumes fS: "finite S" | |
| 53253 | 702 | shows "det (\<chi> i. setsum (a i) S) = | 
| 703 |     setsum (\<lambda>f. det (\<chi> i. a i (f i) :: 'a::comm_ring_1 ^ 'n^'n)) {f. \<forall>i. f i \<in> S}"
 | |
| 704 | proof - | |
| 705 | have th0: "\<And>x y. ((\<chi> i. if i \<in> (UNIV:: 'n set) then x i else y i) :: 'a^'n^'n) = (\<chi> i. x i)" | |
| 706 | by vector | |
| 707 | from det_linear_rows_setsum_lemma[OF fS, of "UNIV :: 'n set" a, unfolded th0, OF finite] | |
| 708 | show ?thesis by simp | |
| 33175 | 709 | qed | 
| 710 | ||
| 711 | lemma matrix_mul_setsum_alt: | |
| 34291 
4e896680897e
finite annotation on cartesian product is now implicit.
 hoelzl parents: 
34289diff
changeset | 712 | fixes A B :: "'a::comm_ring_1^'n^'n" | 
| 33175 | 713 | shows "A ** B = (\<chi> i. setsum (\<lambda>k. A$i$k *s B $ k) (UNIV :: 'n set))" | 
| 714 | by (vector matrix_matrix_mult_def setsum_component) | |
| 715 | ||
| 716 | lemma det_rows_mul: | |
| 34291 
4e896680897e
finite annotation on cartesian product is now implicit.
 hoelzl parents: 
34289diff
changeset | 717 | "det((\<chi> i. c i *s a i)::'a::comm_ring_1^'n^'n) = | 
| 53253 | 718 | setprod (\<lambda>i. c i) (UNIV:: 'n set) * det((\<chi> i. a i)::'a^'n^'n)" | 
| 57418 | 719 | proof (simp add: det_def setsum_right_distrib cong add: setprod.cong, rule setsum.cong) | 
| 33175 | 720 | let ?U = "UNIV :: 'n set" | 
| 721 |   let ?PU = "{p. p permutes ?U}"
 | |
| 53253 | 722 | fix p | 
| 723 | assume pU: "p \<in> ?PU" | |
| 33175 | 724 | let ?s = "of_int (sign p)" | 
| 53253 | 725 | from pU have p: "p permutes ?U" | 
| 726 | by blast | |
| 33175 | 727 | have "setprod (\<lambda>i. c i * a i $ p i) ?U = setprod c ?U * setprod (\<lambda>i. a i $ p i) ?U" | 
| 57418 | 728 | unfolding setprod.distrib .. | 
| 33175 | 729 | then show "?s * (\<Prod>xa\<in>?U. c xa * a xa $ p xa) = | 
| 53854 | 730 | setprod c ?U * (?s* (\<Prod>xa\<in>?U. a xa $ p xa))" | 
| 731 | by (simp add: field_simps) | |
| 57418 | 732 | qed rule | 
| 33175 | 733 | |
| 734 | lemma det_mul: | |
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
34291diff
changeset | 735 | fixes A B :: "'a::linordered_idom^'n^'n" | 
| 33175 | 736 | shows "det (A ** B) = det A * det B" | 
| 53253 | 737 | proof - | 
| 33175 | 738 | let ?U = "UNIV :: 'n set" | 
| 739 |   let ?F = "{f. (\<forall>i\<in> ?U. f i \<in> ?U) \<and> (\<forall>i. i \<notin> ?U \<longrightarrow> f i = i)}"
 | |
| 740 |   let ?PU = "{p. p permutes ?U}"
 | |
| 53854 | 741 | have fU: "finite ?U" | 
| 742 | by simp | |
| 743 | have fF: "finite ?F" | |
| 744 | by (rule finite) | |
| 53253 | 745 |   {
 | 
| 746 | fix p | |
| 747 | assume p: "p permutes ?U" | |
| 33175 | 748 | have "p \<in> ?F" unfolding mem_Collect_eq permutes_in_image[OF p] | 
| 53253 | 749 | using p[unfolded permutes_def] by simp | 
| 750 | } | |
| 53854 | 751 | then have PUF: "?PU \<subseteq> ?F" by blast | 
| 53253 | 752 |   {
 | 
| 753 | fix f | |
| 754 | assume fPU: "f \<in> ?F - ?PU" | |
| 53854 | 755 | have fUU: "f ` ?U \<subseteq> ?U" | 
| 756 | using fPU by auto | |
| 53253 | 757 | from fPU have f: "\<forall>i \<in> ?U. f i \<in> ?U" "\<forall>i. i \<notin> ?U \<longrightarrow> f i = i" "\<not>(\<forall>y. \<exists>!x. f x = y)" | 
| 758 | unfolding permutes_def by auto | |
| 33175 | 759 | |
| 760 | let ?A = "(\<chi> i. A$i$f i *s B$f i) :: 'a^'n^'n" | |
| 761 | let ?B = "(\<chi> i. B$f i) :: 'a^'n^'n" | |
| 53253 | 762 |     {
 | 
| 763 | assume fni: "\<not> inj_on f ?U" | |
| 33175 | 764 | then obtain i j where ij: "f i = f j" "i \<noteq> j" | 
| 765 | unfolding inj_on_def by blast | |
| 766 | from ij | |
| 53854 | 767 | have rth: "row i ?B = row j ?B" | 
| 768 | by (vector row_def) | |
| 33175 | 769 | from det_identical_rows[OF ij(2) rth] | 
| 770 | have "det (\<chi> i. A$i$f i *s B$f i) = 0" | |
| 53253 | 771 | unfolding det_rows_mul by simp | 
| 772 | } | |
| 33175 | 773 | moreover | 
| 53253 | 774 |     {
 | 
| 775 | assume fi: "inj_on f ?U" | |
| 33175 | 776 | from f fi have fith: "\<And>i j. f i = f j \<Longrightarrow> i = j" | 
| 777 | unfolding inj_on_def by metis | |
| 778 | note fs = fi[unfolded surjective_iff_injective_gen[OF fU fU refl fUU, symmetric]] | |
| 53253 | 779 |       {
 | 
| 780 | fix y | |
| 53854 | 781 | from fs f have "\<exists>x. f x = y" | 
| 782 | by blast | |
| 783 | then obtain x where x: "f x = y" | |
| 784 | by blast | |
| 53253 | 785 |         {
 | 
| 786 | fix z | |
| 787 | assume z: "f z = y" | |
| 53854 | 788 | from fith x z have "z = x" | 
| 789 | by metis | |
| 53253 | 790 | } | 
| 53854 | 791 | with x have "\<exists>!x. f x = y" | 
| 792 | by blast | |
| 53253 | 793 | } | 
| 53854 | 794 | with f(3) have "det (\<chi> i. A$i$f i *s B$f i) = 0" | 
| 795 | by blast | |
| 53253 | 796 | } | 
| 53854 | 797 | ultimately have "det (\<chi> i. A$i$f i *s B$f i) = 0" | 
| 798 | by blast | |
| 53253 | 799 | } | 
| 53854 | 800 | then have zth: "\<forall> f\<in> ?F - ?PU. det (\<chi> i. A$i$f i *s B$f i) = 0" | 
| 53253 | 801 | by simp | 
| 802 |   {
 | |
| 803 | fix p | |
| 804 | assume pU: "p \<in> ?PU" | |
| 53854 | 805 | from pU have p: "p permutes ?U" | 
| 806 | by blast | |
| 33175 | 807 | let ?s = "\<lambda>p. of_int (sign p)" | 
| 53253 | 808 | let ?f = "\<lambda>q. ?s p * (\<Prod>i\<in> ?U. A $ i $ p i) * (?s q * (\<Prod>i\<in> ?U. B $ i $ q i))" | 
| 33175 | 809 | have "(setsum (\<lambda>q. ?s q * | 
| 53253 | 810 | (\<Prod>i\<in> ?U. (\<chi> i. A $ i $ p i *s B $ p i :: 'a^'n^'n) $ i $ q i)) ?PU) = | 
| 811 | (setsum (\<lambda>q. ?s p * (\<Prod>i\<in> ?U. A $ i $ p i) * (?s q * (\<Prod>i\<in> ?U. B $ i $ q i))) ?PU)" | |
| 33175 | 812 | unfolding sum_permutations_compose_right[OF permutes_inv[OF p], of ?f] | 
| 57418 | 813 | proof (rule setsum.cong) | 
| 53253 | 814 | fix q | 
| 815 | assume qU: "q \<in> ?PU" | |
| 53854 | 816 | then have q: "q permutes ?U" | 
| 817 | by blast | |
| 33175 | 818 | from p q have pp: "permutation p" and pq: "permutation q" | 
| 819 | unfolding permutation_permutes by auto | |
| 820 | have th00: "of_int (sign p) * of_int (sign p) = (1::'a)" | |
| 821 | "\<And>a. of_int (sign p) * (of_int (sign p) * a) = a" | |
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57418diff
changeset | 822 | unfolding mult.assoc[symmetric] | 
| 53854 | 823 | unfolding of_int_mult[symmetric] | 
| 33175 | 824 | by (simp_all add: sign_idempotent) | 
| 53854 | 825 | have ths: "?s q = ?s p * ?s (q \<circ> inv p)" | 
| 33175 | 826 | using pp pq permutation_inverse[OF pp] sign_inverse[OF pp] | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 827 | by (simp add: th00 ac_simps sign_idempotent sign_compose) | 
| 53854 | 828 | have th001: "setprod (\<lambda>i. B$i$ q (inv p i)) ?U = setprod ((\<lambda>i. B$i$ q (inv p i)) \<circ> p) ?U" | 
| 33175 | 829 | by (rule setprod_permute[OF p]) | 
| 53253 | 830 | have thp: "setprod (\<lambda>i. (\<chi> i. A$i$p i *s B$p i :: 'a^'n^'n) $i $ q i) ?U = | 
| 831 | setprod (\<lambda>i. A$i$p i) ?U * setprod (\<lambda>i. B$i$ q (inv p i)) ?U" | |
| 57418 | 832 | unfolding th001 setprod.distrib[symmetric] o_def permutes_inverses[OF p] | 
| 833 | apply (rule setprod.cong[OF refl]) | |
| 53253 | 834 | using permutes_in_image[OF q] | 
| 835 | apply vector | |
| 836 | done | |
| 837 | show "?s q * setprod (\<lambda>i. (((\<chi> i. A$i$p i *s B$p i) :: 'a^'n^'n)$i$q i)) ?U = | |
| 53854 | 838 | ?s p * (setprod (\<lambda>i. A$i$p i) ?U) * (?s (q \<circ> inv p) * setprod (\<lambda>i. B$i$(q \<circ> inv p) i) ?U)" | 
| 33175 | 839 | using ths thp pp pq permutation_inverse[OF pp] sign_inverse[OF pp] | 
| 36350 | 840 | by (simp add: sign_nz th00 field_simps sign_idempotent sign_compose) | 
| 57418 | 841 | qed rule | 
| 33175 | 842 | } | 
| 843 | then have th2: "setsum (\<lambda>f. det (\<chi> i. A$i$f i *s B$f i)) ?PU = det A * det B" | |
| 844 | unfolding det_def setsum_product | |
| 57418 | 845 | by (rule setsum.cong [OF refl]) | 
| 33175 | 846 | have "det (A**B) = setsum (\<lambda>f. det (\<chi> i. A $ i $ f i *s B $ f i)) ?F" | 
| 53854 | 847 | unfolding matrix_mul_setsum_alt det_linear_rows_setsum[OF fU] | 
| 848 | by simp | |
| 33175 | 849 | also have "\<dots> = setsum (\<lambda>f. det (\<chi> i. A$i$f i *s B$f i)) ?PU" | 
| 57418 | 850 | using setsum.mono_neutral_cong_left[OF fF PUF zth, symmetric] | 
| 33175 | 851 | unfolding det_rows_mul by auto | 
| 852 | finally show ?thesis unfolding th2 . | |
| 853 | qed | |
| 854 | ||
| 53854 | 855 | text {* Relation to invertibility. *}
 | 
| 33175 | 856 | |
| 857 | lemma invertible_left_inverse: | |
| 34291 
4e896680897e
finite annotation on cartesian product is now implicit.
 hoelzl parents: 
34289diff
changeset | 858 | fixes A :: "real^'n^'n" | 
| 33175 | 859 | shows "invertible A \<longleftrightarrow> (\<exists>(B::real^'n^'n). B** A = mat 1)" | 
| 860 | by (metis invertible_def matrix_left_right_inverse) | |
| 861 | ||
| 862 | lemma invertible_righ_inverse: | |
| 34291 
4e896680897e
finite annotation on cartesian product is now implicit.
 hoelzl parents: 
34289diff
changeset | 863 | fixes A :: "real^'n^'n" | 
| 33175 | 864 | shows "invertible A \<longleftrightarrow> (\<exists>(B::real^'n^'n). A** B = mat 1)" | 
| 865 | by (metis invertible_def matrix_left_right_inverse) | |
| 866 | ||
| 867 | lemma invertible_det_nz: | |
| 34291 
4e896680897e
finite annotation on cartesian product is now implicit.
 hoelzl parents: 
34289diff
changeset | 868 | fixes A::"real ^'n^'n" | 
| 33175 | 869 | shows "invertible A \<longleftrightarrow> det A \<noteq> 0" | 
| 53253 | 870 | proof - | 
| 871 |   {
 | |
| 872 | assume "invertible A" | |
| 33175 | 873 | then obtain B :: "real ^'n^'n" where B: "A ** B = mat 1" | 
| 874 | unfolding invertible_righ_inverse by blast | |
| 53854 | 875 | then have "det (A ** B) = det (mat 1 :: real ^'n^'n)" | 
| 876 | by simp | |
| 877 | then have "det A \<noteq> 0" | |
| 878 | by (simp add: det_mul det_I) algebra | |
| 53253 | 879 | } | 
| 33175 | 880 | moreover | 
| 53253 | 881 |   {
 | 
| 882 | assume H: "\<not> invertible A" | |
| 33175 | 883 | let ?U = "UNIV :: 'n set" | 
| 53854 | 884 | have fU: "finite ?U" | 
| 885 | by simp | |
| 33175 | 886 | from H obtain c i where c: "setsum (\<lambda>i. c i *s row i A) ?U = 0" | 
| 53854 | 887 | and iU: "i \<in> ?U" | 
| 888 | and ci: "c i \<noteq> 0" | |
| 33175 | 889 | unfolding invertible_righ_inverse | 
| 53854 | 890 | unfolding matrix_right_invertible_independent_rows | 
| 891 | by blast | |
| 53253 | 892 | have *: "\<And>(a::real^'n) b. a + b = 0 \<Longrightarrow> -a = b" | 
| 33175 | 893 | apply (drule_tac f="op + (- a)" in cong[OF refl]) | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57418diff
changeset | 894 | apply (simp only: ab_left_minus add.assoc[symmetric]) | 
| 33175 | 895 | apply simp | 
| 896 | done | |
| 897 | from c ci | |
| 898 |     have thr0: "- row i A = setsum (\<lambda>j. (1/ c i) *s (c j *s row j A)) (?U - {i})"
 | |
| 57418 | 899 | unfolding setsum.remove[OF fU iU] setsum_cmul | 
| 33175 | 900 | apply - | 
| 901 | apply (rule vector_mul_lcancel_imp[OF ci]) | |
| 44457 
d366fa5551ef
declare euclidean_simps [simp] at the point they are proved;
 huffman parents: 
44260diff
changeset | 902 | apply (auto simp add: field_simps) | 
| 53854 | 903 | unfolding * | 
| 904 | apply rule | |
| 905 | done | |
| 33175 | 906 |     have thr: "- row i A \<in> span {row j A| j. j \<noteq> i}"
 | 
| 907 | unfolding thr0 | |
| 908 | apply (rule span_setsum) | |
| 909 | apply simp | |
| 910 | apply (rule ballI) | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
47108diff
changeset | 911 | apply (rule span_mul [where 'a="real^'n", folded scalar_mult_eq_scaleR])+ | 
| 33175 | 912 | apply (rule span_superset) | 
| 913 | apply auto | |
| 914 | done | |
| 915 | let ?B = "(\<chi> k. if k = i then 0 else row k A) :: real ^'n^'n" | |
| 916 | have thrb: "row i ?B = 0" using iU by (vector row_def) | |
| 917 | have "det A = 0" | |
| 918 | unfolding det_row_span[OF thr, symmetric] right_minus | |
| 53253 | 919 | unfolding det_zero_row[OF thrb] .. | 
| 920 | } | |
| 53854 | 921 | ultimately show ?thesis | 
| 922 | by blast | |
| 33175 | 923 | qed | 
| 924 | ||
| 53854 | 925 | text {* Cramer's rule. *}
 | 
| 33175 | 926 | |
| 35150 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 hoelzl parents: 
35028diff
changeset | 927 | lemma cramer_lemma_transpose: | 
| 53854 | 928 | fixes A:: "real^'n^'n" | 
| 929 | and x :: "real^'n" | |
| 33175 | 930 | shows "det ((\<chi> i. if i = k then setsum (\<lambda>i. x$i *s row i A) (UNIV::'n set) | 
| 53854 | 931 | else row i A)::real^'n^'n) = x$k * det A" | 
| 33175 | 932 | (is "?lhs = ?rhs") | 
| 53253 | 933 | proof - | 
| 33175 | 934 | let ?U = "UNIV :: 'n set" | 
| 935 |   let ?Uk = "?U - {k}"
 | |
| 53854 | 936 | have U: "?U = insert k ?Uk" | 
| 937 | by blast | |
| 938 | have fUk: "finite ?Uk" | |
| 939 | by simp | |
| 940 | have kUk: "k \<notin> ?Uk" | |
| 941 | by simp | |
| 33175 | 942 | have th00: "\<And>k s. x$k *s row k A + s = (x$k - 1) *s row k A + row k A + s" | 
| 36350 | 943 | by (vector field_simps) | 
| 53854 | 944 | have th001: "\<And>f k . (\<lambda>x. if x = k then f k else f x) = f" | 
| 945 | by auto | |
| 33175 | 946 | have "(\<chi> i. row i A) = A" by (vector row_def) | 
| 53253 | 947 | then have thd1: "det (\<chi> i. row i A) = det A" | 
| 948 | by simp | |
| 33175 | 949 | 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" | 
| 950 | apply (rule det_row_span) | |
| 56196 
32b7eafc5a52
remove unnecessary finiteness assumptions from lemmas about setsum
 huffman parents: 
53854diff
changeset | 951 | apply (rule span_setsum) | 
| 33175 | 952 | apply (rule ballI) | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
47108diff
changeset | 953 | apply (rule span_mul [where 'a="real^'n", folded scalar_mult_eq_scaleR])+ | 
| 33175 | 954 | apply (rule span_superset) | 
| 955 | apply auto | |
| 956 | done | |
| 957 | show "?lhs = x$k * det A" | |
| 958 | apply (subst U) | |
| 57418 | 959 | unfolding setsum.insert[OF fUk kUk] | 
| 33175 | 960 | apply (subst th00) | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57418diff
changeset | 961 | unfolding add.assoc | 
| 33175 | 962 | apply (subst det_row_add) | 
| 963 | unfolding thd0 | |
| 964 | unfolding det_row_mul | |
| 965 | unfolding th001[of k "\<lambda>i. row i A"] | |
| 53253 | 966 | unfolding thd1 | 
| 967 | apply (simp add: field_simps) | |
| 968 | done | |
| 33175 | 969 | qed | 
| 970 | ||
| 971 | lemma cramer_lemma: | |
| 36593 
fb69c8cd27bd
define linear algebra concepts using scaleR instead of (op *s); generalized many lemmas, though a few theorems that used to work on type int^'n are a bit less general
 huffman parents: 
36585diff
changeset | 972 | fixes A :: "real^'n^'n" | 
| 
fb69c8cd27bd
define linear algebra concepts using scaleR instead of (op *s); generalized many lemmas, though a few theorems that used to work on type int^'n are a bit less general
 huffman parents: 
36585diff
changeset | 973 | shows "det((\<chi> i j. if j = k then (A *v x)$i else A$i$j):: real^'n^'n) = x$k * det A" | 
| 53253 | 974 | proof - | 
| 33175 | 975 | let ?U = "UNIV :: 'n set" | 
| 53253 | 976 | have *: "\<And>c. setsum (\<lambda>i. c i *s row i (transpose A)) ?U = setsum (\<lambda>i. c i *s column i A) ?U" | 
| 57418 | 977 | by (auto simp add: row_transpose intro: setsum.cong) | 
| 53854 | 978 | show ?thesis | 
| 979 | unfolding matrix_mult_vsum | |
| 53253 | 980 | unfolding cramer_lemma_transpose[of k x "transpose A", unfolded det_transpose, symmetric] | 
| 981 | unfolding *[of "\<lambda>i. x$i"] | |
| 982 | apply (subst det_transpose[symmetric]) | |
| 983 | apply (rule cong[OF refl[of det]]) | |
| 984 | apply (vector transpose_def column_def row_def) | |
| 985 | done | |
| 33175 | 986 | qed | 
| 987 | ||
| 988 | lemma cramer: | |
| 34291 
4e896680897e
finite annotation on cartesian product is now implicit.
 hoelzl parents: 
34289diff
changeset | 989 | fixes A ::"real^'n^'n" | 
| 33175 | 990 | assumes d0: "det A \<noteq> 0" | 
| 36362 
06475a1547cb
fix lots of looping simp calls and other warnings
 huffman parents: 
35542diff
changeset | 991 | shows "A *v x = b \<longleftrightarrow> x = (\<chi> k. det(\<chi> i j. if j=k then b$i else A$i$j) / det A)" | 
| 53253 | 992 | proof - | 
| 33175 | 993 | from d0 obtain B where B: "A ** B = mat 1" "B ** A = mat 1" | 
| 53854 | 994 | unfolding invertible_det_nz[symmetric] invertible_def | 
| 995 | by blast | |
| 996 | have "(A ** B) *v b = b" | |
| 997 | by (simp add: B matrix_vector_mul_lid) | |
| 998 | then have "A *v (B *v b) = b" | |
| 999 | by (simp add: matrix_vector_mul_assoc) | |
| 1000 | then have xe: "\<exists>x. A *v x = b" | |
| 1001 | by blast | |
| 53253 | 1002 |   {
 | 
| 1003 | fix x | |
| 1004 | assume x: "A *v x = b" | |
| 1005 | have "x = (\<chi> k. det(\<chi> i j. if j=k then b$i else A$i$j) / det A)" | |
| 1006 | unfolding x[symmetric] | |
| 1007 | using d0 by (simp add: vec_eq_iff cramer_lemma field_simps) | |
| 1008 | } | |
| 53854 | 1009 | with xe show ?thesis | 
| 1010 | by auto | |
| 33175 | 1011 | qed | 
| 1012 | ||
| 53854 | 1013 | text {* Orthogonality of a transformation and matrix. *}
 | 
| 33175 | 1014 | |
| 1015 | definition "orthogonal_transformation f \<longleftrightarrow> linear f \<and> (\<forall>v w. f v \<bullet> f w = v \<bullet> w)" | |
| 1016 | ||
| 53253 | 1017 | lemma orthogonal_transformation: | 
| 1018 | "orthogonal_transformation f \<longleftrightarrow> linear f \<and> (\<forall>(v::real ^_). norm (f v) = norm v)" | |
| 33175 | 1019 | unfolding orthogonal_transformation_def | 
| 1020 | apply auto | |
| 1021 | apply (erule_tac x=v in allE)+ | |
| 35542 | 1022 | apply (simp add: norm_eq_sqrt_inner) | 
| 53253 | 1023 | apply (simp add: dot_norm linear_add[symmetric]) | 
| 1024 | done | |
| 33175 | 1025 | |
| 53253 | 1026 | definition "orthogonal_matrix (Q::'a::semiring_1^'n^'n) \<longleftrightarrow> | 
| 1027 | transpose Q ** Q = mat 1 \<and> Q ** transpose Q = mat 1" | |
| 33175 | 1028 | |
| 53253 | 1029 | lemma orthogonal_matrix: "orthogonal_matrix (Q:: real ^'n^'n) \<longleftrightarrow> transpose Q ** Q = mat 1" | 
| 33175 | 1030 | by (metis matrix_left_right_inverse orthogonal_matrix_def) | 
| 1031 | ||
| 34291 
4e896680897e
finite annotation on cartesian product is now implicit.
 hoelzl parents: 
34289diff
changeset | 1032 | lemma orthogonal_matrix_id: "orthogonal_matrix (mat 1 :: _^'n^'n)" | 
| 35150 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 hoelzl parents: 
35028diff
changeset | 1033 | by (simp add: orthogonal_matrix_def transpose_mat matrix_mul_lid) | 
| 33175 | 1034 | |
| 1035 | lemma orthogonal_matrix_mul: | |
| 34291 
4e896680897e
finite annotation on cartesian product is now implicit.
 hoelzl parents: 
34289diff
changeset | 1036 | fixes A :: "real ^'n^'n" | 
| 33175 | 1037 | assumes oA : "orthogonal_matrix A" | 
| 53253 | 1038 | and oB: "orthogonal_matrix B" | 
| 33175 | 1039 | shows "orthogonal_matrix(A ** B)" | 
| 1040 | using oA oB | |
| 35150 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 hoelzl parents: 
35028diff
changeset | 1041 | unfolding orthogonal_matrix matrix_transpose_mul | 
| 33175 | 1042 | apply (subst matrix_mul_assoc) | 
| 1043 | apply (subst matrix_mul_assoc[symmetric]) | |
| 53253 | 1044 | apply (simp add: matrix_mul_rid) | 
| 1045 | done | |
| 33175 | 1046 | |
| 1047 | lemma orthogonal_transformation_matrix: | |
| 34291 
4e896680897e
finite annotation on cartesian product is now implicit.
 hoelzl parents: 
34289diff
changeset | 1048 | fixes f:: "real^'n \<Rightarrow> real^'n" | 
| 33175 | 1049 | shows "orthogonal_transformation f \<longleftrightarrow> linear f \<and> orthogonal_matrix(matrix f)" | 
| 1050 | (is "?lhs \<longleftrightarrow> ?rhs") | |
| 53253 | 1051 | proof - | 
| 33175 | 1052 | let ?mf = "matrix f" | 
| 1053 | let ?ot = "orthogonal_transformation f" | |
| 1054 | let ?U = "UNIV :: 'n set" | |
| 1055 | have fU: "finite ?U" by simp | |
| 1056 | let ?m1 = "mat 1 :: real ^'n^'n" | |
| 53253 | 1057 |   {
 | 
| 1058 | assume ot: ?ot | |
| 33175 | 1059 | from ot have lf: "linear f" and fd: "\<forall>v w. f v \<bullet> f w = v \<bullet> w" | 
| 1060 | unfolding orthogonal_transformation_def orthogonal_matrix by blast+ | |
| 53253 | 1061 |     {
 | 
| 1062 | fix i j | |
| 35150 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 hoelzl parents: 
35028diff
changeset | 1063 | let ?A = "transpose ?mf ** ?mf" | 
| 33175 | 1064 | have th0: "\<And>b (x::'a::comm_ring_1). (if b then 1 else 0)*x = (if b then x else 0)" | 
| 1065 | "\<And>b (x::'a::comm_ring_1). x*(if b then 1 else 0) = (if b then x else 0)" | |
| 1066 | by simp_all | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
47108diff
changeset | 1067 | from fd[rule_format, of "axis i 1" "axis j 1", unfolded matrix_works[OF lf, symmetric] dot_matrix_vector_mul] | 
| 33175 | 1068 | have "?A$i$j = ?m1 $ i $ j" | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
47108diff
changeset | 1069 | by (simp add: inner_vec_def matrix_matrix_mult_def columnvector_def rowvector_def | 
| 57418 | 1070 | th0 setsum.delta[OF fU] mat_def axis_def) | 
| 53253 | 1071 | } | 
| 53854 | 1072 | then have "orthogonal_matrix ?mf" | 
| 1073 | unfolding orthogonal_matrix | |
| 53253 | 1074 | by vector | 
| 53854 | 1075 | with lf have ?rhs | 
| 1076 | by blast | |
| 53253 | 1077 | } | 
| 33175 | 1078 | moreover | 
| 53253 | 1079 |   {
 | 
| 1080 | assume lf: "linear f" and om: "orthogonal_matrix ?mf" | |
| 33175 | 1081 | from lf om have ?lhs | 
| 1082 | unfolding orthogonal_matrix_def norm_eq orthogonal_transformation | |
| 1083 | unfolding matrix_works[OF lf, symmetric] | |
| 1084 | apply (subst dot_matrix_vector_mul) | |
| 53253 | 1085 | apply (simp add: dot_matrix_product matrix_mul_lid) | 
| 1086 | done | |
| 1087 | } | |
| 53854 | 1088 | ultimately show ?thesis | 
| 1089 | by blast | |
| 33175 | 1090 | qed | 
| 1091 | ||
| 1092 | lemma det_orthogonal_matrix: | |
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
34291diff
changeset | 1093 | fixes Q:: "'a::linordered_idom^'n^'n" | 
| 33175 | 1094 | assumes oQ: "orthogonal_matrix Q" | 
| 1095 | shows "det Q = 1 \<or> det Q = - 1" | |
| 53253 | 1096 | proof - | 
| 33175 | 1097 | have th: "\<And>x::'a. x = 1 \<or> x = - 1 \<longleftrightarrow> x*x = 1" (is "\<And>x::'a. ?ths x") | 
| 53253 | 1098 | proof - | 
| 33175 | 1099 | fix x:: 'a | 
| 53854 | 1100 | have th0: "x * x - 1 = (x - 1) * (x + 1)" | 
| 53253 | 1101 | by (simp add: field_simps) | 
| 33175 | 1102 | have th1: "\<And>(x::'a) y. x = - y \<longleftrightarrow> x + y = 0" | 
| 53253 | 1103 | apply (subst eq_iff_diff_eq_0) | 
| 1104 | apply simp | |
| 1105 | done | |
| 53854 | 1106 | have "x * x = 1 \<longleftrightarrow> x * x - 1 = 0" | 
| 1107 | by simp | |
| 1108 | also have "\<dots> \<longleftrightarrow> x = 1 \<or> x = - 1" | |
| 1109 | unfolding th0 th1 by simp | |
| 33175 | 1110 | finally show "?ths x" .. | 
| 1111 | qed | |
| 53253 | 1112 | from oQ have "Q ** transpose Q = mat 1" | 
| 1113 | by (metis orthogonal_matrix_def) | |
| 1114 | then have "det (Q ** transpose Q) = det (mat 1:: 'a^'n^'n)" | |
| 1115 | by simp | |
| 1116 | then have "det Q * det Q = 1" | |
| 1117 | by (simp add: det_mul det_I det_transpose) | |
| 33175 | 1118 | then show ?thesis unfolding th . | 
| 1119 | qed | |
| 1120 | ||
| 53854 | 1121 | text {* Linearity of scaling, and hence isometry, that preserves origin. *}
 | 
| 1122 | ||
| 33175 | 1123 | lemma scaling_linear: | 
| 34291 
4e896680897e
finite annotation on cartesian product is now implicit.
 hoelzl parents: 
34289diff
changeset | 1124 | fixes f :: "real ^'n \<Rightarrow> real ^'n" | 
| 53253 | 1125 | assumes f0: "f 0 = 0" | 
| 1126 | and fd: "\<forall>x y. dist (f x) (f y) = c * dist x y" | |
| 33175 | 1127 | shows "linear f" | 
| 53253 | 1128 | proof - | 
| 1129 |   {
 | |
| 1130 | fix v w | |
| 1131 |     {
 | |
| 1132 | fix x | |
| 1133 | note fd[rule_format, of x 0, unfolded dist_norm f0 diff_0_right] | |
| 1134 | } | |
| 33175 | 1135 | note th0 = this | 
| 53077 | 1136 | have "f v \<bullet> f w = c\<^sup>2 * (v \<bullet> w)" | 
| 33175 | 1137 | unfolding dot_norm_neg dist_norm[symmetric] | 
| 1138 | unfolding th0 fd[rule_format] by (simp add: power2_eq_square field_simps)} | |
| 1139 | note fc = this | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
47108diff
changeset | 1140 | show ?thesis | 
| 53600 
8fda7ad57466
make 'linear' into a sublocale of 'bounded_linear';
 huffman parents: 
53253diff
changeset | 1141 | unfolding linear_iff vector_eq[where 'a="real^'n"] scalar_mult_eq_scaleR | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
47108diff
changeset | 1142 | by (simp add: inner_add fc field_simps) | 
| 33175 | 1143 | qed | 
| 1144 | ||
| 1145 | lemma isometry_linear: | |
| 53253 | 1146 | "f (0:: real^'n) = (0:: real^'n) \<Longrightarrow> \<forall>x y. dist(f x) (f y) = dist x y \<Longrightarrow> linear f" | 
| 1147 | by (rule scaling_linear[where c=1]) simp_all | |
| 33175 | 1148 | |
| 53854 | 1149 | text {* Hence another formulation of orthogonal transformation. *}
 | 
| 33175 | 1150 | |
| 1151 | lemma orthogonal_transformation_isometry: | |
| 34291 
4e896680897e
finite annotation on cartesian product is now implicit.
 hoelzl parents: 
34289diff
changeset | 1152 | "orthogonal_transformation f \<longleftrightarrow> f(0::real^'n) = (0::real^'n) \<and> (\<forall>x y. dist(f x) (f y) = dist x y)" | 
| 33175 | 1153 | unfolding orthogonal_transformation | 
| 1154 | apply (rule iffI) | |
| 1155 | apply clarify | |
| 1156 | apply (clarsimp simp add: linear_0 linear_sub[symmetric] dist_norm) | |
| 1157 | apply (rule conjI) | |
| 1158 | apply (rule isometry_linear) | |
| 1159 | apply simp | |
| 1160 | apply simp | |
| 1161 | apply clarify | |
| 1162 | apply (erule_tac x=v in allE) | |
| 1163 | apply (erule_tac x=0 in allE) | |
| 53253 | 1164 | apply (simp add: dist_norm) | 
| 1165 | done | |
| 33175 | 1166 | |
| 53854 | 1167 | text {* Can extend an isometry from unit sphere. *}
 | 
| 33175 | 1168 | |
| 1169 | lemma isometry_sphere_extend: | |
| 34291 
4e896680897e
finite annotation on cartesian product is now implicit.
 hoelzl parents: 
34289diff
changeset | 1170 | fixes f:: "real ^'n \<Rightarrow> real ^'n" | 
| 33175 | 1171 | assumes f1: "\<forall>x. norm x = 1 \<longrightarrow> norm (f x) = 1" | 
| 53253 | 1172 | and fd1: "\<forall> x y. norm x = 1 \<longrightarrow> norm y = 1 \<longrightarrow> dist (f x) (f y) = dist x y" | 
| 33175 | 1173 | shows "\<exists>g. orthogonal_transformation g \<and> (\<forall>x. norm x = 1 \<longrightarrow> g x = f x)" | 
| 53253 | 1174 | proof - | 
| 1175 |   {
 | |
| 1176 | fix x y x' y' x0 y0 x0' y0' :: "real ^'n" | |
| 1177 | assume H: | |
| 1178 | "x = norm x *\<^sub>R x0" | |
| 1179 | "y = norm y *\<^sub>R y0" | |
| 1180 | "x' = norm x *\<^sub>R x0'" "y' = norm y *\<^sub>R y0'" | |
| 1181 | "norm x0 = 1" "norm x0' = 1" "norm y0 = 1" "norm y0' = 1" | |
| 1182 | "norm(x0' - y0') = norm(x0 - y0)" | |
| 53854 | 1183 | then have *: "x0 \<bullet> y0 = x0' \<bullet> y0' + y0' \<bullet> x0' - y0 \<bullet> x0 " | 
| 53253 | 1184 | by (simp add: norm_eq norm_eq_1 inner_add inner_diff) | 
| 33175 | 1185 | have "norm(x' - y') = norm(x - y)" | 
| 1186 | apply (subst H(1)) | |
| 1187 | apply (subst H(2)) | |
| 1188 | apply (subst H(3)) | |
| 1189 | apply (subst H(4)) | |
| 1190 | using H(5-9) | |
| 1191 | apply (simp add: norm_eq norm_eq_1) | |
| 53854 | 1192 | apply (simp add: inner_diff scalar_mult_eq_scaleR) | 
| 1193 | unfolding * | |
| 53253 | 1194 | apply (simp add: field_simps) | 
| 1195 | done | |
| 1196 | } | |
| 33175 | 1197 | note th0 = this | 
| 44228 
5f974bead436
get Multivariate_Analysis/Determinants.thy compiled and working again
 huffman parents: 
41959diff
changeset | 1198 | let ?g = "\<lambda>x. if x = 0 then 0 else norm x *\<^sub>R f (inverse (norm x) *\<^sub>R x)" | 
| 53253 | 1199 |   {
 | 
| 1200 | fix x:: "real ^'n" | |
| 1201 | assume nx: "norm x = 1" | |
| 53854 | 1202 | have "?g x = f x" | 
| 1203 | using nx by auto | |
| 53253 | 1204 | } | 
| 1205 | then have thfg: "\<forall>x. norm x = 1 \<longrightarrow> ?g x = f x" | |
| 1206 | by blast | |
| 53854 | 1207 | have g0: "?g 0 = 0" | 
| 1208 | by simp | |
| 53253 | 1209 |   {
 | 
| 1210 | fix x y :: "real ^'n" | |
| 1211 |     {
 | |
| 1212 | assume "x = 0" "y = 0" | |
| 53854 | 1213 | then have "dist (?g x) (?g y) = dist x y" | 
| 1214 | by simp | |
| 53253 | 1215 | } | 
| 33175 | 1216 | moreover | 
| 53253 | 1217 |     {
 | 
| 1218 | assume "x = 0" "y \<noteq> 0" | |
| 33175 | 1219 | then have "dist (?g x) (?g y) = dist x y" | 
| 36362 
06475a1547cb
fix lots of looping simp calls and other warnings
 huffman parents: 
35542diff
changeset | 1220 | apply (simp add: dist_norm) | 
| 33175 | 1221 | apply (rule f1[rule_format]) | 
| 53253 | 1222 | apply (simp add: field_simps) | 
| 1223 | done | |
| 1224 | } | |
| 33175 | 1225 | moreover | 
| 53253 | 1226 |     {
 | 
| 1227 | assume "x \<noteq> 0" "y = 0" | |
| 33175 | 1228 | then have "dist (?g x) (?g y) = dist x y" | 
| 36362 
06475a1547cb
fix lots of looping simp calls and other warnings
 huffman parents: 
35542diff
changeset | 1229 | apply (simp add: dist_norm) | 
| 33175 | 1230 | apply (rule f1[rule_format]) | 
| 53253 | 1231 | apply (simp add: field_simps) | 
| 1232 | done | |
| 1233 | } | |
| 33175 | 1234 | moreover | 
| 53253 | 1235 |     {
 | 
| 1236 | assume z: "x \<noteq> 0" "y \<noteq> 0" | |
| 1237 | have th00: | |
| 1238 | "x = norm x *\<^sub>R (inverse (norm x) *\<^sub>R x)" | |
| 1239 | "y = norm y *\<^sub>R (inverse (norm y) *\<^sub>R y)" | |
| 1240 | "norm x *\<^sub>R f ((inverse (norm x) *\<^sub>R x)) = norm x *\<^sub>R f (inverse (norm x) *\<^sub>R x)" | |
| 44228 
5f974bead436
get Multivariate_Analysis/Determinants.thy compiled and working again
 huffman parents: 
41959diff
changeset | 1241 | "norm y *\<^sub>R f (inverse (norm y) *\<^sub>R y) = norm y *\<^sub>R f (inverse (norm y) *\<^sub>R y)" | 
| 
5f974bead436
get Multivariate_Analysis/Determinants.thy compiled and working again
 huffman parents: 
41959diff
changeset | 1242 | "norm (inverse (norm x) *\<^sub>R x) = 1" | 
| 
5f974bead436
get Multivariate_Analysis/Determinants.thy compiled and working again
 huffman parents: 
41959diff
changeset | 1243 | "norm (f (inverse (norm x) *\<^sub>R x)) = 1" | 
| 
5f974bead436
get Multivariate_Analysis/Determinants.thy compiled and working again
 huffman parents: 
41959diff
changeset | 1244 | "norm (inverse (norm y) *\<^sub>R y) = 1" | 
| 
5f974bead436
get Multivariate_Analysis/Determinants.thy compiled and working again
 huffman parents: 
41959diff
changeset | 1245 | "norm (f (inverse (norm y) *\<^sub>R y)) = 1" | 
| 
5f974bead436
get Multivariate_Analysis/Determinants.thy compiled and working again
 huffman parents: 
41959diff
changeset | 1246 | "norm (f (inverse (norm x) *\<^sub>R x) - f (inverse (norm y) *\<^sub>R y)) = | 
| 53253 | 1247 | norm (inverse (norm x) *\<^sub>R x - inverse (norm y) *\<^sub>R y)" | 
| 33175 | 1248 | using z | 
| 44457 
d366fa5551ef
declare euclidean_simps [simp] at the point they are proved;
 huffman parents: 
44260diff
changeset | 1249 | by (auto simp add: field_simps intro: f1[rule_format] fd1[rule_format, unfolded dist_norm]) | 
| 33175 | 1250 | from z th0[OF th00] have "dist (?g x) (?g y) = dist x y" | 
| 53253 | 1251 | by (simp add: dist_norm) | 
| 1252 | } | |
| 53854 | 1253 | ultimately have "dist (?g x) (?g y) = dist x y" | 
| 1254 | by blast | |
| 53253 | 1255 | } | 
| 33175 | 1256 | note thd = this | 
| 1257 | show ?thesis | |
| 1258 | apply (rule exI[where x= ?g]) | |
| 1259 | unfolding orthogonal_transformation_isometry | |
| 53253 | 1260 | using g0 thfg thd | 
| 1261 | apply metis | |
| 1262 | done | |
| 33175 | 1263 | qed | 
| 1264 | ||
| 53854 | 1265 | text {* Rotation, reflection, rotoinversion. *}
 | 
| 33175 | 1266 | |
| 1267 | definition "rotation_matrix Q \<longleftrightarrow> orthogonal_matrix Q \<and> det Q = 1" | |
| 1268 | definition "rotoinversion_matrix Q \<longleftrightarrow> orthogonal_matrix Q \<and> det Q = - 1" | |
| 1269 | ||
| 1270 | lemma orthogonal_rotation_or_rotoinversion: | |
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
34291diff
changeset | 1271 | fixes Q :: "'a::linordered_idom^'n^'n" | 
| 33175 | 1272 | shows " orthogonal_matrix Q \<longleftrightarrow> rotation_matrix Q \<or> rotoinversion_matrix Q" | 
| 1273 | by (metis rotoinversion_matrix_def rotation_matrix_def det_orthogonal_matrix) | |
| 53253 | 1274 | |
| 53854 | 1275 | text {* Explicit formulas for low dimensions. *}
 | 
| 33175 | 1276 | |
| 57418 | 1277 | lemma setprod_neutral_const: "setprod f {(1::nat)..1} = f 1"
 | 
| 1278 | by (fact setprod_singleton_nat_seg) | |
| 33175 | 1279 | |
| 1280 | lemma setprod_2: "setprod f {(1::nat)..2} = f 1 * f 2"
 | |
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57418diff
changeset | 1281 | by (simp add: eval_nat_numeral setprod_numseg mult.commute) | 
| 53253 | 1282 | |
| 33175 | 1283 | lemma setprod_3: "setprod f {(1::nat)..3} = f 1 * f 2 * f 3"
 | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57418diff
changeset | 1284 | by (simp add: eval_nat_numeral setprod_numseg mult.commute) | 
| 33175 | 1285 | |
| 1286 | lemma det_1: "det (A::'a::comm_ring_1^1^1) = A$1$1" | |
| 44457 
d366fa5551ef
declare euclidean_simps [simp] at the point they are proved;
 huffman parents: 
44260diff
changeset | 1287 | by (simp add: det_def sign_id) | 
| 33175 | 1288 | |
| 1289 | lemma det_2: "det (A::'a::comm_ring_1^2^2) = A$1$1 * A$2$2 - A$1$2 * A$2$1" | |
| 53253 | 1290 | proof - | 
| 33175 | 1291 |   have f12: "finite {2::2}" "1 \<notin> {2::2}" by auto
 | 
| 1292 | show ?thesis | |
| 53253 | 1293 | unfolding det_def UNIV_2 | 
| 1294 | unfolding setsum_over_permutations_insert[OF f12] | |
| 1295 | unfolding permutes_sing | |
| 1296 | by (simp add: sign_swap_id sign_id swap_id_eq) | |
| 33175 | 1297 | qed | 
| 1298 | ||
| 53253 | 1299 | lemma det_3: | 
| 1300 | "det (A::'a::comm_ring_1^3^3) = | |
| 1301 | A$1$1 * A$2$2 * A$3$3 + | |
| 1302 | A$1$2 * A$2$3 * A$3$1 + | |
| 1303 | A$1$3 * A$2$1 * A$3$2 - | |
| 1304 | A$1$1 * A$2$3 * A$3$2 - | |
| 1305 | A$1$2 * A$2$1 * A$3$3 - | |
| 1306 | A$1$3 * A$2$2 * A$3$1" | |
| 1307 | proof - | |
| 53854 | 1308 |   have f123: "finite {2::3, 3}" "1 \<notin> {2::3, 3}"
 | 
| 1309 | by auto | |
| 1310 |   have f23: "finite {3::3}" "2 \<notin> {3::3}"
 | |
| 1311 | by auto | |
| 33175 | 1312 | |
| 1313 | show ?thesis | |
| 53253 | 1314 | unfolding det_def UNIV_3 | 
| 1315 | unfolding setsum_over_permutations_insert[OF f123] | |
| 1316 | unfolding setsum_over_permutations_insert[OF f23] | |
| 1317 | unfolding permutes_sing | |
| 1318 | by (simp add: sign_swap_id permutation_swap_id sign_compose sign_id swap_id_eq) | |
| 33175 | 1319 | qed | 
| 1320 | ||
| 1321 | end |