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