| author | wenzelm | 
| Sun, 24 Dec 2023 11:58:33 +0100 | |
| changeset 79346 | f86c310327df | 
| parent 76837 | d908a7d3ed1b | 
| child 82802 | 547335b41005 | 
| permissions | -rw-r--r-- | 
| 63627 | 1 | (* Title: HOL/Analysis/Determinants.thy | 
| 68143 | 2 | Author: Amine Chaieb, University of Cambridge; proofs reworked by LCP | 
| 33175 | 3 | *) | 
| 4 | ||
| 71044 | 5 | section \<open>Traces and Determinants of Square Matrices\<close> | 
| 33175 | 6 | |
| 7 | theory Determinants | |
| 44228 
5f974bead436
get Multivariate_Analysis/Determinants.thy compiled and working again
 huffman parents: 
41959diff
changeset | 8 | imports | 
| 73477 | 9 | "HOL-Combinatorics.Permutations" | 
| 69680 | 10 | Cartesian_Space | 
| 33175 | 11 | begin | 
| 12 | ||
| 69683 | 13 | subsection \<open>Trace\<close> | 
| 33175 | 14 | |
| 70136 | 15 | definition\<^marker>\<open>tag important\<close> trace :: "'a::semiring_1^'n^'n \<Rightarrow> 'a" | 
| 64267 | 16 | where "trace A = sum (\<lambda>i. ((A$i)$i)) (UNIV::'n set)" | 
| 33175 | 17 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 18 | lemma trace_0: "trace (mat 0) = 0" | 
| 33175 | 19 | by (simp add: trace_def mat_def) | 
| 20 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 21 | lemma  trace_I: "trace (mat 1 :: 'a::semiring_1^'n^'n) = of_nat(CARD('n))"
 | 
| 33175 | 22 | by (simp add: trace_def mat_def) | 
| 23 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 24 | lemma trace_add: "trace ((A::'a::comm_semiring_1^'n^'n) + B) = trace A + trace B" | 
| 64267 | 25 | by (simp add: trace_def sum.distrib) | 
| 33175 | 26 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 27 | lemma trace_sub: "trace ((A::'a::comm_ring_1^'n^'n) - B) = trace A - trace B" | 
| 64267 | 28 | by (simp add: trace_def sum_subtractf) | 
| 33175 | 29 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 30 | lemma trace_mul_sym: "trace ((A::'a::comm_semiring_1^'n^'m) ** B) = trace (B**A)" | 
| 33175 | 31 | apply (simp add: trace_def matrix_matrix_mult_def) | 
| 66804 
3f9bb52082c4
avoid name clashes on interpretation of abstract locales
 haftmann parents: 
66453diff
changeset | 32 | apply (subst sum.swap) | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57418diff
changeset | 33 | apply (simp add: mult.commute) | 
| 53253 | 34 | done | 
| 33175 | 35 | |
| 70136 | 36 | subsubsection\<^marker>\<open>tag important\<close> \<open>Definition of determinant\<close> | 
| 33175 | 37 | |
| 70136 | 38 | definition\<^marker>\<open>tag important\<close> det:: "'a::comm_ring_1^'n^'n \<Rightarrow> 'a" where | 
| 53253 | 39 | "det A = | 
| 64272 | 40 | sum (\<lambda>p. of_int (sign p) * prod (\<lambda>i. A$i$p i) (UNIV :: 'n set)) | 
| 53253 | 41 |       {p. p permutes (UNIV :: 'n set)}"
 | 
| 33175 | 42 | |
| 68134 | 43 | text \<open>Basic determinant properties\<close> | 
| 33175 | 44 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 45 | lemma det_transpose [simp]: "det (transpose A) = det (A::'a::comm_ring_1 ^'n^'n)" | 
| 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 46 | proof - | 
| 33175 | 47 | let ?di = "\<lambda>A i j. A$i$j" | 
| 48 | let ?U = "(UNIV :: 'n set)" | |
| 49 | have fU: "finite ?U" by simp | |
| 53253 | 50 |   {
 | 
| 51 | fix p | |
| 52 |     assume p: "p \<in> {p. p permutes ?U}"
 | |
| 53854 | 53 | from p have pU: "p permutes ?U" | 
| 54 | by blast | |
| 33175 | 55 | have sth: "sign (inv p) = sign p" | 
| 44260 
7784fa3232ce
Determinants.thy: avoid using mem_def/Collect_def
 huffman parents: 
44228diff
changeset | 56 | by (metis sign_inverse fU p mem_Collect_eq permutation_permutes) | 
| 33175 | 57 | from permutes_inj[OF pU] | 
| 53854 | 58 | have pi: "inj_on p ?U" | 
| 59 | by (blast intro: subset_inj_on) | |
| 33175 | 60 | from permutes_image[OF pU] | 
| 64272 | 61 | have "prod (\<lambda>i. ?di (transpose A) i (inv p i)) ?U = | 
| 62 | prod (\<lambda>i. ?di (transpose A) i (inv p i)) (p ` ?U)" | |
| 53854 | 63 | by simp | 
| 64272 | 64 | also have "\<dots> = prod ((\<lambda>i. ?di (transpose A) i (inv p i)) \<circ> p) ?U" | 
| 65 | unfolding prod.reindex[OF pi] .. | |
| 66 | also have "\<dots> = prod (\<lambda>i. ?di A i (p i)) ?U" | |
| 53253 | 67 | proof - | 
| 68134 | 68 | have "((\<lambda>i. ?di (transpose A) i (inv p i)) \<circ> p) i = ?di A i (p i)" if "i \<in> ?U" for i | 
| 69 | using that permutes_inv_o[OF pU] permutes_in_image[OF pU] | |
| 70 | unfolding transpose_def by (simp add: fun_eq_iff) | |
| 71 | then show "prod ((\<lambda>i. ?di (transpose A) i (inv p i)) \<circ> p) ?U = prod (\<lambda>i. ?di A i (p i)) ?U" | |
| 64272 | 72 | by (auto intro: prod.cong) | 
| 33175 | 73 | qed | 
| 64272 | 74 | finally have "of_int (sign (inv p)) * (prod (\<lambda>i. ?di (transpose A) i (inv p i)) ?U) = | 
| 75 | of_int (sign p) * (prod (\<lambda>i. ?di A i (p i)) ?U)" | |
| 53854 | 76 | using sth by simp | 
| 53253 | 77 | } | 
| 78 | then show ?thesis | |
| 79 | unfolding det_def | |
| 68138 | 80 | by (subst sum_permutations_inverse) (blast intro: sum.cong) | 
| 33175 | 81 | qed | 
| 82 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 83 | lemma det_lowerdiagonal: | 
| 34291 
4e896680897e
finite annotation on cartesian product is now implicit.
 hoelzl parents: 
34289diff
changeset | 84 |   fixes A :: "'a::comm_ring_1^('n::{finite,wellorder})^('n::{finite,wellorder})"
 | 
| 33175 | 85 | assumes ld: "\<And>i j. i < j \<Longrightarrow> A$i$j = 0" | 
| 64272 | 86 | shows "det A = prod (\<lambda>i. A$i$i) (UNIV:: 'n set)" | 
| 53253 | 87 | proof - | 
| 33175 | 88 | let ?U = "UNIV:: 'n set" | 
| 89 |   let ?PU = "{p. p permutes ?U}"
 | |
| 64272 | 90 | let ?pp = "\<lambda>p. of_int (sign p) * prod (\<lambda>i. A$i$p i) (UNIV :: 'n set)" | 
| 53854 | 91 | have fU: "finite ?U" | 
| 92 | by simp | |
| 93 |   have id0: "{id} \<subseteq> ?PU"
 | |
| 68138 | 94 | by (auto simp: permutes_id) | 
| 68134 | 95 |   have p0: "\<forall>p \<in> ?PU - {id}. ?pp p = 0"
 | 
| 96 | proof | |
| 53253 | 97 | fix p | 
| 68134 | 98 |     assume "p \<in> ?PU - {id}"
 | 
| 99 | then obtain i where i: "p i > i" | |
| 100 | by clarify (meson leI permutes_natset_le) | |
| 101 | from ld[OF i] have "\<exists>i \<in> ?U. A$i$p i = 0" | |
| 53253 | 102 | by blast | 
| 68134 | 103 | with prod_zero[OF fU] show "?pp p = 0" | 
| 104 | by force | |
| 105 | qed | |
| 106 | from sum.mono_neutral_cong_left[OF finite_permutations[OF fU] id0 p0] show ?thesis | |
| 33175 | 107 | unfolding det_def by (simp add: sign_id) | 
| 108 | qed | |
| 109 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 110 | lemma det_upperdiagonal: | 
| 34291 
4e896680897e
finite annotation on cartesian product is now implicit.
 hoelzl parents: 
34289diff
changeset | 111 |   fixes A :: "'a::comm_ring_1^'n::{finite,wellorder}^'n::{finite,wellorder}"
 | 
| 33175 | 112 | assumes ld: "\<And>i j. i > j \<Longrightarrow> A$i$j = 0" | 
| 64272 | 113 | shows "det A = prod (\<lambda>i. A$i$i) (UNIV:: 'n set)" | 
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 114 | proof - | 
| 33175 | 115 | let ?U = "UNIV:: 'n set" | 
| 116 |   let ?PU = "{p. p permutes ?U}"
 | |
| 64272 | 117 | let ?pp = "(\<lambda>p. of_int (sign p) * prod (\<lambda>i. A$i$p i) (UNIV :: 'n set))" | 
| 53854 | 118 | have fU: "finite ?U" | 
| 119 | by simp | |
| 120 |   have id0: "{id} \<subseteq> ?PU"
 | |
| 68138 | 121 | by (auto simp: permutes_id) | 
| 68134 | 122 |   have p0: "\<forall>p \<in> ?PU -{id}. ?pp p = 0"
 | 
| 123 | proof | |
| 53253 | 124 | fix p | 
| 53854 | 125 |     assume p: "p \<in> ?PU - {id}"
 | 
| 68134 | 126 | then obtain i where i: "p i < i" | 
| 127 | by clarify (meson leI permutes_natset_ge) | |
| 128 | from ld[OF i] have "\<exists>i \<in> ?U. A$i$p i = 0" | |
| 53854 | 129 | by blast | 
| 68134 | 130 | with prod_zero[OF fU] show "?pp p = 0" | 
| 131 | by force | |
| 132 | qed | |
| 133 | from sum.mono_neutral_cong_left[OF finite_permutations[OF fU] id0 p0] show ?thesis | |
| 33175 | 134 | unfolding det_def by (simp add: sign_id) | 
| 135 | qed | |
| 136 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 137 | proposition det_diagonal: | 
| 34291 
4e896680897e
finite annotation on cartesian product is now implicit.
 hoelzl parents: 
34289diff
changeset | 138 | fixes A :: "'a::comm_ring_1^'n^'n" | 
| 33175 | 139 | assumes ld: "\<And>i j. i \<noteq> j \<Longrightarrow> A$i$j = 0" | 
| 64272 | 140 | shows "det A = prod (\<lambda>i. A$i$i) (UNIV::'n set)" | 
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 141 | proof - | 
| 33175 | 142 | let ?U = "UNIV:: 'n set" | 
| 143 |   let ?PU = "{p. p permutes ?U}"
 | |
| 64272 | 144 | let ?pp = "\<lambda>p. of_int (sign p) * prod (\<lambda>i. A$i$p i) (UNIV :: 'n set)" | 
| 33175 | 145 | have fU: "finite ?U" by simp | 
| 146 | from finite_permutations[OF fU] have fPU: "finite ?PU" . | |
| 53854 | 147 |   have id0: "{id} \<subseteq> ?PU"
 | 
| 68138 | 148 | by (auto simp: permutes_id) | 
| 68134 | 149 |   have p0: "\<forall>p \<in> ?PU - {id}. ?pp p = 0"
 | 
| 150 | proof | |
| 53253 | 151 | fix p | 
| 152 |     assume p: "p \<in> ?PU - {id}"
 | |
| 53854 | 153 | then obtain i where i: "p i \<noteq> i" | 
| 68134 | 154 | by fastforce | 
| 155 | with ld have "\<exists>i \<in> ?U. A$i$p i = 0" | |
| 156 | by (metis UNIV_I) | |
| 157 | with prod_zero [OF fU] show "?pp p = 0" | |
| 158 | by force | |
| 159 | qed | |
| 64267 | 160 | from sum.mono_neutral_cong_left[OF fPU id0 p0] show ?thesis | 
| 33175 | 161 | unfolding det_def by (simp add: sign_id) | 
| 162 | qed | |
| 163 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 164 | lemma det_I [simp]: "det (mat 1 :: 'a::comm_ring_1^'n^'n) = 1" | 
| 67673 
c8caefb20564
lots of new material, ultimately related to measure theory
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 165 | by (simp add: det_diagonal mat_def) | 
| 33175 | 166 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 167 | lemma det_0 [simp]: "det (mat 0 :: 'a::comm_ring_1^'n^'n) = 0" | 
| 67970 | 168 | by (simp add: det_def prod_zero power_0_left) | 
| 33175 | 169 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 170 | lemma det_permute_rows: | 
| 34291 
4e896680897e
finite annotation on cartesian product is now implicit.
 hoelzl parents: 
34289diff
changeset | 171 | fixes A :: "'a::comm_ring_1^'n^'n" | 
| 33175 | 172 | assumes p: "p permutes (UNIV :: 'n::finite set)" | 
| 53854 | 173 | shows "det (\<chi> i. A$p i :: 'a^'n^'n) = of_int (sign p) * det A" | 
| 68134 | 174 | proof - | 
| 33175 | 175 | let ?U = "UNIV :: 'n set" | 
| 176 |   let ?PU = "{p. p permutes ?U}"
 | |
| 68134 | 177 | have *: "(\<Sum>q\<in>?PU. of_int (sign (q \<circ> p)) * (\<Prod>i\<in>?U. A $ p i $ (q \<circ> p) i)) = | 
| 178 | (\<Sum>n\<in>?PU. of_int (sign p) * of_int (sign n) * (\<Prod>i\<in>?U. A $ i $ n i))" | |
| 179 | proof (rule sum.cong) | |
| 180 | fix q | |
| 181 | assume qPU: "q \<in> ?PU" | |
| 182 | have fU: "finite ?U" | |
| 183 | by simp | |
| 184 | from qPU have q: "q permutes ?U" | |
| 185 | by blast | |
| 186 | have "prod (\<lambda>i. A$p i$ (q \<circ> p) i) ?U = prod ((\<lambda>i. A$p i$(q \<circ> p) i) \<circ> inv p) ?U" | |
| 187 | by (simp only: prod.permute[OF permutes_inv[OF p], symmetric]) | |
| 188 | also have "\<dots> = prod (\<lambda>i. A $ (p \<circ> inv p) i $ (q \<circ> (p \<circ> inv p)) i) ?U" | |
| 189 | by (simp only: o_def) | |
| 190 | also have "\<dots> = prod (\<lambda>i. A$i$q i) ?U" | |
| 191 | by (simp only: o_def permutes_inverses[OF p]) | |
| 192 | finally have thp: "prod (\<lambda>i. A$p i$ (q \<circ> p) i) ?U = prod (\<lambda>i. A$i$q i) ?U" | |
| 193 | by blast | |
| 194 | from p q have pp: "permutation p" and qp: "permutation q" | |
| 195 | by (metis fU permutation_permutes)+ | |
| 196 | show "of_int (sign (q \<circ> p)) * prod (\<lambda>i. A$ p i$ (q \<circ> p) i) ?U = | |
| 197 | of_int (sign p) * of_int (sign q) * prod (\<lambda>i. A$i$q i) ?U" | |
| 198 | by (simp only: thp sign_compose[OF qp pp] mult.commute of_int_mult) | |
| 199 | qed auto | |
| 200 | show ?thesis | |
| 201 | apply (simp add: det_def sum_distrib_left mult.assoc[symmetric]) | |
| 202 | apply (subst sum_permutations_compose_right[OF p]) | |
| 203 | apply (rule *) | |
| 204 | done | |
| 68143 | 205 | qed | 
| 33175 | 206 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 207 | lemma det_permute_columns: | 
| 34291 
4e896680897e
finite annotation on cartesian product is now implicit.
 hoelzl parents: 
34289diff
changeset | 208 | fixes A :: "'a::comm_ring_1^'n^'n" | 
| 33175 | 209 | assumes p: "p permutes (UNIV :: 'n set)" | 
| 210 | shows "det(\<chi> i j. A$i$ p j :: 'a^'n^'n) = of_int (sign p) * det A" | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 211 | proof - | 
| 33175 | 212 | 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 | 213 | let ?At = "transpose A" | 
| 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 hoelzl parents: 
35028diff
changeset | 214 | 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 | 215 | unfolding det_permute_rows[OF p, of ?At] det_transpose .. | 
| 33175 | 216 | moreover | 
| 35150 
082fa4bd403d
Rename transp to transpose in HOL-Multivariate_Analysis. (by himmelma)
 hoelzl parents: 
35028diff
changeset | 217 | have "?Ap = transpose (\<chi> i. transpose A $ p i)" | 
| 44228 
5f974bead436
get Multivariate_Analysis/Determinants.thy compiled and working again
 huffman parents: 
41959diff
changeset | 218 | by (simp add: transpose_def vec_eq_iff) | 
| 53854 | 219 | ultimately show ?thesis | 
| 220 | by simp | |
| 33175 | 221 | qed | 
| 222 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 223 | lemma det_identical_columns: | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 224 | fixes A :: "'a::comm_ring_1^'n^'n" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 225 | assumes jk: "j \<noteq> k" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 226 | and r: "column j A = column k A" | 
| 33175 | 227 | shows "det A = 0" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 228 | proof - | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 229 | let ?U="UNIV::'n set" | 
| 73648 | 230 | let ?t_jk="Transposition.transpose j k" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 231 |   let ?PU="{p. p permutes ?U}"
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 232 |   let ?S1="{p. p\<in>?PU \<and> evenperm p}"
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 233 |   let ?S2="{(?t_jk \<circ> p) |p. p \<in>?S1}"
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 234 | let ?f="\<lambda>p. of_int (sign p) * (\<Prod>i\<in>UNIV. A $ i $ p i)" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 235 | let ?g="\<lambda>p. ?t_jk \<circ> p" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 236 | have g_S1: "?S2 = ?g` ?S1" by auto | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 237 | have inj_g: "inj_on ?g ?S1" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 238 | proof (unfold inj_on_def, auto) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 239 | fix x y assume x: "x permutes ?U" and even_x: "evenperm x" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 240 | and y: "y permutes ?U" and even_y: "evenperm y" and eq: "?t_jk \<circ> x = ?t_jk \<circ> y" | 
| 73932 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 desharna parents: 
73648diff
changeset | 241 | show "x = y" by (metis (opaque_lifting, no_types) comp_assoc eq id_comp swap_id_idempotent) | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 242 | qed | 
| 73648 | 243 | have tjk_permutes: "?t_jk permutes ?U" | 
| 244 | by (auto simp add: permutes_def dest: transpose_eq_imp_eq) (meson transpose_involutory) | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 245 | have tjk_eq: "\<forall>i l. A $ i $ ?t_jk l = A $ i $ l" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 246 | using r jk | 
| 73648 | 247 | unfolding column_def vec_eq_iff by (simp add: Transposition.transpose_def) | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 248 | have sign_tjk: "sign ?t_jk = -1" using sign_swap_id[of j k] jk by auto | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 249 |   {fix x
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 250 | assume x: "x\<in> ?S1" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 251 | have "sign (?t_jk \<circ> x) = sign (?t_jk) * sign x" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 252 | by (metis (lifting) finite_class.finite_UNIV mem_Collect_eq | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 253 | permutation_permutes permutation_swap_id sign_compose x) | 
| 68138 | 254 | also have "\<dots> = - sign x" using sign_tjk by simp | 
| 255 | also have "\<dots> \<noteq> sign x" unfolding sign_def by simp | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 256 | finally have "sign (?t_jk \<circ> x) \<noteq> sign x" and "(?t_jk \<circ> x) \<in> ?S2" | 
| 68134 | 257 | using x by force+ | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 258 | } | 
| 68134 | 259 |   hence disjoint: "?S1 \<inter> ?S2 = {}"
 | 
| 260 | by (force simp: sign_def) | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 261 | have PU_decomposition: "?PU = ?S1 \<union> ?S2" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 262 | proof (auto) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 263 | fix x | 
| 73648 | 264 | assume x: "x permutes ?U" and "\<forall>p. p permutes ?U \<longrightarrow> x = Transposition.transpose j k \<circ> p \<longrightarrow> \<not> evenperm p" | 
| 265 | then obtain p where p: "p permutes UNIV" and x_eq: "x = Transposition.transpose j k \<circ> p" | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 266 | and odd_p: "\<not> evenperm p" | 
| 68134 | 267 | by (metis (mono_tags) id_o o_assoc permutes_compose swap_id_idempotent tjk_permutes) | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 268 | thus "evenperm x" | 
| 68134 | 269 | by (meson evenperm_comp evenperm_swap finite_class.finite_UNIV | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 270 | jk permutation_permutes permutation_swap_id) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 271 | next | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 272 | fix p assume p: "p permutes ?U" | 
| 73648 | 273 | show "Transposition.transpose j k \<circ> p permutes UNIV" by (metis p permutes_compose tjk_permutes) | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 274 | qed | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 275 | have "sum ?f ?S2 = sum ((\<lambda>p. of_int (sign p) * (\<Prod>i\<in>UNIV. A $ i $ p i)) | 
| 73648 | 276 |   \<circ> (\<circ>) (Transposition.transpose j k)) {p \<in> {p. p permutes UNIV}. evenperm p}"
 | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 277 | unfolding g_S1 by (rule sum.reindex[OF inj_g]) | 
| 68138 | 278 | also have "\<dots> = sum (\<lambda>p. of_int (sign (?t_jk \<circ> p)) * (\<Prod>i\<in>UNIV. A $ i $ p i)) ?S1" | 
| 279 | unfolding o_def by (rule sum.cong, auto simp: tjk_eq) | |
| 280 | also have "\<dots> = sum (\<lambda>p. - ?f p) ?S1" | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 281 | proof (rule sum.cong, auto) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 282 | fix x assume x: "x permutes ?U" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 283 | and even_x: "evenperm x" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 284 | hence perm_x: "permutation x" and perm_tjk: "permutation ?t_jk" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 285 | using permutation_permutes[of x] permutation_permutes[of ?t_jk] permutation_swap_id | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 286 | by (metis finite_code)+ | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 287 | have "(sign (?t_jk \<circ> x)) = - (sign x)" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 288 | unfolding sign_compose[OF perm_tjk perm_x] sign_tjk by auto | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 289 | thus "of_int (sign (?t_jk \<circ> x)) * (\<Prod>i\<in>UNIV. A $ i $ x i) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 290 | = - (of_int (sign x) * (\<Prod>i\<in>UNIV. A $ i $ x i))" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 291 | by auto | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 292 | qed | 
| 68138 | 293 | also have "\<dots>= - sum ?f ?S1" unfolding sum_negf .. | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 294 | finally have *: "sum ?f ?S2 = - sum ?f ?S1" . | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 295 | have "det A = (\<Sum>p | p permutes UNIV. of_int (sign p) * (\<Prod>i\<in>UNIV. A $ i $ p i))" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 296 | unfolding det_def .. | 
| 68138 | 297 | also have "\<dots>= sum ?f ?S1 + sum ?f ?S2" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 298 | by (subst PU_decomposition, rule sum.union_disjoint[OF _ _ disjoint], auto) | 
| 68138 | 299 | also have "\<dots>= sum ?f ?S1 - sum ?f ?S1 " unfolding * by auto | 
| 300 | also have "\<dots>= 0" by simp | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 301 | finally show "det A = 0" by simp | 
| 33175 | 302 | qed | 
| 303 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 304 | lemma det_identical_rows: | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 305 | fixes A :: "'a::comm_ring_1^'n^'n" | 
| 68134 | 306 | assumes ij: "i \<noteq> j" and r: "row i A = row j A" | 
| 33175 | 307 | shows "det A = 0" | 
| 68134 | 308 | by (metis column_transpose det_identical_columns det_transpose ij r) | 
| 33175 | 309 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 310 | lemma det_zero_row: | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 311 |   fixes A :: "'a::{idom, ring_char_0}^'n^'n" and F :: "'b::{field}^'m^'m"
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 312 | shows "row i A = 0 \<Longrightarrow> det A = 0" and "row j F = 0 \<Longrightarrow> det F = 0" | 
| 68138 | 313 | by (force simp: row_def det_def vec_eq_iff sign_nz intro!: sum.neutral)+ | 
| 33175 | 314 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 315 | lemma det_zero_column: | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 316 |   fixes A :: "'a::{idom, ring_char_0}^'n^'n" and F :: "'b::{field}^'m^'m"
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 317 | shows "column i A = 0 \<Longrightarrow> det A = 0" and "column j F = 0 \<Longrightarrow> det F = 0" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 318 | unfolding atomize_conj atomize_imp | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 319 | by (metis det_transpose det_zero_row row_transpose) | 
| 33175 | 320 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 321 | lemma det_row_add: | 
| 33175 | 322 | fixes a b c :: "'n::finite \<Rightarrow> _ ^ 'n" | 
| 323 | shows "det((\<chi> i. if i = k then a i + b i else c i)::'a::comm_ring_1^'n^'n) = | |
| 53253 | 324 | det((\<chi> i. if i = k then a i else c i)::'a::comm_ring_1^'n^'n) + | 
| 325 | det((\<chi> i. if i = k then b i else c i)::'a::comm_ring_1^'n^'n)" | |
| 64267 | 326 | unfolding det_def vec_lambda_beta sum.distrib[symmetric] | 
| 327 | proof (rule sum.cong) | |
| 33175 | 328 | let ?U = "UNIV :: 'n set" | 
| 329 |   let ?pU = "{p. p permutes ?U}"
 | |
| 330 | let ?f = "(\<lambda>i. if i = k then a i + b i else c i)::'n \<Rightarrow> 'a::comm_ring_1^'n" | |
| 331 | let ?g = "(\<lambda> i. if i = k then a i else c i)::'n \<Rightarrow> 'a::comm_ring_1^'n" | |
| 332 | let ?h = "(\<lambda> i. if i = k then b i else c i)::'n \<Rightarrow> 'a::comm_ring_1^'n" | |
| 53253 | 333 | fix p | 
| 334 | assume p: "p \<in> ?pU" | |
| 33175 | 335 |   let ?Uk = "?U - {k}"
 | 
| 53854 | 336 | from p have pU: "p permutes ?U" | 
| 337 | by blast | |
| 338 | have kU: "?U = insert k ?Uk" | |
| 339 | by blast | |
| 68134 | 340 | have eq: "prod (\<lambda>i. ?f i $ p i) ?Uk = prod (\<lambda>i. ?g i $ p i) ?Uk" | 
| 341 | "prod (\<lambda>i. ?f i $ p i) ?Uk = prod (\<lambda>i. ?h i $ p i) ?Uk" | |
| 342 | by auto | |
| 343 | have Uk: "finite ?Uk" "k \<notin> ?Uk" | |
| 53854 | 344 | by auto | 
| 64272 | 345 | have "prod (\<lambda>i. ?f i $ p i) ?U = prod (\<lambda>i. ?f i $ p i) (insert k ?Uk)" | 
| 33175 | 346 | unfolding kU[symmetric] .. | 
| 64272 | 347 | also have "\<dots> = ?f k $ p k * prod (\<lambda>i. ?f i $ p i) ?Uk" | 
| 68134 | 348 | by (rule prod.insert) auto | 
| 64272 | 349 | also have "\<dots> = (a k $ p k * prod (\<lambda>i. ?f i $ p i) ?Uk) + (b k$ p k * prod (\<lambda>i. ?f i $ p i) ?Uk)" | 
| 53253 | 350 | by (simp add: field_simps) | 
| 64272 | 351 | also have "\<dots> = (a k $ p k * prod (\<lambda>i. ?g i $ p i) ?Uk) + (b k$ p k * prod (\<lambda>i. ?h i $ p i) ?Uk)" | 
| 68134 | 352 | by (metis eq) | 
| 64272 | 353 | also have "\<dots> = prod (\<lambda>i. ?g i $ p i) (insert k ?Uk) + prod (\<lambda>i. ?h i $ p i) (insert k ?Uk)" | 
| 68134 | 354 | unfolding prod.insert[OF Uk] by simp | 
| 64272 | 355 | finally have "prod (\<lambda>i. ?f i $ p i) ?U = prod (\<lambda>i. ?g i $ p i) ?U + prod (\<lambda>i. ?h i $ p i) ?U" | 
| 53854 | 356 | unfolding kU[symmetric] . | 
| 64272 | 357 | then show "of_int (sign p) * prod (\<lambda>i. ?f i $ p i) ?U = | 
| 358 | of_int (sign p) * prod (\<lambda>i. ?g i $ p i) ?U + of_int (sign p) * prod (\<lambda>i. ?h i $ p i) ?U" | |
| 36350 | 359 | by (simp add: field_simps) | 
| 68134 | 360 | qed auto | 
| 33175 | 361 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 362 | lemma det_row_mul: | 
| 33175 | 363 | fixes a b :: "'n::finite \<Rightarrow> _ ^ 'n" | 
| 364 | shows "det((\<chi> i. if i = k then c *s a i else b i)::'a::comm_ring_1^'n^'n) = | |
| 53253 | 365 | c * det((\<chi> i. if i = k then a i else b i)::'a::comm_ring_1^'n^'n)" | 
| 64267 | 366 | unfolding det_def vec_lambda_beta sum_distrib_left | 
| 367 | proof (rule sum.cong) | |
| 33175 | 368 | let ?U = "UNIV :: 'n set" | 
| 369 |   let ?pU = "{p. p permutes ?U}"
 | |
| 370 | let ?f = "(\<lambda>i. if i = k then c*s a i else b i)::'n \<Rightarrow> 'a::comm_ring_1^'n" | |
| 371 | let ?g = "(\<lambda> i. if i = k then a i else b i)::'n \<Rightarrow> 'a::comm_ring_1^'n" | |
| 53253 | 372 | fix p | 
| 373 | assume p: "p \<in> ?pU" | |
| 33175 | 374 |   let ?Uk = "?U - {k}"
 | 
| 53854 | 375 | from p have pU: "p permutes ?U" | 
| 376 | by blast | |
| 377 | have kU: "?U = insert k ?Uk" | |
| 378 | by blast | |
| 68134 | 379 | have eq: "prod (\<lambda>i. ?f i $ p i) ?Uk = prod (\<lambda>i. ?g i $ p i) ?Uk" | 
| 68138 | 380 | by auto | 
| 68134 | 381 | have Uk: "finite ?Uk" "k \<notin> ?Uk" | 
| 53854 | 382 | by auto | 
| 64272 | 383 | have "prod (\<lambda>i. ?f i $ p i) ?U = prod (\<lambda>i. ?f i $ p i) (insert k ?Uk)" | 
| 33175 | 384 | unfolding kU[symmetric] .. | 
| 64272 | 385 | also have "\<dots> = ?f k $ p k * prod (\<lambda>i. ?f i $ p i) ?Uk" | 
| 68134 | 386 | by (rule prod.insert) auto | 
| 64272 | 387 | also have "\<dots> = (c*s a k) $ p k * prod (\<lambda>i. ?f i $ p i) ?Uk" | 
| 53253 | 388 | by (simp add: field_simps) | 
| 64272 | 389 | also have "\<dots> = c* (a k $ p k * prod (\<lambda>i. ?g i $ p i) ?Uk)" | 
| 68134 | 390 | unfolding eq by (simp add: ac_simps) | 
| 64272 | 391 | also have "\<dots> = c* (prod (\<lambda>i. ?g i $ p i) (insert k ?Uk))" | 
| 68134 | 392 | unfolding prod.insert[OF Uk] by simp | 
| 64272 | 393 | finally have "prod (\<lambda>i. ?f i $ p i) ?U = c* (prod (\<lambda>i. ?g i $ p i) ?U)" | 
| 53253 | 394 | unfolding kU[symmetric] . | 
| 68134 | 395 | then show "of_int (sign p) * prod (\<lambda>i. ?f i $ p i) ?U = c * (of_int (sign p) * prod (\<lambda>i. ?g i $ p i) ?U)" | 
| 36350 | 396 | by (simp add: field_simps) | 
| 68134 | 397 | qed auto | 
| 33175 | 398 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 399 | lemma det_row_0: | 
| 33175 | 400 | fixes b :: "'n::finite \<Rightarrow> _ ^ 'n" | 
| 401 | shows "det((\<chi> i. if i = k then 0 else b i)::'a::comm_ring_1^'n^'n) = 0" | |
| 53253 | 402 | using det_row_mul[of k 0 "\<lambda>i. 1" b] | 
| 403 | apply simp | |
| 404 | apply (simp only: vector_smult_lzero) | |
| 405 | done | |
| 33175 | 406 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 407 | lemma det_row_operation: | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 408 |   fixes A :: "'a::{comm_ring_1}^'n^'n"
 | 
| 33175 | 409 | assumes ij: "i \<noteq> j" | 
| 410 | shows "det (\<chi> k. if k = i then row i A + c *s row j A else row k A) = det A" | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 411 | proof - | 
| 33175 | 412 | let ?Z = "(\<chi> k. if k = i then row j A else row k A) :: 'a ^'n^'n" | 
| 413 | have th: "row i ?Z = row j ?Z" by (vector row_def) | |
| 414 | have th2: "((\<chi> k. if k = i then row i A else row k A) :: 'a^'n^'n) = A" | |
| 415 | by (vector row_def) | |
| 416 | show ?thesis | |
| 417 | unfolding det_row_add [of i] det_row_mul[of i] det_identical_rows[OF ij th] th2 | |
| 418 | by simp | |
| 419 | qed | |
| 420 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 421 | lemma det_row_span: | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 422 |   fixes A :: "'a::{field}^'n^'n"
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 423 |   assumes x: "x \<in> vec.span {row j A |j. j \<noteq> i}"
 | 
| 33175 | 424 | shows "det (\<chi> k. if k = i then row i A + x else row k A) = det A" | 
| 68069 
36209dfb981e
tidying up and using real induction methods
 paulson <lp15@cam.ac.uk> parents: 
68050diff
changeset | 425 | using x | 
| 68074 | 426 | proof (induction rule: vec.span_induct_alt) | 
| 68069 
36209dfb981e
tidying up and using real induction methods
 paulson <lp15@cam.ac.uk> parents: 
68050diff
changeset | 427 | case base | 
| 68134 | 428 | have "(if k = i then row i A + 0 else row k A) = row k A" for k | 
| 429 | by simp | |
| 68069 
36209dfb981e
tidying up and using real induction methods
 paulson <lp15@cam.ac.uk> parents: 
68050diff
changeset | 430 | then show ?case | 
| 68134 | 431 | by (simp add: row_def) | 
| 68069 
36209dfb981e
tidying up and using real induction methods
 paulson <lp15@cam.ac.uk> parents: 
68050diff
changeset | 432 | next | 
| 
36209dfb981e
tidying up and using real induction methods
 paulson <lp15@cam.ac.uk> parents: 
68050diff
changeset | 433 | case (step c z y) | 
| 
36209dfb981e
tidying up and using real induction methods
 paulson <lp15@cam.ac.uk> parents: 
68050diff
changeset | 434 | then obtain j where j: "z = row j A" "i \<noteq> j" | 
| 
36209dfb981e
tidying up and using real induction methods
 paulson <lp15@cam.ac.uk> parents: 
68050diff
changeset | 435 | by blast | 
| 
36209dfb981e
tidying up and using real induction methods
 paulson <lp15@cam.ac.uk> parents: 
68050diff
changeset | 436 | let ?w = "row i A + y" | 
| 
36209dfb981e
tidying up and using real induction methods
 paulson <lp15@cam.ac.uk> parents: 
68050diff
changeset | 437 | have th0: "row i A + (c*s z + y) = ?w + c*s z" | 
| 
36209dfb981e
tidying up and using real induction methods
 paulson <lp15@cam.ac.uk> parents: 
68050diff
changeset | 438 | by vector | 
| 
36209dfb981e
tidying up and using real induction methods
 paulson <lp15@cam.ac.uk> parents: 
68050diff
changeset | 439 | let ?d = "\<lambda>x. det (\<chi> k. if k = i then x else row k A)" | 
| 
36209dfb981e
tidying up and using real induction methods
 paulson <lp15@cam.ac.uk> parents: 
68050diff
changeset | 440 | have thz: "?d z = 0" | 
| 
36209dfb981e
tidying up and using real induction methods
 paulson <lp15@cam.ac.uk> parents: 
68050diff
changeset | 441 | apply (rule det_identical_rows[OF j(2)]) | 
| 
36209dfb981e
tidying up and using real induction methods
 paulson <lp15@cam.ac.uk> parents: 
68050diff
changeset | 442 | using j | 
| 
36209dfb981e
tidying up and using real induction methods
 paulson <lp15@cam.ac.uk> parents: 
68050diff
changeset | 443 | apply (vector row_def) | 
| 33175 | 444 | done | 
| 68069 
36209dfb981e
tidying up and using real induction methods
 paulson <lp15@cam.ac.uk> parents: 
68050diff
changeset | 445 | have "?d (row i A + (c*s z + y)) = ?d (?w + c*s z)" | 
| 
36209dfb981e
tidying up and using real induction methods
 paulson <lp15@cam.ac.uk> parents: 
68050diff
changeset | 446 | unfolding th0 .. | 
| 
36209dfb981e
tidying up and using real induction methods
 paulson <lp15@cam.ac.uk> parents: 
68050diff
changeset | 447 | then have "?d (row i A + (c*s z + y)) = det A" | 
| 
36209dfb981e
tidying up and using real induction methods
 paulson <lp15@cam.ac.uk> parents: 
68050diff
changeset | 448 | unfolding thz step.IH det_row_mul[of i] det_row_add[of i] by simp | 
| 
36209dfb981e
tidying up and using real induction methods
 paulson <lp15@cam.ac.uk> parents: 
68050diff
changeset | 449 | then show ?case | 
| 
36209dfb981e
tidying up and using real induction methods
 paulson <lp15@cam.ac.uk> parents: 
68050diff
changeset | 450 | unfolding scalar_mult_eq_scaleR . | 
| 68143 | 451 | qed | 
| 33175 | 452 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 453 | lemma matrix_id [simp]: "det (matrix id) = 1" | 
| 67673 
c8caefb20564
lots of new material, ultimately related to measure theory
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 454 | by (simp add: matrix_id_mat_1) | 
| 
c8caefb20564
lots of new material, ultimately related to measure theory
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 455 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 456 | proposition  det_matrix_scaleR [simp]: "det (matrix (((*\<^sub>R) r)) :: real^'n^'n) = r ^ CARD('n::finite)"
 | 
| 67673 
c8caefb20564
lots of new material, ultimately related to measure theory
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 457 | apply (subst det_diagonal) | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 458 | apply (auto simp: matrix_def mat_def) | 
| 67673 
c8caefb20564
lots of new material, ultimately related to measure theory
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 459 | apply (simp add: cart_eq_inner_axis inner_axis_axis) | 
| 
c8caefb20564
lots of new material, ultimately related to measure theory
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 460 | done | 
| 
c8caefb20564
lots of new material, ultimately related to measure theory
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 461 | |
| 60420 | 462 | text \<open> | 
| 53854 | 463 | May as well do this, though it's a bit unsatisfactory since it ignores | 
| 464 | exact duplicates by considering the rows/columns as a set. | |
| 60420 | 465 | \<close> | 
| 33175 | 466 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 467 | lemma det_dependent_rows: | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 468 |   fixes A:: "'a::{field}^'n^'n"
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 469 | assumes d: "vec.dependent (rows A)" | 
| 33175 | 470 | shows "det A = 0" | 
| 53253 | 471 | proof - | 
| 33175 | 472 | let ?U = "UNIV :: 'n set" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 473 |   from d obtain i where i: "row i A \<in> vec.span (rows A - {row i A})"
 | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 474 | unfolding vec.dependent_def rows_def by blast | 
| 68134 | 475 | show ?thesis | 
| 476 | proof (cases "\<forall>i j. i \<noteq> j \<longrightarrow> row i A \<noteq> row j A") | |
| 477 | case True | |
| 478 |     with i have "vec.span (rows A - {row i A}) \<subseteq> vec.span {row j A |j. j \<noteq> i}"
 | |
| 68138 | 479 | by (auto simp: rows_def intro!: vec.span_mono) | 
| 68134 | 480 |     then have "- row i A \<in> vec.span {row j A|j. j \<noteq> i}"
 | 
| 481 | by (meson i subsetCE vec.span_neg) | |
| 482 | from det_row_span[OF this] | |
| 33175 | 483 | have "det A = det (\<chi> k. if k = i then 0 *s 1 else row k A)" | 
| 484 | unfolding right_minus vector_smult_lzero .. | |
| 68134 | 485 | with det_row_mul[of i 0 "\<lambda>i. 1"] | 
| 486 | show ?thesis by simp | |
| 487 | next | |
| 488 | case False | |
| 489 | then obtain j k where jk: "j \<noteq> k" "row j A = row k A" | |
| 490 | by auto | |
| 491 | from det_identical_rows[OF jk] show ?thesis . | |
| 492 | qed | |
| 33175 | 493 | qed | 
| 494 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 495 | lemma det_dependent_columns: | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 496 | assumes d: "vec.dependent (columns (A::real^'n^'n))" | 
| 53253 | 497 | shows "det A = 0" | 
| 498 | by (metis d det_dependent_rows rows_transpose det_transpose) | |
| 33175 | 499 | |
| 68134 | 500 | text \<open>Multilinearity and the multiplication formula\<close> | 
| 33175 | 501 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 502 | lemma Cart_lambda_cong: "(\<And>x. f x = g x) \<Longrightarrow> (vec_lambda f::'a^'n) = (vec_lambda g :: 'a^'n)" | 
| 68134 | 503 | by auto | 
| 33175 | 504 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 505 | lemma det_linear_row_sum: | 
| 33175 | 506 | assumes fS: "finite S" | 
| 64267 | 507 | shows "det ((\<chi> i. if i = k then sum (a i) S else c i)::'a::comm_ring_1^'n^'n) = | 
| 508 | sum (\<lambda>j. det ((\<chi> i. if i = k then a i j else c i)::'a^'n^'n)) S" | |
| 68134 | 509 | using fS by (induct rule: finite_induct; simp add: det_row_0 det_row_add cong: if_cong) | 
| 33175 | 510 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 511 | lemma finite_bounded_functions: | 
| 33175 | 512 | assumes fS: "finite S" | 
| 513 |   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 | 514 | proof (induct k) | 
| 33175 | 515 | case 0 | 
| 68134 | 516 |   have *: "{f. \<forall>i. f i = i} = {id}"
 | 
| 53854 | 517 | by auto | 
| 518 | show ?case | |
| 68138 | 519 | by (auto simp: *) | 
| 33175 | 520 | next | 
| 521 | case (Suc k) | |
| 522 | let ?f = "\<lambda>(y::nat,g) i. if i = Suc k then y else g i" | |
| 523 |   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)})"
 | |
| 524 |   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)}"
 | |
| 68138 | 525 | apply (auto simp: image_iff) | 
| 68134 | 526 | apply (rename_tac f) | 
| 527 | apply (rule_tac x="f (Suc k)" in bexI) | |
| 68138 | 528 | apply (rule_tac x = "\<lambda>i. if i = Suc k then i else f i" in exI, auto) | 
| 33175 | 529 | done | 
| 530 | with finite_imageI[OF finite_cartesian_product[OF fS Suc.hyps(1)], of ?f] | |
| 53854 | 531 | show ?case | 
| 532 | by metis | |
| 33175 | 533 | qed | 
| 534 | ||
| 535 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 536 | lemma det_linear_rows_sum_lemma: | 
| 53854 | 537 | assumes fS: "finite S" | 
| 538 | and fT: "finite T" | |
| 64267 | 539 | shows "det ((\<chi> i. if i \<in> T then sum (a i) S else c i):: 'a::comm_ring_1^'n^'n) = | 
| 540 | sum (\<lambda>f. det((\<chi> i. if i \<in> T then a i (f i) else c i)::'a^'n^'n)) | |
| 53253 | 541 |       {f. (\<forall>i \<in> T. f i \<in> S) \<and> (\<forall>i. i \<notin> T \<longrightarrow> f i = i)}"
 | 
| 542 | using fT | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 543 | proof (induct T arbitrary: a c set: finite) | 
| 33175 | 544 | case empty | 
| 53253 | 545 |   have th0: "\<And>x y. (\<chi> i. if i \<in> {} then x i else y i) = (\<chi> i. y i)"
 | 
| 546 | by vector | |
| 53854 | 547 | from empty.prems show ?case | 
| 62408 
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
 paulson <lp15@cam.ac.uk> parents: 
61286diff
changeset | 548 | unfolding th0 by (simp add: eq_id_iff) | 
| 33175 | 549 | next | 
| 550 | case (insert z T a c) | |
| 551 |   let ?F = "\<lambda>T. {f. (\<forall>i \<in> T. f i \<in> S) \<and> (\<forall>i. i \<notin> T \<longrightarrow> f i = i)}"
 | |
| 552 | let ?h = "\<lambda>(y,g) i. if i = z then y else g i" | |
| 553 | let ?k = "\<lambda>h. (h(z),(\<lambda>i. if i = z then i else h i))" | |
| 554 | 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 | 555 | let ?c = "\<lambda>j i. if i = z then a i j else c i" | 
| 53253 | 556 | 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)" | 
| 557 | by simp | |
| 33175 | 558 | have thif2: "\<And>a b c d e. (if a then b else if c then d else e) = | 
| 53253 | 559 | (if c then (if a then b else d) else (if a then b else e))" | 
| 560 | by simp | |
| 68134 | 561 | from \<open>z \<notin> T\<close> have nz: "\<And>i. i \<in> T \<Longrightarrow> i \<noteq> z" | 
| 53253 | 562 | by auto | 
| 64267 | 563 | have "det (\<chi> i. if i \<in> insert z T then sum (a i) S else c i) = | 
| 564 | det (\<chi> i. if i = z then sum (a i) S else if i \<in> T then sum (a i) S else c i)" | |
| 33175 | 565 | unfolding insert_iff thif .. | 
| 64267 | 566 | also have "\<dots> = (\<Sum>j\<in>S. det (\<chi> i. if i \<in> T then sum (a i) S else if i = z then a i j else c i))" | 
| 567 | unfolding det_linear_row_sum[OF fS] | |
| 68134 | 568 | by (subst thif2) (simp add: nz cong: if_cong) | 
| 33175 | 569 | finally have tha: | 
| 64267 | 570 | "det (\<chi> i. if i \<in> insert z T then sum (a i) S else c i) = | 
| 33175 | 571 | (\<Sum>(j, f)\<in>S \<times> ?F T. det (\<chi> i. if i \<in> T then a i (f i) | 
| 572 | else if i = z then a i j | |
| 573 | else c i))" | |
| 64267 | 574 | unfolding insert.hyps unfolding sum.cartesian_product by blast | 
| 33175 | 575 | show ?case unfolding tha | 
| 60420 | 576 | using \<open>z \<notin> T\<close> | 
| 64267 | 577 | by (intro sum.reindex_bij_witness[where i="?k" and j="?h"]) | 
| 57129 
7edb7550663e
introduce more powerful reindexing rules for big operators
 hoelzl parents: 
56545diff
changeset | 578 | (auto intro!: cong[OF refl[of det]] simp: vec_eq_iff) | 
| 33175 | 579 | qed | 
| 580 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 581 | lemma det_linear_rows_sum: | 
| 53854 | 582 | fixes S :: "'n::finite set" | 
| 583 | assumes fS: "finite S" | |
| 64267 | 584 | shows "det (\<chi> i. sum (a i) S) = | 
| 585 |     sum (\<lambda>f. det (\<chi> i. a i (f i) :: 'a::comm_ring_1 ^ 'n^'n)) {f. \<forall>i. f i \<in> S}"
 | |
| 53253 | 586 | proof - | 
| 587 | 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)" | |
| 588 | by vector | |
| 64267 | 589 | from det_linear_rows_sum_lemma[OF fS, of "UNIV :: 'n set" a, unfolded th0, OF finite] | 
| 53253 | 590 | show ?thesis by simp | 
| 33175 | 591 | qed | 
| 592 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 593 | lemma matrix_mul_sum_alt: | 
| 34291 
4e896680897e
finite annotation on cartesian product is now implicit.
 hoelzl parents: 
34289diff
changeset | 594 | fixes A B :: "'a::comm_ring_1^'n^'n" | 
| 64267 | 595 | shows "A ** B = (\<chi> i. sum (\<lambda>k. A$i$k *s B $ k) (UNIV :: 'n set))" | 
| 596 | by (vector matrix_matrix_mult_def sum_component) | |
| 33175 | 597 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 598 | lemma det_rows_mul: | 
| 34291 
4e896680897e
finite annotation on cartesian product is now implicit.
 hoelzl parents: 
34289diff
changeset | 599 | "det((\<chi> i. c i *s a i)::'a::comm_ring_1^'n^'n) = | 
| 64272 | 600 | prod (\<lambda>i. c i) (UNIV:: 'n set) * det((\<chi> i. a i)::'a^'n^'n)" | 
| 601 | proof (simp add: det_def sum_distrib_left cong add: prod.cong, rule sum.cong) | |
| 33175 | 602 | let ?U = "UNIV :: 'n set" | 
| 603 |   let ?PU = "{p. p permutes ?U}"
 | |
| 53253 | 604 | fix p | 
| 605 | assume pU: "p \<in> ?PU" | |
| 33175 | 606 | let ?s = "of_int (sign p)" | 
| 53253 | 607 | from pU have p: "p permutes ?U" | 
| 608 | by blast | |
| 64272 | 609 | have "prod (\<lambda>i. c i * a i $ p i) ?U = prod c ?U * prod (\<lambda>i. a i $ p i) ?U" | 
| 610 | unfolding prod.distrib .. | |
| 33175 | 611 | then show "?s * (\<Prod>xa\<in>?U. c xa * a xa $ p xa) = | 
| 64272 | 612 | prod c ?U * (?s* (\<Prod>xa\<in>?U. a xa $ p xa))" | 
| 53854 | 613 | by (simp add: field_simps) | 
| 57418 | 614 | qed rule | 
| 33175 | 615 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 616 | proposition det_mul: | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 617 | fixes A B :: "'a::comm_ring_1^'n^'n" | 
| 33175 | 618 | shows "det (A ** B) = det A * det B" | 
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 619 | proof - | 
| 33175 | 620 | let ?U = "UNIV :: 'n set" | 
| 68134 | 621 |   let ?F = "{f. (\<forall>i \<in> ?U. f i \<in> ?U) \<and> (\<forall>i. i \<notin> ?U \<longrightarrow> f i = i)}"
 | 
| 33175 | 622 |   let ?PU = "{p. p permutes ?U}"
 | 
| 68134 | 623 | have "p \<in> ?F" if "p permutes ?U" for p | 
| 53854 | 624 | by simp | 
| 625 | then have PUF: "?PU \<subseteq> ?F" by blast | |
| 53253 | 626 |   {
 | 
| 627 | fix f | |
| 628 | assume fPU: "f \<in> ?F - ?PU" | |
| 53854 | 629 | have fUU: "f ` ?U \<subseteq> ?U" | 
| 630 | using fPU by auto | |
| 53253 | 631 | 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)" | 
| 632 | unfolding permutes_def by auto | |
| 33175 | 633 | |
| 634 | let ?A = "(\<chi> i. A$i$f i *s B$f i) :: 'a^'n^'n" | |
| 635 | let ?B = "(\<chi> i. B$f i) :: 'a^'n^'n" | |
| 53253 | 636 |     {
 | 
| 637 | assume fni: "\<not> inj_on f ?U" | |
| 33175 | 638 | then obtain i j where ij: "f i = f j" "i \<noteq> j" | 
| 639 | unfolding inj_on_def by blast | |
| 68134 | 640 | then have "row i ?B = row j ?B" | 
| 53854 | 641 | by (vector row_def) | 
| 68134 | 642 | with det_identical_rows[OF ij(2)] | 
| 33175 | 643 | have "det (\<chi> i. A$i$f i *s B$f i) = 0" | 
| 68134 | 644 | unfolding det_rows_mul by force | 
| 53253 | 645 | } | 
| 33175 | 646 | moreover | 
| 53253 | 647 |     {
 | 
| 648 | assume fi: "inj_on f ?U" | |
| 33175 | 649 | from f fi have fith: "\<And>i j. f i = f j \<Longrightarrow> i = j" | 
| 650 | unfolding inj_on_def by metis | |
| 68134 | 651 | note fs = fi[unfolded surjective_iff_injective_gen[OF finite finite refl fUU, symmetric]] | 
| 652 | have "\<exists>!x. f x = y" for y | |
| 653 | using fith fs by blast | |
| 53854 | 654 | with f(3) have "det (\<chi> i. A$i$f i *s B$f i) = 0" | 
| 655 | by blast | |
| 53253 | 656 | } | 
| 53854 | 657 | ultimately have "det (\<chi> i. A$i$f i *s B$f i) = 0" | 
| 658 | by blast | |
| 53253 | 659 | } | 
| 53854 | 660 | then have zth: "\<forall> f\<in> ?F - ?PU. det (\<chi> i. A$i$f i *s B$f i) = 0" | 
| 53253 | 661 | by simp | 
| 662 |   {
 | |
| 663 | fix p | |
| 664 | assume pU: "p \<in> ?PU" | |
| 53854 | 665 | from pU have p: "p permutes ?U" | 
| 666 | by blast | |
| 33175 | 667 | let ?s = "\<lambda>p. of_int (sign p)" | 
| 53253 | 668 | let ?f = "\<lambda>q. ?s p * (\<Prod>i\<in> ?U. A $ i $ p i) * (?s q * (\<Prod>i\<in> ?U. B $ i $ q i))" | 
| 64267 | 669 | have "(sum (\<lambda>q. ?s q * | 
| 53253 | 670 | (\<Prod>i\<in> ?U. (\<chi> i. A $ i $ p i *s B $ p i :: 'a^'n^'n) $ i $ q i)) ?PU) = | 
| 64267 | 671 | (sum (\<lambda>q. ?s p * (\<Prod>i\<in> ?U. A $ i $ p i) * (?s q * (\<Prod>i\<in> ?U. B $ i $ q i))) ?PU)" | 
| 33175 | 672 | unfolding sum_permutations_compose_right[OF permutes_inv[OF p], of ?f] | 
| 64267 | 673 | proof (rule sum.cong) | 
| 53253 | 674 | fix q | 
| 675 | assume qU: "q \<in> ?PU" | |
| 53854 | 676 | then have q: "q permutes ?U" | 
| 677 | by blast | |
| 33175 | 678 | from p q have pp: "permutation p" and pq: "permutation q" | 
| 679 | unfolding permutation_permutes by auto | |
| 680 | have th00: "of_int (sign p) * of_int (sign p) = (1::'a)" | |
| 681 | "\<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 | 682 | unfolding mult.assoc[symmetric] | 
| 53854 | 683 | unfolding of_int_mult[symmetric] | 
| 33175 | 684 | by (simp_all add: sign_idempotent) | 
| 53854 | 685 | have ths: "?s q = ?s p * ?s (q \<circ> inv p)" | 
| 33175 | 686 | using pp pq permutation_inverse[OF pp] sign_inverse[OF pp] | 
| 68134 | 687 | by (simp add: th00 ac_simps sign_idempotent sign_compose) | 
| 64272 | 688 | have th001: "prod (\<lambda>i. B$i$ q (inv p i)) ?U = prod ((\<lambda>i. B$i$ q (inv p i)) \<circ> p) ?U" | 
| 68134 | 689 | by (rule prod.permute[OF p]) | 
| 64272 | 690 | have thp: "prod (\<lambda>i. (\<chi> i. A$i$p i *s B$p i :: 'a^'n^'n) $i $ q i) ?U = | 
| 691 | prod (\<lambda>i. A$i$p i) ?U * prod (\<lambda>i. B$i$ q (inv p i)) ?U" | |
| 692 | unfolding th001 prod.distrib[symmetric] o_def permutes_inverses[OF p] | |
| 693 | apply (rule prod.cong[OF refl]) | |
| 53253 | 694 | using permutes_in_image[OF q] | 
| 695 | apply vector | |
| 696 | done | |
| 64272 | 697 | show "?s q * prod (\<lambda>i. (((\<chi> i. A$i$p i *s B$p i) :: 'a^'n^'n)$i$q i)) ?U = | 
| 698 | ?s p * (prod (\<lambda>i. A$i$p i) ?U) * (?s (q \<circ> inv p) * prod (\<lambda>i. B$i$(q \<circ> inv p) i) ?U)" | |
| 33175 | 699 | using ths thp pp pq permutation_inverse[OF pp] sign_inverse[OF pp] | 
| 36350 | 700 | by (simp add: sign_nz th00 field_simps sign_idempotent sign_compose) | 
| 57418 | 701 | qed rule | 
| 33175 | 702 | } | 
| 64267 | 703 | then have th2: "sum (\<lambda>f. det (\<chi> i. A$i$f i *s B$f i)) ?PU = det A * det B" | 
| 704 | unfolding det_def sum_product | |
| 705 | by (rule sum.cong [OF refl]) | |
| 706 | have "det (A**B) = sum (\<lambda>f. det (\<chi> i. A $ i $ f i *s B $ f i)) ?F" | |
| 68134 | 707 | unfolding matrix_mul_sum_alt det_linear_rows_sum[OF finite] | 
| 53854 | 708 | by simp | 
| 64267 | 709 | also have "\<dots> = sum (\<lambda>f. det (\<chi> i. A$i$f i *s B$f i)) ?PU" | 
| 68134 | 710 | using sum.mono_neutral_cong_left[OF finite PUF zth, symmetric] | 
| 33175 | 711 | unfolding det_rows_mul by auto | 
| 712 | finally show ?thesis unfolding th2 . | |
| 713 | qed | |
| 714 | ||
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 715 | |
| 69683 | 716 | subsection \<open>Relation to invertibility\<close> | 
| 33175 | 717 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 718 | proposition invertible_det_nz: | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 719 |   fixes A::"'a::{field}^'n^'n"
 | 
| 33175 | 720 | shows "invertible A \<longleftrightarrow> det A \<noteq> 0" | 
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 721 | proof (cases "invertible A") | 
| 68134 | 722 | case True | 
| 723 | then obtain B :: "'a^'n^'n" where B: "A ** B = mat 1" | |
| 724 | unfolding invertible_right_inverse by blast | |
| 725 | then have "det (A ** B) = det (mat 1 :: 'a^'n^'n)" | |
| 726 | by simp | |
| 727 | then show ?thesis | |
| 728 | by (metis True det_I det_mul mult_zero_left one_neq_zero) | |
| 729 | next | |
| 730 | case False | |
| 731 | let ?U = "UNIV :: 'n set" | |
| 732 | have fU: "finite ?U" | |
| 733 | by simp | |
| 734 | from False obtain c i where c: "sum (\<lambda>i. c i *s row i A) ?U = 0" and iU: "i \<in> ?U" and ci: "c i \<noteq> 0" | |
| 735 | unfolding invertible_right_inverse matrix_right_invertible_independent_rows | |
| 53854 | 736 | by blast | 
| 68134 | 737 |   have thr0: "- row i A = sum (\<lambda>j. (1/ c i) *s (c j *s row j A)) (?U - {i})"
 | 
| 68143 | 738 | unfolding sum_cmul using c ci | 
| 68138 | 739 | by (auto simp: sum.remove[OF fU iU] eq_vector_fraction_iff add_eq_0_iff) | 
| 68134 | 740 |   have thr: "- row i A \<in> vec.span {row j A| j. j \<noteq> i}"
 | 
| 741 | unfolding thr0 by (auto intro: vec.span_base vec.span_scale vec.span_sum) | |
| 742 | let ?B = "(\<chi> k. if k = i then 0 else row k A) :: 'a^'n^'n" | |
| 743 | have thrb: "row i ?B = 0" using iU by (vector row_def) | |
| 744 | have "det A = 0" | |
| 745 | unfolding det_row_span[OF thr, symmetric] right_minus | |
| 746 | unfolding det_zero_row(2)[OF thrb] .. | |
| 747 | then show ?thesis | |
| 748 | by (simp add: False) | |
| 33175 | 749 | qed | 
| 750 | ||
| 68134 | 751 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 752 | lemma det_nz_iff_inj_gen: | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 753 | fixes f :: "'a::field^'n \<Rightarrow> 'a::field^'n" | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 754 | assumes "Vector_Spaces.linear (*s) (*s) f" | 
| 67990 | 755 | shows "det (matrix f) \<noteq> 0 \<longleftrightarrow> inj f" | 
| 756 | proof | |
| 757 | assume "det (matrix f) \<noteq> 0" | |
| 758 | then show "inj f" | |
| 759 | using assms invertible_det_nz inj_matrix_vector_mult by force | |
| 760 | next | |
| 761 | assume "inj f" | |
| 762 | show "det (matrix f) \<noteq> 0" | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 763 | using vec.linear_injective_left_inverse [OF assms \<open>inj f\<close>] | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 764 | by (metis assms invertible_det_nz invertible_left_inverse matrix_compose_gen matrix_id_mat_1) | 
| 67990 | 765 | qed | 
| 766 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 767 | lemma det_nz_iff_inj: | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 768 | fixes f :: "real^'n \<Rightarrow> real^'n" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 769 | assumes "linear f" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 770 | shows "det (matrix f) \<noteq> 0 \<longleftrightarrow> inj f" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 771 | using det_nz_iff_inj_gen[of f] assms | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 772 | unfolding linear_matrix_vector_mul_eq . | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 773 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 774 | lemma det_eq_0_rank: | 
| 67990 | 775 | fixes A :: "real^'n^'n" | 
| 776 |   shows "det A = 0 \<longleftrightarrow> rank A < CARD('n)"
 | |
| 777 | using invertible_det_nz [of A] | |
| 778 | by (auto simp: matrix_left_invertible_injective invertible_left_inverse less_rank_noninjective) | |
| 779 | ||
| 70136 | 780 | subsubsection\<^marker>\<open>tag important\<close> \<open>Invertibility of matrices and corresponding linear functions\<close> | 
| 67981 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 781 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 782 | lemma matrix_left_invertible_gen: | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 783 | fixes f :: "'a::field^'m \<Rightarrow> 'a::field^'n" | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 784 | assumes "Vector_Spaces.linear (*s) (*s) f" | 
| 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 785 | shows "((\<exists>B. B ** matrix f = mat 1) \<longleftrightarrow> (\<exists>g. Vector_Spaces.linear (*s) (*s) g \<and> g \<circ> f = id))" | 
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 786 | proof safe | 
| 67981 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 787 | fix B | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 788 | assume 1: "B ** matrix f = mat 1" | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 789 | show "\<exists>g. Vector_Spaces.linear (*s) (*s) g \<and> g \<circ> f = id" | 
| 67981 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 790 | proof (intro exI conjI) | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 791 | show "Vector_Spaces.linear (*s) (*s) (\<lambda>y. B *v y)" | 
| 68138 | 792 | by simp | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 793 | show "((*v) B) \<circ> f = id" | 
| 67981 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 794 | unfolding o_def | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 795 | by (metis assms 1 eq_id_iff matrix_vector_mul(1) matrix_vector_mul_assoc matrix_vector_mul_lid) | 
| 67981 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 796 | qed | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 797 | next | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 798 | fix g | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 799 | assume "Vector_Spaces.linear (*s) (*s) g" "g \<circ> f = id" | 
| 67981 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 800 | then have "matrix g ** matrix f = mat 1" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 801 | by (metis assms matrix_compose_gen matrix_id_mat_1) | 
| 67981 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 802 | then show "\<exists>B. B ** matrix f = mat 1" .. | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 803 | qed | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 804 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 805 | lemma matrix_left_invertible: | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 806 | "linear f \<Longrightarrow> ((\<exists>B. B ** matrix f = mat 1) \<longleftrightarrow> (\<exists>g. linear g \<and> g \<circ> f = id))" for f::"real^'m \<Rightarrow> real^'n" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 807 | using matrix_left_invertible_gen[of f] | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 808 | by (auto simp: linear_matrix_vector_mul_eq) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 809 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 810 | lemma matrix_right_invertible_gen: | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 811 | fixes f :: "'a::field^'m \<Rightarrow> 'a^'n" | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 812 | assumes "Vector_Spaces.linear (*s) (*s) f" | 
| 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 813 | shows "((\<exists>B. matrix f ** B = mat 1) \<longleftrightarrow> (\<exists>g. Vector_Spaces.linear (*s) (*s) g \<and> f \<circ> g = id))" | 
| 67981 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 814 | proof safe | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 815 | fix B | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 816 | assume 1: "matrix f ** B = mat 1" | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 817 | show "\<exists>g. Vector_Spaces.linear (*s) (*s) g \<and> f \<circ> g = id" | 
| 67981 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 818 | proof (intro exI conjI) | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 819 | show "Vector_Spaces.linear (*s) (*s) ((*v) B)" | 
| 68138 | 820 | by simp | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 821 | show "f \<circ> (*v) B = id" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 822 | using 1 assms comp_apply eq_id_iff vec.linear_id matrix_id_mat_1 matrix_vector_mul_assoc matrix_works | 
| 73932 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 desharna parents: 
73648diff
changeset | 823 | by (metis (no_types, opaque_lifting)) | 
| 67981 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 824 | qed | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 825 | next | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 826 | fix g | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 827 | assume "Vector_Spaces.linear (*s) (*s) g" and "f \<circ> g = id" | 
| 67981 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 828 | then have "matrix f ** matrix g = mat 1" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 829 | by (metis assms matrix_compose_gen matrix_id_mat_1) | 
| 67981 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 830 | then show "\<exists>B. matrix f ** B = mat 1" .. | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 831 | qed | 
| 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 832 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 833 | lemma matrix_right_invertible: | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 834 | "linear f \<Longrightarrow> ((\<exists>B. matrix f ** B = mat 1) \<longleftrightarrow> (\<exists>g. linear g \<and> f \<circ> g = id))" for f::"real^'m \<Rightarrow> real^'n" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 835 | using matrix_right_invertible_gen[of f] | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 836 | by (auto simp: linear_matrix_vector_mul_eq) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 837 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 838 | lemma matrix_invertible_gen: | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 839 | fixes f :: "'a::field^'m \<Rightarrow> 'a::field^'n" | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 840 | assumes "Vector_Spaces.linear (*s) (*s) f" | 
| 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 841 | shows "invertible (matrix f) \<longleftrightarrow> (\<exists>g. Vector_Spaces.linear (*s) (*s) g \<and> f \<circ> g = id \<and> g \<circ> f = id)" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 842 | (is "?lhs = ?rhs") | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 843 | proof | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 844 | assume ?lhs then show ?rhs | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 845 | by (metis assms invertible_def left_right_inverse_eq matrix_left_invertible_gen matrix_right_invertible_gen) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 846 | next | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 847 | assume ?rhs then show ?lhs | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 848 | by (metis assms invertible_def matrix_compose_gen matrix_id_mat_1) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 849 | qed | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 850 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 851 | lemma matrix_invertible: | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 852 | "linear f \<Longrightarrow> invertible (matrix f) \<longleftrightarrow> (\<exists>g. linear g \<and> f \<circ> g = id \<and> g \<circ> f = id)" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 853 | for f::"real^'m \<Rightarrow> real^'n" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 854 | using matrix_invertible_gen[of f] | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 855 | by (auto simp: linear_matrix_vector_mul_eq) | 
| 67981 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 856 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 857 | lemma invertible_eq_bij: | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 858 | fixes m :: "'a::field^'m^'n" | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68833diff
changeset | 859 | shows "invertible m \<longleftrightarrow> bij ((*v) m)" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 860 | using matrix_invertible_gen[OF matrix_vector_mul_linear_gen, of m, simplified matrix_of_matrix_vector_mul] | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 861 | by (metis bij_betw_def left_right_inverse_eq matrix_vector_mul_linear_gen o_bij | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 862 | vec.linear_injective_left_inverse vec.linear_surjective_right_inverse) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 863 | |
| 67981 
349c639e593c
more new theorems on real^1, matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67971diff
changeset | 864 | |
| 69683 | 865 | subsection \<open>Cramer's rule\<close> | 
| 33175 | 866 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 867 | lemma cramer_lemma_transpose: | 
| 68263 | 868 |   fixes A:: "'a::{field}^'n^'n"
 | 
| 869 |     and x :: "'a::{field}^'n"
 | |
| 64267 | 870 | shows "det ((\<chi> i. if i = k then sum (\<lambda>i. x$i *s row i A) (UNIV::'n set) | 
| 68263 | 871 |                              else row i A)::'a::{field}^'n^'n) = x$k * det A"
 | 
| 33175 | 872 | (is "?lhs = ?rhs") | 
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 873 | proof - | 
| 33175 | 874 | let ?U = "UNIV :: 'n set" | 
| 875 |   let ?Uk = "?U - {k}"
 | |
| 53854 | 876 | have U: "?U = insert k ?Uk" | 
| 877 | by blast | |
| 878 | have kUk: "k \<notin> ?Uk" | |
| 879 | by simp | |
| 33175 | 880 | have th00: "\<And>k s. x$k *s row k A + s = (x$k - 1) *s row k A + row k A + s" | 
| 36350 | 881 | by (vector field_simps) | 
| 53854 | 882 | have th001: "\<And>f k . (\<lambda>x. if x = k then f k else f x) = f" | 
| 883 | by auto | |
| 33175 | 884 | have "(\<chi> i. row i A) = A" by (vector row_def) | 
| 53253 | 885 | then have thd1: "det (\<chi> i. row i A) = det A" | 
| 886 | by simp | |
| 33175 | 887 | 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" | 
| 68134 | 888 | by (force intro: det_row_span vec.span_sum vec.span_scale vec.span_base) | 
| 33175 | 889 | show "?lhs = x$k * det A" | 
| 890 | apply (subst U) | |
| 68134 | 891 | unfolding sum.insert[OF finite kUk] | 
| 33175 | 892 | apply (subst th00) | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57418diff
changeset | 893 | unfolding add.assoc | 
| 33175 | 894 | apply (subst det_row_add) | 
| 895 | unfolding thd0 | |
| 896 | unfolding det_row_mul | |
| 897 | unfolding th001[of k "\<lambda>i. row i A"] | |
| 53253 | 898 | unfolding thd1 | 
| 899 | apply (simp add: field_simps) | |
| 900 | done | |
| 33175 | 901 | qed | 
| 902 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 903 | proposition cramer_lemma: | 
| 68263 | 904 |   fixes A :: "'a::{field}^'n^'n"
 | 
| 905 |   shows "det((\<chi> i j. if j = k then (A *v x)$i else A$i$j):: 'a::{field}^'n^'n) = x$k * det A"
 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 906 | proof - | 
| 33175 | 907 | let ?U = "UNIV :: 'n set" | 
| 64267 | 908 | have *: "\<And>c. sum (\<lambda>i. c i *s row i (transpose A)) ?U = sum (\<lambda>i. c i *s column i A) ?U" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 909 | by (auto intro: sum.cong) | 
| 53854 | 910 | show ?thesis | 
| 67673 
c8caefb20564
lots of new material, ultimately related to measure theory
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 911 | unfolding matrix_mult_sum | 
| 53253 | 912 | unfolding cramer_lemma_transpose[of k x "transpose A", unfolded det_transpose, symmetric] | 
| 913 | unfolding *[of "\<lambda>i. x$i"] | |
| 914 | apply (subst det_transpose[symmetric]) | |
| 915 | apply (rule cong[OF refl[of det]]) | |
| 916 | apply (vector transpose_def column_def row_def) | |
| 917 | done | |
| 33175 | 918 | qed | 
| 919 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 920 | proposition cramer: | 
| 68263 | 921 |   fixes A ::"'a::{field}^'n^'n"
 | 
| 33175 | 922 | assumes d0: "det A \<noteq> 0" | 
| 36362 
06475a1547cb
fix lots of looping simp calls and other warnings
 huffman parents: 
35542diff
changeset | 923 | 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)" | 
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 924 | proof - | 
| 33175 | 925 | from d0 obtain B where B: "A ** B = mat 1" "B ** A = mat 1" | 
| 53854 | 926 | unfolding invertible_det_nz[symmetric] invertible_def | 
| 927 | by blast | |
| 928 | have "(A ** B) *v b = b" | |
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 929 | by (simp add: B) | 
| 53854 | 930 | then have "A *v (B *v b) = b" | 
| 931 | by (simp add: matrix_vector_mul_assoc) | |
| 932 | then have xe: "\<exists>x. A *v x = b" | |
| 933 | by blast | |
| 53253 | 934 |   {
 | 
| 935 | fix x | |
| 936 | assume x: "A *v x = b" | |
| 937 | have "x = (\<chi> k. det(\<chi> i j. if j=k then b$i else A$i$j) / det A)" | |
| 938 | unfolding x[symmetric] | |
| 939 | using d0 by (simp add: vec_eq_iff cramer_lemma field_simps) | |
| 940 | } | |
| 53854 | 941 | with xe show ?thesis | 
| 942 | by auto | |
| 33175 | 943 | qed | 
| 944 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 945 | lemma det_1: "det (A::'a::comm_ring_1^1^1) = A$1$1" | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 946 | by (simp add: det_def sign_id) | 
| 33175 | 947 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 948 | lemma det_2: "det (A::'a::comm_ring_1^2^2) = A$1$1 * A$2$2 - A$1$2 * A$2$1" | 
| 53253 | 949 | proof - | 
| 33175 | 950 |   have f12: "finite {2::2}" "1 \<notin> {2::2}" by auto
 | 
| 951 | show ?thesis | |
| 53253 | 952 | unfolding det_def UNIV_2 | 
| 64267 | 953 | unfolding sum_over_permutations_insert[OF f12] | 
| 53253 | 954 | unfolding permutes_sing | 
| 955 | by (simp add: sign_swap_id sign_id swap_id_eq) | |
| 33175 | 956 | qed | 
| 957 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 958 | lemma det_3: | 
| 53253 | 959 | "det (A::'a::comm_ring_1^3^3) = | 
| 960 | A$1$1 * A$2$2 * A$3$3 + | |
| 961 | A$1$2 * A$2$3 * A$3$1 + | |
| 962 | A$1$3 * A$2$1 * A$3$2 - | |
| 963 | A$1$1 * A$2$3 * A$3$2 - | |
| 964 | A$1$2 * A$2$1 * A$3$3 - | |
| 965 | A$1$3 * A$2$2 * A$3$1" | |
| 966 | proof - | |
| 53854 | 967 |   have f123: "finite {2::3, 3}" "1 \<notin> {2::3, 3}"
 | 
| 968 | by auto | |
| 969 |   have f23: "finite {3::3}" "2 \<notin> {3::3}"
 | |
| 970 | by auto | |
| 33175 | 971 | |
| 972 | show ?thesis | |
| 53253 | 973 | unfolding det_def UNIV_3 | 
| 64267 | 974 | unfolding sum_over_permutations_insert[OF f123] | 
| 975 | unfolding sum_over_permutations_insert[OF f23] | |
| 53253 | 976 | unfolding permutes_sing | 
| 977 | by (simp add: sign_swap_id permutation_swap_id sign_compose sign_id swap_id_eq) | |
| 33175 | 978 | qed | 
| 979 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 980 | proposition det_orthogonal_matrix: | 
| 69680 | 981 | fixes Q:: "'a::linordered_idom^'n^'n" | 
| 982 | assumes oQ: "orthogonal_matrix Q" | |
| 983 | shows "det Q = 1 \<or> det Q = - 1" | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 984 | proof - | 
| 69680 | 985 | have "Q ** transpose Q = mat 1" | 
| 986 | by (metis oQ orthogonal_matrix_def) | |
| 987 | then have "det (Q ** transpose Q) = det (mat 1:: 'a^'n^'n)" | |
| 988 | by simp | |
| 989 | then have "det Q * det Q = 1" | |
| 990 | by (simp add: det_mul) | |
| 991 | then show ?thesis | |
| 992 | by (simp add: square_eq_1_iff) | |
| 993 | qed | |
| 994 | ||
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 995 | proposition orthogonal_transformation_det [simp]: | 
| 69680 | 996 | fixes f :: "real^'n \<Rightarrow> real^'n" | 
| 997 | shows "orthogonal_transformation f \<Longrightarrow> \<bar>det (matrix f)\<bar> = 1" | |
| 70136 | 998 | using det_orthogonal_matrix orthogonal_transformation_matrix by fastforce | 
| 69680 | 999 | |
| 69683 | 1000 | subsection \<open>Rotation, reflection, rotoinversion\<close> | 
| 69680 | 1001 | |
| 70136 | 1002 | definition\<^marker>\<open>tag important\<close> "rotation_matrix Q \<longleftrightarrow> orthogonal_matrix Q \<and> det Q = 1" | 
| 1003 | definition\<^marker>\<open>tag important\<close> "rotoinversion_matrix Q \<longleftrightarrow> orthogonal_matrix Q \<and> det Q = - 1" | |
| 69680 | 1004 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 1005 | lemma orthogonal_rotation_or_rotoinversion: | 
| 69680 | 1006 | fixes Q :: "'a::linordered_idom^'n^'n" | 
| 1007 | shows " orthogonal_matrix Q \<longleftrightarrow> rotation_matrix Q \<or> rotoinversion_matrix Q" | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 1008 | by (metis rotoinversion_matrix_def rotation_matrix_def det_orthogonal_matrix) | 
| 69680 | 1009 | |
| 68134 | 1010 | text\<open> Slightly stronger results giving rotation, but only in two or more dimensions\<close> | 
| 67683 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1011 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 1012 | lemma rotation_matrix_exists_basis: | 
| 67683 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1013 | fixes a :: "real^'n" | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1014 |   assumes 2: "2 \<le> CARD('n)" and "norm a = 1"
 | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1015 | obtains A where "rotation_matrix A" "A *v (axis k 1) = a" | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1016 | proof - | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1017 | obtain A where "orthogonal_matrix A" and A: "A *v (axis k 1) = a" | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1018 | using orthogonal_matrix_exists_basis assms by metis | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1019 | with orthogonal_rotation_or_rotoinversion | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1020 | consider "rotation_matrix A" | "rotoinversion_matrix A" | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1021 | by metis | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1022 | then show thesis | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1023 | proof cases | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1024 | assume "rotation_matrix A" | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1025 | then show ?thesis | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1026 | using \<open>A *v axis k 1 = a\<close> that by auto | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1027 | next | 
| 76836 
30182f9e1818
Big simplifications of old proofs
 paulson <lp15@cam.ac.uk> parents: 
73932diff
changeset | 1028 | from obtain_subset_with_card_n[OF 2] obtain h i::'n where "h \<noteq> i" | 
| 76837 | 1029 | by (fastforce simp add: eval_nat_numeral card_Suc_eq) | 
| 69680 | 1030 | then obtain j where "j \<noteq> k" | 
| 1031 | by (metis (full_types)) | |
| 67683 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1032 | let ?TA = "transpose A" | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1033 | let ?A = "\<chi> i. if i = j then - 1 *\<^sub>R (?TA $ i) else ?TA $i" | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1034 | assume "rotoinversion_matrix A" | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1035 | then have [simp]: "det A = -1" | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1036 | by (simp add: rotoinversion_matrix_def) | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1037 | show ?thesis | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1038 | proof | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1039 | have [simp]: "row i (\<chi> i. if i = j then - 1 *\<^sub>R ?TA $ i else ?TA $ i) = (if i = j then - row i ?TA else row i ?TA)" for i | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1040 | by (auto simp: row_def) | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1041 | have "orthogonal_matrix ?A" | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1042 | unfolding orthogonal_matrix_orthonormal_rows | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1043 | using \<open>orthogonal_matrix A\<close> by (auto simp: orthogonal_matrix_orthonormal_columns orthogonal_clauses) | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1044 | then show "rotation_matrix (transpose ?A)" | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1045 | unfolding rotation_matrix_def | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1046 | by (simp add: det_row_mul[of j _ "\<lambda>i. ?TA $ i", unfolded scalar_mult_eq_scaleR]) | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1047 | show "transpose ?A *v axis k 1 = a" | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1048 | using \<open>j \<noteq> k\<close> A by (simp add: matrix_vector_column axis_def scalar_mult_eq_scaleR if_distrib [of "\<lambda>z. z *\<^sub>R c" for c] cong: if_cong) | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1049 | qed | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1050 | qed | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1051 | qed | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1052 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 1053 | lemma rotation_exists_1: | 
| 67683 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1054 | fixes a :: "real^'n" | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1055 |   assumes "2 \<le> CARD('n)" "norm a = 1" "norm b = 1"
 | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1056 | obtains f where "orthogonal_transformation f" "det(matrix f) = 1" "f a = b" | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1057 | proof - | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1058 | obtain k::'n where True | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1059 | by simp | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1060 | obtain A B where AB: "rotation_matrix A" "rotation_matrix B" | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1061 | and eq: "A *v (axis k 1) = a" "B *v (axis k 1) = b" | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1062 | using rotation_matrix_exists_basis assms by metis | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1063 | let ?f = "\<lambda>x. (B ** transpose A) *v x" | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1064 | show thesis | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1065 | proof | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1066 | show "orthogonal_transformation ?f" | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1067 | using AB orthogonal_matrix_mul orthogonal_transformation_matrix rotation_matrix_def matrix_vector_mul_linear by force | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1068 | show "det (matrix ?f) = 1" | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1069 | using AB by (auto simp: det_mul rotation_matrix_def) | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1070 | show "?f a = b" | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1071 | using AB unfolding orthogonal_matrix_def rotation_matrix_def | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1072 | by (metis eq matrix_mul_rid matrix_vector_mul_assoc) | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1073 | qed | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1074 | qed | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1075 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 1076 | lemma rotation_exists: | 
| 67683 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1077 | fixes a :: "real^'n" | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1078 |   assumes 2: "2 \<le> CARD('n)" and eq: "norm a = norm b"
 | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1079 | obtains f where "orthogonal_transformation f" "det(matrix f) = 1" "f a = b" | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1080 | proof (cases "a = 0 \<or> b = 0") | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1081 | case True | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1082 | with assms have "a = 0" "b = 0" | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1083 | by auto | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1084 | then show ?thesis | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1085 | by (metis eq_id_iff matrix_id orthogonal_transformation_id that) | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1086 | next | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1087 | case False | 
| 68072 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 1088 | then obtain f where f: "orthogonal_transformation f" "det (matrix f) = 1" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 1089 | and f': "f (a /\<^sub>R norm a) = b /\<^sub>R norm b" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 1090 | using rotation_exists_1 [of "a /\<^sub>R norm a" "b /\<^sub>R norm b", OF 2] by auto | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 1091 | then interpret linear f by (simp add: orthogonal_transformation) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 1092 | have "f a = b" | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 1093 | using f' False | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 1094 | by (simp add: eq scale) | 
| 
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
 immler parents: 
67990diff
changeset | 1095 | with f show thesis .. | 
| 67683 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1096 | qed | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1097 | |
| 69720 
be6634e99e09
redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69683diff
changeset | 1098 | lemma rotation_rightward_line: | 
| 67683 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1099 | fixes a :: "real^'n" | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1100 |   obtains f where "orthogonal_transformation f" "2 \<le> CARD('n) \<Longrightarrow> det(matrix f) = 1"
 | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1101 | "f(norm a *\<^sub>R axis k 1) = a" | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1102 | proof (cases "CARD('n) = 1")
 | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1103 | case True | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1104 | obtain f where "orthogonal_transformation f" "f (norm a *\<^sub>R axis k (1::real)) = a" | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1105 | proof (rule orthogonal_transformation_exists) | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1106 | show "norm (norm a *\<^sub>R axis k (1::real)) = norm a" | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1107 | by simp | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1108 | qed auto | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1109 | then show thesis | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1110 | using True that by auto | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1111 | next | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1112 | case False | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1113 | obtain f where "orthogonal_transformation f" "det(matrix f) = 1" "f (norm a *\<^sub>R axis k 1) = a" | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1114 | proof (rule rotation_exists) | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1115 |     show "2 \<le> CARD('n)"
 | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1116 | using False one_le_card_finite [where 'a='n] by linarith | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1117 | show "norm (norm a *\<^sub>R axis k (1::real)) = norm a" | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1118 | by simp | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1119 | qed auto | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1120 | then show thesis | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1121 | using that by blast | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1122 | qed | 
| 
817944aeac3f
Lots of new material about matrices, etc.
 paulson <lp15@cam.ac.uk> parents: 
67673diff
changeset | 1123 | |
| 33175 | 1124 | end |