1 section \<open>Instantiates the finite Cartesian product of Euclidean spaces as a Euclidean space.\<close>
3 theory Cartesian_Euclidean_Space
4 imports Finite_Cartesian_Product Integration
7 lemma subspace_special_hyperplane: "subspace {x. x $ k = 0}"
8 by (simp add: subspace_def)
10 lemma delta_mult_idempotent:
11 "(if k=a then 1 else (0::'a::semiring_1)) * (if k=a then 1 else 0) = (if k=a then 1 else 0)"
14 lemma setsum_UNIV_sum:
15 fixes g :: "'a::finite + 'b::finite \<Rightarrow> _"
16 shows "(\<Sum>x\<in>UNIV. g x) = (\<Sum>x\<in>UNIV. g (Inl x)) + (\<Sum>x\<in>UNIV. g (Inr x))"
17 apply (subst UNIV_Plus_UNIV [symmetric])
18 apply (subst setsum.Plus)
22 lemma setsum_mult_product:
23 "setsum h {..<A * B :: nat} = (\<Sum>i\<in>{..<A}. \<Sum>j\<in>{..<B}. h (j + i * B))"
24 unfolding setsum_nat_group[of h B A, unfolded atLeast0LessThan, symmetric]
25 proof (rule setsum.cong, simp, rule setsum.reindex_cong)
27 show "inj_on (\<lambda>j. j + i * B) {..<B}" by (auto intro!: inj_onI)
28 show "{i * B..<i * B + B} = (\<lambda>j. j + i * B) ` {..<B}"
30 fix j assume "j \<in> {i * B..<i * B + B}"
31 then show "j \<in> (\<lambda>j. j + i * B) ` {..<B}"
32 by (auto intro!: image_eqI[of _ _ "j - i * B"])
37 subsection\<open>Basic componentwise operations on vectors.\<close>
39 instantiation vec :: (times, finite) times
42 definition "op * \<equiv> (\<lambda> x y. (\<chi> i. (x$i) * (y$i)))"
47 instantiation vec :: (one, finite) one
50 definition "1 \<equiv> (\<chi> i. 1)"
55 instantiation vec :: (ord, finite) ord
58 definition "x \<le> y \<longleftrightarrow> (\<forall>i. x$i \<le> y$i)"
59 definition "x < (y::'a^'b) \<longleftrightarrow> x \<le> y \<and> \<not> y \<le> x"
64 text\<open>The ordering on one-dimensional vectors is linear.\<close>
67 assumes UNIV_one: "card (UNIV :: 'a set) = Suc 0"
72 from UNIV_one show "finite (UNIV :: 'a set)"
73 by (auto intro!: card_ge_0_finite)
78 instance vec:: (order, finite) order
79 by standard (auto simp: less_eq_vec_def less_vec_def vec_eq_iff
80 intro: order.trans order.antisym order.strict_implies_order)
82 instance vec :: (linorder, cart_one) linorder
84 obtain a :: 'b where all: "\<And>P. (\<forall>i. P i) \<longleftrightarrow> P a"
86 have "card (UNIV :: 'b set) = Suc 0" by (rule UNIV_one)
87 then obtain b :: 'b where "UNIV = {b}" by (auto iff: card_Suc_eq)
88 then have "\<And>P. (\<forall>i\<in>UNIV. P i) \<longleftrightarrow> P b" by auto
89 then show thesis by (auto intro: that)
91 fix x y :: "'a^'b::cart_one"
92 note [simp] = less_eq_vec_def less_vec_def all vec_eq_iff field_simps
93 show "x \<le> y \<or> y \<le> x" by auto
96 text\<open>Constant Vectors\<close>
98 definition "vec x = (\<chi> i. x)"
100 lemma interval_cbox_cart: "{a::real^'n..b} = cbox a b"
101 by (auto simp add: less_eq_vec_def mem_box Basis_vec_def inner_axis)
103 text\<open>Also the scalar-vector multiplication.\<close>
105 definition vector_scalar_mult:: "'a::times \<Rightarrow> 'a ^ 'n \<Rightarrow> 'a ^ 'n" (infixl "*s" 70)
106 where "c *s x = (\<chi> i. c * (x$i))"
109 subsection \<open>A naive proof procedure to lift really trivial arithmetic stuff from the basis of the vector space.\<close>
111 lemma setsum_cong_aux:
112 "(\<And>x. x \<in> A \<Longrightarrow> f x = g x) \<Longrightarrow> setsum f A = setsum g A"
113 by (auto intro: setsum.cong)
115 hide_fact (open) setsum_cong_aux
117 method_setup vector = \<open>
120 simpset_of (put_simpset HOL_basic_ss @{context}
121 addsimps [@{thm setsum.distrib} RS sym,
122 @{thm setsum_subtractf} RS sym, @{thm setsum_right_distrib},
123 @{thm setsum_left_distrib}, @{thm setsum_negf} RS sym])
125 simpset_of (@{context} addsimps
126 [@{thm plus_vec_def}, @{thm times_vec_def},
127 @{thm minus_vec_def}, @{thm uminus_vec_def},
128 @{thm one_vec_def}, @{thm zero_vec_def}, @{thm vec_def},
129 @{thm scaleR_vec_def},
130 @{thm vec_lambda_beta}, @{thm vector_scalar_mult_def}])
131 fun vector_arith_tac ctxt ths =
132 simp_tac (put_simpset ss1 ctxt)
133 THEN' (fn i => resolve_tac ctxt @{thms Cartesian_Euclidean_Space.setsum_cong_aux} i
134 ORELSE resolve_tac ctxt @{thms setsum.neutral} i
135 ORELSE simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm vec_eq_iff}]) i)
136 (* THEN' TRY o clarify_tac HOL_cs THEN' (TRY o rtac @{thm iffI}) *)
137 THEN' asm_full_simp_tac (put_simpset ss2 ctxt addsimps ths)
139 Attrib.thms >> (fn ths => fn ctxt => SIMPLE_METHOD' (vector_arith_tac ctxt ths))
141 \<close> "lift trivial vector statements to real arith statements"
143 lemma vec_0[simp]: "vec 0 = 0" by vector
144 lemma vec_1[simp]: "vec 1 = 1" by vector
146 lemma vec_inj[simp]: "vec x = vec y \<longleftrightarrow> x = y" by vector
148 lemma vec_in_image_vec: "vec x \<in> (vec ` S) \<longleftrightarrow> x \<in> S" by auto
150 lemma vec_add: "vec(x + y) = vec x + vec y" by vector
151 lemma vec_sub: "vec(x - y) = vec x - vec y" by vector
152 lemma vec_cmul: "vec(c * x) = c *s vec x " by vector
153 lemma vec_neg: "vec(- x) = - vec x " by vector
157 shows "vec(setsum f S) = setsum (vec \<circ> f) S"
161 then show ?case by simp
164 then show ?case by (auto simp add: vec_add)
167 text\<open>Obvious "component-pushing".\<close>
169 lemma vec_component [simp]: "vec x $ i = x"
172 lemma vector_mult_component [simp]: "(x * y)$i = x$i * y$i"
175 lemma vector_smult_component [simp]: "(c *s y)$i = c * (y$i)"
178 lemma cond_component: "(if b then x else y)$i = (if b then x$i else y$i)" by vector
180 lemmas vector_component =
181 vec_component vector_add_component vector_mult_component
182 vector_smult_component vector_minus_component vector_uminus_component
183 vector_scaleR_component cond_component
186 subsection \<open>Some frequently useful arithmetic lemmas over vectors.\<close>
188 instance vec :: (semigroup_mult, finite) semigroup_mult
189 by standard (vector mult.assoc)
191 instance vec :: (monoid_mult, finite) monoid_mult
194 instance vec :: (ab_semigroup_mult, finite) ab_semigroup_mult
195 by standard (vector mult.commute)
197 instance vec :: (comm_monoid_mult, finite) comm_monoid_mult
200 instance vec :: (semiring, finite) semiring
201 by standard (vector field_simps)+
203 instance vec :: (semiring_0, finite) semiring_0
204 by standard (vector field_simps)+
205 instance vec :: (semiring_1, finite) semiring_1
207 instance vec :: (comm_semiring, finite) comm_semiring
208 by standard (vector field_simps)+
210 instance vec :: (comm_semiring_0, finite) comm_semiring_0 ..
211 instance vec :: (cancel_comm_monoid_add, finite) cancel_comm_monoid_add ..
212 instance vec :: (semiring_0_cancel, finite) semiring_0_cancel ..
213 instance vec :: (comm_semiring_0_cancel, finite) comm_semiring_0_cancel ..
214 instance vec :: (ring, finite) ring ..
215 instance vec :: (semiring_1_cancel, finite) semiring_1_cancel ..
216 instance vec :: (comm_semiring_1, finite) comm_semiring_1 ..
218 instance vec :: (ring_1, finite) ring_1 ..
220 instance vec :: (real_algebra, finite) real_algebra
221 by standard (simp_all add: vec_eq_iff)
223 instance vec :: (real_algebra_1, finite) real_algebra_1 ..
225 lemma of_nat_index: "(of_nat n :: 'a::semiring_1 ^'n)$i = of_nat n"
228 then show ?case by vector
231 then show ?case by vector
234 lemma one_index [simp]: "(1 :: 'a :: one ^ 'n) $ i = 1"
237 lemma neg_one_index [simp]: "(- 1 :: 'a :: {one, uminus} ^ 'n) $ i = - 1"
240 instance vec :: (semiring_char_0, finite) semiring_char_0
243 show "inj (of_nat :: nat \<Rightarrow> 'a ^ 'b)"
244 by (auto intro!: injI simp add: vec_eq_iff of_nat_index)
247 instance vec :: (numeral, finite) numeral ..
248 instance vec :: (semiring_numeral, finite) semiring_numeral ..
250 lemma numeral_index [simp]: "numeral w $ i = numeral w"
251 by (induct w) (simp_all only: numeral.simps vector_add_component one_index)
253 lemma neg_numeral_index [simp]: "- numeral w $ i = - numeral w"
254 by (simp only: vector_uminus_component numeral_index)
256 instance vec :: (comm_ring_1, finite) comm_ring_1 ..
257 instance vec :: (ring_char_0, finite) ring_char_0 ..
259 lemma vector_smult_assoc: "a *s (b *s x) = ((a::'a::semigroup_mult) * b) *s x"
260 by (vector mult.assoc)
261 lemma vector_sadd_rdistrib: "((a::'a::semiring) + b) *s x = a *s x + b *s x"
262 by (vector field_simps)
263 lemma vector_add_ldistrib: "(c::'a::semiring) *s (x + y) = c *s x + c *s y"
264 by (vector field_simps)
265 lemma vector_smult_lzero[simp]: "(0::'a::mult_zero) *s x = 0" by vector
266 lemma vector_smult_lid[simp]: "(1::'a::monoid_mult) *s x = x" by vector
267 lemma vector_ssub_ldistrib: "(c::'a::ring) *s (x - y) = c *s x - c *s y"
268 by (vector field_simps)
269 lemma vector_smult_rneg: "(c::'a::ring) *s -x = -(c *s x)" by vector
270 lemma vector_smult_lneg: "- (c::'a::ring) *s x = -(c *s x)" by vector
271 lemma vector_sneg_minus1: "-x = (-1::'a::ring_1) *s x" by vector
272 lemma vector_smult_rzero[simp]: "c *s 0 = (0::'a::mult_zero ^ 'n)" by vector
273 lemma vector_sub_rdistrib: "((a::'a::ring) - b) *s x = a *s x - b *s x"
274 by (vector field_simps)
276 lemma vec_eq[simp]: "(vec m = vec n) \<longleftrightarrow> (m = n)"
277 by (simp add: vec_eq_iff)
279 lemma norm_eq_0_imp: "norm x = 0 ==> x = (0::real ^'n)" by (metis norm_eq_zero)
280 lemma vector_mul_eq_0[simp]: "(a *s x = 0) \<longleftrightarrow> a = (0::'a::idom) \<or> x = 0"
282 lemma vector_mul_lcancel[simp]: "a *s x = a *s y \<longleftrightarrow> a = (0::real) \<or> x = y"
283 by (metis eq_iff_diff_eq_0 vector_mul_eq_0 vector_ssub_ldistrib)
284 lemma vector_mul_rcancel[simp]: "a *s x = b *s x \<longleftrightarrow> (a::real) = b \<or> x = 0"
285 by (metis eq_iff_diff_eq_0 vector_mul_eq_0 vector_sub_rdistrib)
286 lemma vector_mul_lcancel_imp: "a \<noteq> (0::real) ==> a *s x = a *s y ==> (x = y)"
287 by (metis vector_mul_lcancel)
288 lemma vector_mul_rcancel_imp: "x \<noteq> 0 \<Longrightarrow> (a::real) *s x = b *s x ==> a = b"
289 by (metis vector_mul_rcancel)
291 lemma component_le_norm_cart: "\<bar>x$i\<bar> <= norm x"
292 apply (simp add: norm_vec_def)
293 apply (rule member_le_setL2, simp_all)
296 lemma norm_bound_component_le_cart: "norm x <= e ==> \<bar>x$i\<bar> <= e"
297 by (metis component_le_norm_cart order_trans)
299 lemma norm_bound_component_lt_cart: "norm x < e ==> \<bar>x$i\<bar> < e"
300 by (metis component_le_norm_cart le_less_trans)
302 lemma norm_le_l1_cart: "norm x <= setsum(\<lambda>i. \<bar>x$i\<bar>) UNIV"
303 by (simp add: norm_vec_def setL2_le_setsum)
305 lemma scalar_mult_eq_scaleR: "c *s x = c *\<^sub>R x"
306 unfolding scaleR_vec_def vector_scalar_mult_def by simp
308 lemma dist_mul[simp]: "dist (c *s x) (c *s y) = \<bar>c\<bar> * dist x y"
309 unfolding dist_norm scalar_mult_eq_scaleR
310 unfolding scaleR_right_diff_distrib[symmetric] by simp
312 lemma setsum_component [simp]:
313 fixes f:: " 'a \<Rightarrow> ('b::comm_monoid_add) ^'n"
314 shows "(setsum f S)$i = setsum (\<lambda>x. (f x)$i) S"
315 proof (cases "finite S")
317 then show ?thesis by induct simp_all
320 then show ?thesis by simp
323 lemma setsum_eq: "setsum f S = (\<chi> i. setsum (\<lambda>x. (f x)$i ) S)"
324 by (simp add: vec_eq_iff)
327 fixes f:: "'c \<Rightarrow> ('a::semiring_1)^'n"
328 shows "setsum (\<lambda>x. c *s f x) S = c *s setsum f S"
329 by (simp add: vec_eq_iff setsum_right_distrib)
331 lemma setsum_norm_allsubsets_bound_cart:
332 fixes f:: "'a \<Rightarrow> real ^'n"
333 assumes fP: "finite P" and fPs: "\<And>Q. Q \<subseteq> P \<Longrightarrow> norm (setsum f Q) \<le> e"
334 shows "setsum (\<lambda>x. norm (f x)) P \<le> 2 * real CARD('n) * e"
335 using setsum_norm_allsubsets_bound[OF assms]
338 subsection\<open>Closures and interiors of halfspaces\<close>
340 lemma interior_halfspace_le [simp]:
341 assumes "a \<noteq> 0"
342 shows "interior {x. a \<bullet> x \<le> b} = {x. a \<bullet> x < b}"
344 have *: "a \<bullet> x < b" if x: "x \<in> S" and S: "S \<subseteq> {x. a \<bullet> x \<le> b}" and "open S" for S x
346 obtain e where "e>0" and e: "cball x e \<subseteq> S"
347 using \<open>open S\<close> open_contains_cball x by blast
348 then have "x + (e / norm a) *\<^sub>R a \<in> cball x e"
349 by (simp add: dist_norm)
350 then have "x + (e / norm a) *\<^sub>R a \<in> S"
352 then have "x + (e / norm a) *\<^sub>R a \<in> {x. a \<bullet> x \<le> b}"
354 moreover have "e * (a \<bullet> a) / norm a > 0"
355 by (simp add: \<open>0 < e\<close> assms)
356 ultimately show ?thesis
357 by (simp add: algebra_simps)
360 by (rule interior_unique) (auto simp: open_halfspace_lt *)
363 lemma interior_halfspace_ge [simp]:
364 "a \<noteq> 0 \<Longrightarrow> interior {x. a \<bullet> x \<ge> b} = {x. a \<bullet> x > b}"
365 using interior_halfspace_le [of "-a" "-b"] by simp
367 lemma interior_halfspace_component_le [simp]:
368 "interior {x. x$k \<le> a} = {x :: (real,'n::finite) vec. x$k < a}" (is "?LE")
369 and interior_halfspace_component_ge [simp]:
370 "interior {x. x$k \<ge> a} = {x :: (real,'n::finite) vec. x$k > a}" (is "?GE")
372 have "axis k (1::real) \<noteq> 0"
373 by (simp add: axis_def vec_eq_iff)
374 moreover have "axis k (1::real) \<bullet> x = x$k" for x
375 by (simp add: cart_eq_inner_axis inner_commute)
376 ultimately show ?LE ?GE
377 using interior_halfspace_le [of "axis k (1::real)" a]
378 interior_halfspace_ge [of "axis k (1::real)" a] by auto
381 lemma closure_halfspace_lt [simp]:
382 assumes "a \<noteq> 0"
383 shows "closure {x. a \<bullet> x < b} = {x. a \<bullet> x \<le> b}"
385 have [simp]: "-{x. a \<bullet> x < b} = {x. a \<bullet> x \<ge> b}"
388 using interior_halfspace_ge [of a b] assms
389 by (force simp: closure_interior)
392 lemma closure_halfspace_gt [simp]:
393 "a \<noteq> 0 \<Longrightarrow> closure {x. a \<bullet> x > b} = {x. a \<bullet> x \<ge> b}"
394 using closure_halfspace_lt [of "-a" "-b"] by simp
396 lemma closure_halfspace_component_lt [simp]:
397 "closure {x. x$k < a} = {x :: (real,'n::finite) vec. x$k \<le> a}" (is "?LE")
398 and closure_halfspace_component_gt [simp]:
399 "closure {x. x$k > a} = {x :: (real,'n::finite) vec. x$k \<ge> a}" (is "?GE")
401 have "axis k (1::real) \<noteq> 0"
402 by (simp add: axis_def vec_eq_iff)
403 moreover have "axis k (1::real) \<bullet> x = x$k" for x
404 by (simp add: cart_eq_inner_axis inner_commute)
405 ultimately show ?LE ?GE
406 using closure_halfspace_lt [of "axis k (1::real)" a]
407 closure_halfspace_gt [of "axis k (1::real)" a] by auto
410 lemma interior_hyperplane [simp]:
411 assumes "a \<noteq> 0"
412 shows "interior {x. a \<bullet> x = b} = {}"
414 have [simp]: "{x. a \<bullet> x = b} = {x. a \<bullet> x \<le> b} \<inter> {x. a \<bullet> x \<ge> b}"
417 by (auto simp: assms)
420 lemma frontier_halfspace_le:
421 assumes "a \<noteq> 0 \<or> b \<noteq> 0"
422 shows "frontier {x. a \<bullet> x \<le> b} = {x. a \<bullet> x = b}"
423 proof (cases "a = 0")
424 case True with assms show ?thesis by simp
426 case False then show ?thesis
427 by (force simp: frontier_def closed_halfspace_le)
430 lemma frontier_halfspace_ge:
431 assumes "a \<noteq> 0 \<or> b \<noteq> 0"
432 shows "frontier {x. a \<bullet> x \<ge> b} = {x. a \<bullet> x = b}"
433 proof (cases "a = 0")
434 case True with assms show ?thesis by simp
436 case False then show ?thesis
437 by (force simp: frontier_def closed_halfspace_ge)
440 lemma frontier_halfspace_lt:
441 assumes "a \<noteq> 0 \<or> b \<noteq> 0"
442 shows "frontier {x. a \<bullet> x < b} = {x. a \<bullet> x = b}"
443 proof (cases "a = 0")
444 case True with assms show ?thesis by simp
446 case False then show ?thesis
447 by (force simp: frontier_def interior_open open_halfspace_lt)
450 lemma frontier_halfspace_gt:
451 assumes "a \<noteq> 0 \<or> b \<noteq> 0"
452 shows "frontier {x. a \<bullet> x > b} = {x. a \<bullet> x = b}"
453 proof (cases "a = 0")
454 case True with assms show ?thesis by simp
456 case False then show ?thesis
457 by (force simp: frontier_def interior_open open_halfspace_gt)
460 lemma interior_standard_hyperplane:
461 "interior {x :: (real,'n::finite) vec. x$k = a} = {}"
463 have "axis k (1::real) \<noteq> 0"
464 by (simp add: axis_def vec_eq_iff)
465 moreover have "axis k (1::real) \<bullet> x = x$k" for x
466 by (simp add: cart_eq_inner_axis inner_commute)
467 ultimately show ?thesis
468 using interior_hyperplane [of "axis k (1::real)" a]
472 subsection \<open>Matrix operations\<close>
474 text\<open>Matrix notation. NB: an MxN matrix is of type @{typ "'a^'n^'m"}, not @{typ "'a^'m^'n"}\<close>
476 definition matrix_matrix_mult :: "('a::semiring_1) ^'n^'m \<Rightarrow> 'a ^'p^'n \<Rightarrow> 'a ^ 'p ^'m"
478 where "m ** m' == (\<chi> i j. setsum (\<lambda>k. ((m$i)$k) * ((m'$k)$j)) (UNIV :: 'n set)) ::'a ^ 'p ^'m"
480 definition matrix_vector_mult :: "('a::semiring_1) ^'n^'m \<Rightarrow> 'a ^'n \<Rightarrow> 'a ^ 'm"
482 where "m *v x \<equiv> (\<chi> i. setsum (\<lambda>j. ((m$i)$j) * (x$j)) (UNIV ::'n set)) :: 'a^'m"
484 definition vector_matrix_mult :: "'a ^ 'm \<Rightarrow> ('a::semiring_1) ^'n^'m \<Rightarrow> 'a ^'n "
486 where "v v* m == (\<chi> j. setsum (\<lambda>i. ((m$i)$j) * (v$i)) (UNIV :: 'm set)) :: 'a^'n"
488 definition "(mat::'a::zero => 'a ^'n^'n) k = (\<chi> i j. if i = j then k else 0)"
489 definition transpose where
490 "(transpose::'a^'n^'m \<Rightarrow> 'a^'m^'n) A = (\<chi> i j. ((A$j)$i))"
491 definition "(row::'m => 'a ^'n^'m \<Rightarrow> 'a ^'n) i A = (\<chi> j. ((A$i)$j))"
492 definition "(column::'n =>'a^'n^'m =>'a^'m) j A = (\<chi> i. ((A$i)$j))"
493 definition "rows(A::'a^'n^'m) = { row i A | i. i \<in> (UNIV :: 'm set)}"
494 definition "columns(A::'a^'n^'m) = { column i A | i. i \<in> (UNIV :: 'n set)}"
496 lemma mat_0[simp]: "mat 0 = 0" by (vector mat_def)
497 lemma matrix_add_ldistrib: "(A ** (B + C)) = (A ** B) + (A ** C)"
498 by (vector matrix_matrix_mult_def setsum.distrib[symmetric] field_simps)
500 lemma matrix_mul_lid:
501 fixes A :: "'a::semiring_1 ^ 'm ^ 'n"
502 shows "mat 1 ** A = A"
503 apply (simp add: matrix_matrix_mult_def mat_def)
505 apply (auto simp only: if_distrib cond_application_beta setsum.delta'[OF finite]
506 mult_1_left mult_zero_left if_True UNIV_I)
510 lemma matrix_mul_rid:
511 fixes A :: "'a::semiring_1 ^ 'm ^ 'n"
512 shows "A ** mat 1 = A"
513 apply (simp add: matrix_matrix_mult_def mat_def)
515 apply (auto simp only: if_distrib cond_application_beta setsum.delta[OF finite]
516 mult_1_right mult_zero_right if_True UNIV_I cong: if_cong)
519 lemma matrix_mul_assoc: "A ** (B ** C) = (A ** B) ** C"
520 apply (vector matrix_matrix_mult_def setsum_right_distrib setsum_left_distrib mult.assoc)
521 apply (subst setsum.commute)
525 lemma matrix_vector_mul_assoc: "A *v (B *v x) = (A ** B) *v x"
526 apply (vector matrix_matrix_mult_def matrix_vector_mult_def
527 setsum_right_distrib setsum_left_distrib mult.assoc)
528 apply (subst setsum.commute)
532 lemma matrix_vector_mul_lid: "mat 1 *v x = (x::'a::semiring_1 ^ 'n)"
533 apply (vector matrix_vector_mult_def mat_def)
534 apply (simp add: if_distrib cond_application_beta setsum.delta' cong del: if_weak_cong)
537 lemma matrix_transpose_mul:
538 "transpose(A ** B) = transpose B ** transpose (A::'a::comm_semiring_1^_^_)"
539 by (simp add: matrix_matrix_mult_def transpose_def vec_eq_iff mult.commute)
542 fixes A B :: "'a::semiring_1 ^ 'n ^ 'm"
543 shows "A = B \<longleftrightarrow> (\<forall>x. A *v x = B *v x)" (is "?lhs \<longleftrightarrow> ?rhs")
545 apply (subst vec_eq_iff)
547 apply (clarsimp simp add: matrix_vector_mult_def if_distrib cond_application_beta vec_eq_iff cong del: if_weak_cong)
548 apply (erule_tac x="axis ia 1" in allE)
549 apply (erule_tac x="i" in allE)
550 apply (auto simp add: if_distrib cond_application_beta axis_def
551 setsum.delta[OF finite] cong del: if_weak_cong)
554 lemma matrix_vector_mul_component: "((A::real^_^_) *v x)$k = (A$k) \<bullet> x"
555 by (simp add: matrix_vector_mult_def inner_vec_def)
557 lemma dot_lmul_matrix: "((x::real ^_) v* A) \<bullet> y = x \<bullet> (A *v y)"
558 apply (simp add: inner_vec_def matrix_vector_mult_def vector_matrix_mult_def setsum_left_distrib setsum_right_distrib ac_simps)
559 apply (subst setsum.commute)
563 lemma transpose_mat: "transpose (mat n) = mat n"
564 by (vector transpose_def mat_def)
566 lemma transpose_transpose: "transpose(transpose A) = A"
567 by (vector transpose_def)
570 fixes A:: "'a::semiring_1^_^_"
571 shows "row i (transpose A) = column i A"
572 by (simp add: row_def column_def transpose_def vec_eq_iff)
574 lemma column_transpose:
575 fixes A:: "'a::semiring_1^_^_"
576 shows "column i (transpose A) = row i A"
577 by (simp add: row_def column_def transpose_def vec_eq_iff)
579 lemma rows_transpose: "rows(transpose (A::'a::semiring_1^_^_)) = columns A"
580 by (auto simp add: rows_def columns_def row_transpose intro: set_eqI)
582 lemma columns_transpose: "columns(transpose (A::'a::semiring_1^_^_)) = rows A"
583 by (metis transpose_transpose rows_transpose)
585 text\<open>Two sometimes fruitful ways of looking at matrix-vector multiplication.\<close>
587 lemma matrix_mult_dot: "A *v x = (\<chi> i. A$i \<bullet> x)"
588 by (simp add: matrix_vector_mult_def inner_vec_def)
590 lemma matrix_mult_vsum:
591 "(A::'a::comm_semiring_1^'n^'m) *v x = setsum (\<lambda>i. (x$i) *s column i A) (UNIV:: 'n set)"
592 by (simp add: matrix_vector_mult_def vec_eq_iff column_def mult.commute)
594 lemma vector_componentwise:
595 "(x::'a::ring_1^'n) = (\<chi> j. \<Sum>i\<in>UNIV. (x$i) * (axis i 1 :: 'a^'n) $ j)"
596 by (simp add: axis_def if_distrib setsum.If_cases vec_eq_iff)
598 lemma basis_expansion: "setsum (\<lambda>i. (x$i) *s axis i 1) UNIV = (x::('a::ring_1) ^'n)"
599 by (auto simp add: axis_def vec_eq_iff if_distrib setsum.If_cases cong del: if_weak_cong)
601 lemma linear_componentwise:
602 fixes f:: "real ^'m \<Rightarrow> real ^ _"
603 assumes lf: "linear f"
604 shows "(f x)$j = setsum (\<lambda>i. (x$i) * (f (axis i 1)$j)) (UNIV :: 'm set)" (is "?lhs = ?rhs")
606 let ?M = "(UNIV :: 'm set)"
607 let ?N = "(UNIV :: 'n set)"
608 have "?rhs = (setsum (\<lambda>i.(x$i) *\<^sub>R f (axis i 1) ) ?M)$j"
609 unfolding setsum_component by simp
611 unfolding linear_setsum_mul[OF lf, symmetric]
612 unfolding scalar_mult_eq_scaleR[symmetric]
613 unfolding basis_expansion
617 text\<open>Inverse matrices (not necessarily square)\<close>
620 "invertible(A::'a::semiring_1^'n^'m) \<longleftrightarrow> (\<exists>A'::'a^'m^'n. A ** A' = mat 1 \<and> A' ** A = mat 1)"
623 "matrix_inv(A:: 'a::semiring_1^'n^'m) =
624 (SOME A'::'a^'m^'n. A ** A' = mat 1 \<and> A' ** A = mat 1)"
626 text\<open>Correspondence between matrices and linear operators.\<close>
628 definition matrix :: "('a::{plus,times, one, zero}^'m \<Rightarrow> 'a ^ 'n) \<Rightarrow> 'a^'m^'n"
629 where "matrix f = (\<chi> i j. (f(axis j 1))$i)"
631 lemma matrix_vector_mul_linear: "linear(\<lambda>x. A *v (x::real ^ _))"
632 by (simp add: linear_iff matrix_vector_mult_def vec_eq_iff
633 field_simps setsum_right_distrib setsum.distrib)
636 assumes lf: "linear f"
637 shows "matrix f *v x = f (x::real ^ 'n)"
638 apply (simp add: matrix_def matrix_vector_mult_def vec_eq_iff mult.commute)
640 apply (rule linear_componentwise[OF lf, symmetric])
643 lemma matrix_vector_mul: "linear f ==> f = (\<lambda>x. matrix f *v (x::real ^ 'n))"
644 by (simp add: ext matrix_works)
646 lemma matrix_of_matrix_vector_mul: "matrix(\<lambda>x. A *v (x :: real ^ 'n)) = A"
647 by (simp add: matrix_eq matrix_vector_mul_linear matrix_works)
649 lemma matrix_compose:
650 assumes lf: "linear (f::real^'n \<Rightarrow> real^'m)"
651 and lg: "linear (g::real^'m \<Rightarrow> real^_)"
652 shows "matrix (g \<circ> f) = matrix g ** matrix f"
653 using lf lg linear_compose[OF lf lg] matrix_works[OF linear_compose[OF lf lg]]
654 by (simp add: matrix_eq matrix_works matrix_vector_mul_assoc[symmetric] o_def)
656 lemma matrix_vector_column:
657 "(A::'a::comm_semiring_1^'n^_) *v x = setsum (\<lambda>i. (x$i) *s ((transpose A)$i)) (UNIV:: 'n set)"
658 by (simp add: matrix_vector_mult_def transpose_def vec_eq_iff mult.commute)
660 lemma adjoint_matrix: "adjoint(\<lambda>x. (A::real^'n^'m) *v x) = (\<lambda>x. transpose A *v x)"
661 apply (rule adjoint_unique)
662 apply (simp add: transpose_def inner_vec_def matrix_vector_mult_def
663 setsum_left_distrib setsum_right_distrib)
664 apply (subst setsum.commute)
665 apply (auto simp add: ac_simps)
668 lemma matrix_adjoint: assumes lf: "linear (f :: real^'n \<Rightarrow> real ^'m)"
669 shows "matrix(adjoint f) = transpose(matrix f)"
670 apply (subst matrix_vector_mul[OF lf])
671 unfolding adjoint_matrix matrix_of_matrix_vector_mul
676 subsection \<open>lambda skolemization on cartesian products\<close>
678 (* FIXME: rename do choice_cart *)
680 lemma lambda_skolem: "(\<forall>i. \<exists>x. P i x) \<longleftrightarrow>
681 (\<exists>x::'a ^ 'n. \<forall>i. P i (x $ i))" (is "?lhs \<longleftrightarrow> ?rhs")
683 let ?S = "(UNIV :: 'n set)"
685 then have ?lhs by auto }
688 then obtain f where f:"\<forall>i. P i (f i)" unfolding choice_iff by metis
689 let ?x = "(\<chi> i. (f i)) :: 'a ^ 'n"
691 from f have "P i (f i)" by metis
692 then have "P i (?x $ i)" by auto
694 hence "\<forall>i. P i (?x$i)" by metis
695 hence ?rhs by metis }
696 ultimately show ?thesis by metis
699 lemma vector_sub_project_orthogonal_cart: "(b::real^'n) \<bullet> (x - ((b \<bullet> x) / (b \<bullet> b)) *s b) = 0"
700 unfolding inner_simps scalar_mult_eq_scaleR by auto
702 lemma left_invertible_transpose:
703 "(\<exists>(B). B ** transpose (A) = mat (1::'a::comm_semiring_1)) \<longleftrightarrow> (\<exists>(B). A ** B = mat 1)"
704 by (metis matrix_transpose_mul transpose_mat transpose_transpose)
706 lemma right_invertible_transpose:
707 "(\<exists>(B). transpose (A) ** B = mat (1::'a::comm_semiring_1)) \<longleftrightarrow> (\<exists>(B). B ** A = mat 1)"
708 by (metis matrix_transpose_mul transpose_mat transpose_transpose)
710 lemma matrix_left_invertible_injective:
711 "(\<exists>B. (B::real^'m^'n) ** (A::real^'n^'m) = mat 1) \<longleftrightarrow> (\<forall>x y. A *v x = A *v y \<longrightarrow> x = y)"
713 { fix B:: "real^'m^'n" and x y assume B: "B ** A = mat 1" and xy: "A *v x = A*v y"
714 from xy have "B*v (A *v x) = B *v (A*v y)" by simp
716 unfolding matrix_vector_mul_assoc B matrix_vector_mul_lid . }
718 { assume A: "\<forall>x y. A *v x = A *v y \<longrightarrow> x = y"
719 hence i: "inj (op *v A)" unfolding inj_on_def by auto
720 from linear_injective_left_inverse[OF matrix_vector_mul_linear i]
721 obtain g where g: "linear g" "g \<circ> op *v A = id" by blast
722 have "matrix g ** A = mat 1"
723 unfolding matrix_eq matrix_vector_mul_lid matrix_vector_mul_assoc[symmetric] matrix_works[OF g(1)]
724 using g(2) by (simp add: fun_eq_iff)
725 then have "\<exists>B. (B::real ^'m^'n) ** A = mat 1" by blast }
726 ultimately show ?thesis by blast
729 lemma matrix_left_invertible_ker:
730 "(\<exists>B. (B::real ^'m^'n) ** (A::real^'n^'m) = mat 1) \<longleftrightarrow> (\<forall>x. A *v x = 0 \<longrightarrow> x = 0)"
731 unfolding matrix_left_invertible_injective
732 using linear_injective_0[OF matrix_vector_mul_linear, of A]
733 by (simp add: inj_on_def)
735 lemma matrix_right_invertible_surjective:
736 "(\<exists>B. (A::real^'n^'m) ** (B::real^'m^'n) = mat 1) \<longleftrightarrow> surj (\<lambda>x. A *v x)"
738 { fix B :: "real ^'m^'n"
739 assume AB: "A ** B = mat 1"
740 { fix x :: "real ^ 'm"
741 have "A *v (B *v x) = x"
742 by (simp add: matrix_vector_mul_lid matrix_vector_mul_assoc AB) }
743 hence "surj (op *v A)" unfolding surj_def by metis }
745 { assume sf: "surj (op *v A)"
746 from linear_surjective_right_inverse[OF matrix_vector_mul_linear sf]
747 obtain g:: "real ^'m \<Rightarrow> real ^'n" where g: "linear g" "op *v A \<circ> g = id"
750 have "A ** (matrix g) = mat 1"
751 unfolding matrix_eq matrix_vector_mul_lid
752 matrix_vector_mul_assoc[symmetric] matrix_works[OF g(1)]
753 using g(2) unfolding o_def fun_eq_iff id_def
755 hence "\<exists>B. A ** (B::real^'m^'n) = mat 1" by blast
757 ultimately show ?thesis unfolding surj_def by blast
760 lemma matrix_left_invertible_independent_columns:
761 fixes A :: "real^'n^'m"
762 shows "(\<exists>(B::real ^'m^'n). B ** A = mat 1) \<longleftrightarrow>
763 (\<forall>c. setsum (\<lambda>i. c i *s column i A) (UNIV :: 'n set) = 0 \<longrightarrow> (\<forall>i. c i = 0))"
764 (is "?lhs \<longleftrightarrow> ?rhs")
766 let ?U = "UNIV :: 'n set"
767 { assume k: "\<forall>x. A *v x = 0 \<longrightarrow> x = 0"
769 assume c: "setsum (\<lambda>i. c i *s column i A) ?U = 0" and i: "i \<in> ?U"
770 let ?x = "\<chi> i. c i"
771 have th0:"A *v ?x = 0"
773 unfolding matrix_mult_vsum vec_eq_iff
775 from k[rule_format, OF th0] i
776 have "c i = 0" by (vector vec_eq_iff)}
777 hence ?rhs by blast }
780 { fix x assume x: "A *v x = 0"
781 let ?c = "\<lambda>i. ((x$i ):: real)"
782 from H[rule_format, of ?c, unfolded matrix_mult_vsum[symmetric], OF x]
783 have "x = 0" by vector }
785 ultimately show ?thesis unfolding matrix_left_invertible_ker by blast
788 lemma matrix_right_invertible_independent_rows:
789 fixes A :: "real^'n^'m"
790 shows "(\<exists>(B::real^'m^'n). A ** B = mat 1) \<longleftrightarrow>
791 (\<forall>c. setsum (\<lambda>i. c i *s row i A) (UNIV :: 'm set) = 0 \<longrightarrow> (\<forall>i. c i = 0))"
792 unfolding left_invertible_transpose[symmetric]
793 matrix_left_invertible_independent_columns
794 by (simp add: column_transpose)
796 lemma matrix_right_invertible_span_columns:
797 "(\<exists>(B::real ^'n^'m). (A::real ^'m^'n) ** B = mat 1) \<longleftrightarrow>
798 span (columns A) = UNIV" (is "?lhs = ?rhs")
800 let ?U = "UNIV :: 'm set"
801 have fU: "finite ?U" by simp
802 have lhseq: "?lhs \<longleftrightarrow> (\<forall>y. \<exists>(x::real^'m). setsum (\<lambda>i. (x$i) *s column i A) ?U = y)"
803 unfolding matrix_right_invertible_surjective matrix_mult_vsum surj_def
804 apply (subst eq_commute)
807 have rhseq: "?rhs \<longleftrightarrow> (\<forall>x. x \<in> span (columns A))" by blast
810 from h[unfolded lhseq, rule_format, of x] obtain y :: "real ^'m"
811 where y: "setsum (\<lambda>i. (y$i) *s column i A) ?U = x" by blast
812 have "x \<in> span (columns A)"
813 unfolding y[symmetric]
814 apply (rule span_setsum)
815 unfolding scalar_mult_eq_scaleR
816 apply (rule span_mul)
817 apply (rule span_superset)
818 unfolding columns_def
822 then have ?rhs unfolding rhseq by blast }
825 let ?P = "\<lambda>(y::real ^'n). \<exists>(x::real^'m). setsum (\<lambda>i. (x$i) *s column i A) ?U = y"
828 proof (rule span_induct_alt[of ?P "columns A", folded scalar_mult_eq_scaleR])
829 show "\<exists>x::real ^ 'm. setsum (\<lambda>i. (x$i) *s column i A) ?U = 0"
830 by (rule exI[where x=0], simp)
833 assume y1: "y1 \<in> columns A" and y2: "?P y2"
834 from y1 obtain i where i: "i \<in> ?U" "y1 = column i A"
835 unfolding columns_def by blast
836 from y2 obtain x:: "real ^'m" where
837 x: "setsum (\<lambda>i. (x$i) *s column i A) ?U = y2" by blast
838 let ?x = "(\<chi> j. if j = i then c + (x$i) else (x$j))::real^'m"
839 show "?P (c*s y1 + y2)"
840 proof (rule exI[where x= "?x"], vector, auto simp add: i x[symmetric] if_distrib distrib_left cond_application_beta cong del: if_weak_cong)
842 have th: "\<forall>xa \<in> ?U. (if xa = i then (c + (x$i)) * ((column xa A)$j)
843 else (x$xa) * ((column xa A$j))) = (if xa = i then c * ((column i A)$j) else 0) + ((x$xa) * ((column xa A)$j))"
844 using i(1) by (simp add: field_simps)
845 have "setsum (\<lambda>xa. if xa = i then (c + (x$i)) * ((column xa A)$j)
846 else (x$xa) * ((column xa A$j))) ?U = setsum (\<lambda>xa. (if xa = i then c * ((column i A)$j) else 0) + ((x$xa) * ((column xa A)$j))) ?U"
847 apply (rule setsum.cong[OF refl])
850 also have "\<dots> = setsum (\<lambda>xa. if xa = i then c * ((column i A)$j) else 0) ?U + setsum (\<lambda>xa. ((x$xa) * ((column xa A)$j))) ?U"
851 by (simp add: setsum.distrib)
852 also have "\<dots> = c * ((column i A)$j) + setsum (\<lambda>xa. ((x$xa) * ((column xa A)$j))) ?U"
853 unfolding setsum.delta[OF fU]
855 finally show "setsum (\<lambda>xa. if xa = i then (c + (x$i)) * ((column xa A)$j)
856 else (x$xa) * ((column xa A$j))) ?U = c * ((column i A)$j) + setsum (\<lambda>xa. ((x$xa) * ((column xa A)$j))) ?U" .
859 show "y \<in> span (columns A)"
863 then have ?lhs unfolding lhseq ..
865 ultimately show ?thesis by blast
868 lemma matrix_left_invertible_span_rows:
869 "(\<exists>(B::real^'m^'n). B ** (A::real^'n^'m) = mat 1) \<longleftrightarrow> span (rows A) = UNIV"
870 unfolding right_invertible_transpose[symmetric]
871 unfolding columns_transpose[symmetric]
872 unfolding matrix_right_invertible_span_columns
875 text \<open>The same result in terms of square matrices.\<close>
877 lemma matrix_left_right_inverse:
878 fixes A A' :: "real ^'n^'n"
879 shows "A ** A' = mat 1 \<longleftrightarrow> A' ** A = mat 1"
881 { fix A A' :: "real ^'n^'n"
882 assume AA': "A ** A' = mat 1"
883 have sA: "surj (op *v A)"
886 apply (rule_tac x="(A' *v y)" in exI)
887 apply (simp add: matrix_vector_mul_assoc AA' matrix_vector_mul_lid)
889 from linear_surjective_isomorphism[OF matrix_vector_mul_linear sA]
890 obtain f' :: "real ^'n \<Rightarrow> real ^'n"
891 where f': "linear f'" "\<forall>x. f' (A *v x) = x" "\<forall>x. A *v f' x = x" by blast
892 have th: "matrix f' ** A = mat 1"
893 by (simp add: matrix_eq matrix_works[OF f'(1)]
894 matrix_vector_mul_assoc[symmetric] matrix_vector_mul_lid f'(2)[rule_format])
895 hence "(matrix f' ** A) ** A' = mat 1 ** A'" by simp
896 hence "matrix f' = A'"
897 by (simp add: matrix_mul_assoc[symmetric] AA' matrix_mul_rid matrix_mul_lid)
898 hence "matrix f' ** A = A' ** A" by simp
899 hence "A' ** A = mat 1" by (simp add: th)
901 then show ?thesis by blast
904 text \<open>Considering an n-element vector as an n-by-1 or 1-by-n matrix.\<close>
906 definition "rowvector v = (\<chi> i j. (v$j))"
908 definition "columnvector v = (\<chi> i j. (v$i))"
910 lemma transpose_columnvector: "transpose(columnvector v) = rowvector v"
911 by (simp add: transpose_def rowvector_def columnvector_def vec_eq_iff)
913 lemma transpose_rowvector: "transpose(rowvector v) = columnvector v"
914 by (simp add: transpose_def columnvector_def rowvector_def vec_eq_iff)
916 lemma dot_rowvector_columnvector: "columnvector (A *v v) = A ** columnvector v"
917 by (vector columnvector_def matrix_matrix_mult_def matrix_vector_mult_def)
919 lemma dot_matrix_product:
920 "(x::real^'n) \<bullet> y = (((rowvector x ::real^'n^1) ** (columnvector y :: real^1^'n))$1)$1"
921 by (vector matrix_matrix_mult_def rowvector_def columnvector_def inner_vec_def)
923 lemma dot_matrix_vector_mul:
924 fixes A B :: "real ^'n ^'n" and x y :: "real ^'n"
925 shows "(A *v x) \<bullet> (B *v y) =
926 (((rowvector x :: real^'n^1) ** ((transpose A ** B) ** (columnvector y :: real ^1^'n)))$1)$1"
927 unfolding dot_matrix_product transpose_columnvector[symmetric]
928 dot_rowvector_columnvector matrix_transpose_mul matrix_mul_assoc ..
931 lemma infnorm_cart:"infnorm (x::real^'n) = Sup {\<bar>x$i\<bar> |i. i\<in>UNIV}"
932 by (simp add: infnorm_def inner_axis Basis_vec_def) (metis (lifting) inner_axis real_inner_1_right)
934 lemma component_le_infnorm_cart: "\<bar>x$i\<bar> \<le> infnorm (x::real^'n)"
935 using Basis_le_infnorm[of "axis i 1" x]
936 by (simp add: Basis_vec_def axis_eq_axis inner_axis)
938 lemma continuous_component: "continuous F f \<Longrightarrow> continuous F (\<lambda>x. f x $ i)"
939 unfolding continuous_def by (rule tendsto_vec_nth)
941 lemma continuous_on_component: "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. f x $ i)"
942 unfolding continuous_on_def by (fast intro: tendsto_vec_nth)
944 lemma closed_positive_orthant: "closed {x::real^'n. \<forall>i. 0 \<le>x$i}"
945 by (simp add: Collect_all_eq closed_INT closed_Collect_le continuous_on_const continuous_on_id continuous_on_component)
947 lemma bounded_component_cart: "bounded s \<Longrightarrow> bounded ((\<lambda>x. x $ i) ` s)"
948 unfolding bounded_def
950 apply (rule_tac x="x $ i" in exI)
951 apply (rule_tac x="e" in exI)
953 apply (rule order_trans [OF dist_vec_nth_le], simp)
956 lemma compact_lemma_cart:
957 fixes f :: "nat \<Rightarrow> 'a::heine_borel ^ 'n"
958 assumes f: "bounded (range f)"
959 shows "\<exists>l r. subseq r \<and>
960 (\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r n) $ i) (l $ i) < e) sequentially)"
963 have "\<forall>d' \<subseteq> d. ?th d'"
964 by (rule compact_lemma_general[where unproj=vec_lambda])
965 (auto intro!: f bounded_component_cart simp: vec_lambda_eta)
966 then show "?th d" by simp
969 instance vec :: (heine_borel, finite) heine_borel
971 fix f :: "nat \<Rightarrow> 'a ^ 'b"
972 assume f: "bounded (range f)"
973 then obtain l r where r: "subseq r"
974 and l: "\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>UNIV. dist (f (r n) $ i) (l $ i) < e) sequentially"
975 using compact_lemma_cart [OF f] by blast
976 let ?d = "UNIV::'b set"
977 { fix e::real assume "e>0"
978 hence "0 < e / (real_of_nat (card ?d))"
979 using zero_less_card_finite divide_pos_pos[of e, of "real_of_nat (card ?d)"] by auto
980 with l have "eventually (\<lambda>n. \<forall>i. dist (f (r n) $ i) (l $ i) < e / (real_of_nat (card ?d))) sequentially"
984 assume n: "\<forall>i. dist (f (r n) $ i) (l $ i) < e / (real_of_nat (card ?d))"
985 have "dist (f (r n)) l \<le> (\<Sum>i\<in>?d. dist (f (r n) $ i) (l $ i))"
986 unfolding dist_vec_def using zero_le_dist by (rule setL2_le_setsum)
987 also have "\<dots> < (\<Sum>i\<in>?d. e / (real_of_nat (card ?d)))"
988 by (rule setsum_strict_mono) (simp_all add: n)
989 finally have "dist (f (r n)) l < e" by simp
991 ultimately have "eventually (\<lambda>n. dist (f (r n)) l < e) sequentially"
992 by (rule eventually_mono)
994 hence "((f \<circ> r) \<longlongrightarrow> l) sequentially" unfolding o_def tendsto_iff by simp
995 with r show "\<exists>l r. subseq r \<and> ((f \<circ> r) \<longlongrightarrow> l) sequentially" by auto
1000 shows "box a b = {x::real^'n. \<forall>i. a$i < x$i \<and> x$i < b$i}"
1001 and "cbox a b = {x::real^'n. \<forall>i. a$i \<le> x$i \<and> x$i \<le> b$i}"
1002 by (auto simp add: set_eq_iff less_vec_def less_eq_vec_def mem_box Basis_vec_def inner_axis)
1004 lemma mem_interval_cart:
1005 fixes a :: "real^'n"
1006 shows "x \<in> box a b \<longleftrightarrow> (\<forall>i. a$i < x$i \<and> x$i < b$i)"
1007 and "x \<in> cbox a b \<longleftrightarrow> (\<forall>i. a$i \<le> x$i \<and> x$i \<le> b$i)"
1008 using interval_cart[of a b] by (auto simp add: set_eq_iff less_vec_def less_eq_vec_def)
1010 lemma interval_eq_empty_cart:
1011 fixes a :: "real^'n"
1012 shows "(box a b = {} \<longleftrightarrow> (\<exists>i. b$i \<le> a$i))" (is ?th1)
1013 and "(cbox a b = {} \<longleftrightarrow> (\<exists>i. b$i < a$i))" (is ?th2)
1015 { fix i x assume as:"b$i \<le> a$i" and x:"x\<in>box a b"
1016 hence "a $ i < x $ i \<and> x $ i < b $ i" unfolding mem_interval_cart by auto
1017 hence "a$i < b$i" by auto
1018 hence False using as by auto }
1020 { assume as:"\<forall>i. \<not> (b$i \<le> a$i)"
1021 let ?x = "(1/2) *\<^sub>R (a + b)"
1023 have "a$i < b$i" using as[THEN spec[where x=i]] by auto
1024 hence "a$i < ((1/2) *\<^sub>R (a+b)) $ i" "((1/2) *\<^sub>R (a+b)) $ i < b$i"
1025 unfolding vector_smult_component and vector_add_component
1027 hence "box a b \<noteq> {}" using mem_interval_cart(1)[of "?x" a b] by auto }
1028 ultimately show ?th1 by blast
1030 { fix i x assume as:"b$i < a$i" and x:"x\<in>cbox a b"
1031 hence "a $ i \<le> x $ i \<and> x $ i \<le> b $ i" unfolding mem_interval_cart by auto
1032 hence "a$i \<le> b$i" by auto
1033 hence False using as by auto }
1035 { assume as:"\<forall>i. \<not> (b$i < a$i)"
1036 let ?x = "(1/2) *\<^sub>R (a + b)"
1038 have "a$i \<le> b$i" using as[THEN spec[where x=i]] by auto
1039 hence "a$i \<le> ((1/2) *\<^sub>R (a+b)) $ i" "((1/2) *\<^sub>R (a+b)) $ i \<le> b$i"
1040 unfolding vector_smult_component and vector_add_component
1042 hence "cbox a b \<noteq> {}" using mem_interval_cart(2)[of "?x" a b] by auto }
1043 ultimately show ?th2 by blast
1046 lemma interval_ne_empty_cart:
1047 fixes a :: "real^'n"
1048 shows "cbox a b \<noteq> {} \<longleftrightarrow> (\<forall>i. a$i \<le> b$i)"
1049 and "box a b \<noteq> {} \<longleftrightarrow> (\<forall>i. a$i < b$i)"
1050 unfolding interval_eq_empty_cart[of a b] by (auto simp add: not_less not_le)
1051 (* BH: Why doesn't just "auto" work here? *)
1053 lemma subset_interval_imp_cart:
1054 fixes a :: "real^'n"
1055 shows "(\<forall>i. a$i \<le> c$i \<and> d$i \<le> b$i) \<Longrightarrow> cbox c d \<subseteq> cbox a b"
1056 and "(\<forall>i. a$i < c$i \<and> d$i < b$i) \<Longrightarrow> cbox c d \<subseteq> box a b"
1057 and "(\<forall>i. a$i \<le> c$i \<and> d$i \<le> b$i) \<Longrightarrow> box c d \<subseteq> cbox a b"
1058 and "(\<forall>i. a$i \<le> c$i \<and> d$i \<le> b$i) \<Longrightarrow> box c d \<subseteq> box a b"
1059 unfolding subset_eq[unfolded Ball_def] unfolding mem_interval_cart
1060 by (auto intro: order_trans less_le_trans le_less_trans less_imp_le) (* BH: Why doesn't just "auto" work here? *)
1062 lemma interval_sing:
1063 fixes a :: "'a::linorder^'n"
1064 shows "{a .. a} = {a} \<and> {a<..<a} = {}"
1065 apply (auto simp add: set_eq_iff less_vec_def less_eq_vec_def vec_eq_iff)
1068 lemma subset_interval_cart:
1069 fixes a :: "real^'n"
1070 shows "cbox c d \<subseteq> cbox a b \<longleftrightarrow> (\<forall>i. c$i \<le> d$i) --> (\<forall>i. a$i \<le> c$i \<and> d$i \<le> b$i)" (is ?th1)
1071 and "cbox c d \<subseteq> box a b \<longleftrightarrow> (\<forall>i. c$i \<le> d$i) --> (\<forall>i. a$i < c$i \<and> d$i < b$i)" (is ?th2)
1072 and "box c d \<subseteq> cbox a b \<longleftrightarrow> (\<forall>i. c$i < d$i) --> (\<forall>i. a$i \<le> c$i \<and> d$i \<le> b$i)" (is ?th3)
1073 and "box c d \<subseteq> box a b \<longleftrightarrow> (\<forall>i. c$i < d$i) --> (\<forall>i. a$i \<le> c$i \<and> d$i \<le> b$i)" (is ?th4)
1074 using subset_box[of c d a b] by (simp_all add: Basis_vec_def inner_axis)
1076 lemma disjoint_interval_cart:
1078 shows "cbox a b \<inter> cbox c d = {} \<longleftrightarrow> (\<exists>i. (b$i < a$i \<or> d$i < c$i \<or> b$i < c$i \<or> d$i < a$i))" (is ?th1)
1079 and "cbox a b \<inter> box c d = {} \<longleftrightarrow> (\<exists>i. (b$i < a$i \<or> d$i \<le> c$i \<or> b$i \<le> c$i \<or> d$i \<le> a$i))" (is ?th2)
1080 and "box a b \<inter> cbox c d = {} \<longleftrightarrow> (\<exists>i. (b$i \<le> a$i \<or> d$i < c$i \<or> b$i \<le> c$i \<or> d$i \<le> a$i))" (is ?th3)
1081 and "box a b \<inter> box c d = {} \<longleftrightarrow> (\<exists>i. (b$i \<le> a$i \<or> d$i \<le> c$i \<or> b$i \<le> c$i \<or> d$i \<le> a$i))" (is ?th4)
1082 using disjoint_interval[of a b c d] by (simp_all add: Basis_vec_def inner_axis)
1084 lemma inter_interval_cart:
1085 fixes a :: "real^'n"
1086 shows "cbox a b \<inter> cbox c d = {(\<chi> i. max (a$i) (c$i)) .. (\<chi> i. min (b$i) (d$i))}"
1087 unfolding inter_interval
1088 by (auto simp: mem_box less_eq_vec_def)
1089 (auto simp: Basis_vec_def inner_axis)
1091 lemma closed_interval_left_cart:
1092 fixes b :: "real^'n"
1093 shows "closed {x::real^'n. \<forall>i. x$i \<le> b$i}"
1094 by (simp add: Collect_all_eq closed_INT closed_Collect_le continuous_on_const continuous_on_id continuous_on_component)
1096 lemma closed_interval_right_cart:
1098 shows "closed {x::real^'n. \<forall>i. a$i \<le> x$i}"
1099 by (simp add: Collect_all_eq closed_INT closed_Collect_le continuous_on_const continuous_on_id continuous_on_component)
1101 lemma is_interval_cart:
1102 "is_interval (s::(real^'n) set) \<longleftrightarrow>
1103 (\<forall>a\<in>s. \<forall>b\<in>s. \<forall>x. (\<forall>i. ((a$i \<le> x$i \<and> x$i \<le> b$i) \<or> (b$i \<le> x$i \<and> x$i \<le> a$i))) \<longrightarrow> x \<in> s)"
1104 by (simp add: is_interval_def Ball_def Basis_vec_def inner_axis imp_ex)
1106 lemma closed_halfspace_component_le_cart: "closed {x::real^'n. x$i \<le> a}"
1107 by (simp add: closed_Collect_le continuous_on_const continuous_on_id continuous_on_component)
1109 lemma closed_halfspace_component_ge_cart: "closed {x::real^'n. x$i \<ge> a}"
1110 by (simp add: closed_Collect_le continuous_on_const continuous_on_id continuous_on_component)
1112 lemma open_halfspace_component_lt_cart: "open {x::real^'n. x$i < a}"
1113 by (simp add: open_Collect_less continuous_on_const continuous_on_id continuous_on_component)
1115 lemma open_halfspace_component_gt_cart: "open {x::real^'n. x$i > a}"
1116 by (simp add: open_Collect_less continuous_on_const continuous_on_id continuous_on_component)
1118 lemma Lim_component_le_cart:
1119 fixes f :: "'a \<Rightarrow> real^'n"
1120 assumes "(f \<longlongrightarrow> l) net" "\<not> (trivial_limit net)" "eventually (\<lambda>x. f x $i \<le> b) net"
1122 by (rule tendsto_le[OF assms(2) tendsto_const tendsto_vec_nth, OF assms(1, 3)])
1124 lemma Lim_component_ge_cart:
1125 fixes f :: "'a \<Rightarrow> real^'n"
1126 assumes "(f \<longlongrightarrow> l) net" "\<not> (trivial_limit net)" "eventually (\<lambda>x. b \<le> (f x)$i) net"
1128 by (rule tendsto_le[OF assms(2) tendsto_vec_nth tendsto_const, OF assms(1, 3)])
1130 lemma Lim_component_eq_cart:
1131 fixes f :: "'a \<Rightarrow> real^'n"
1132 assumes net: "(f \<longlongrightarrow> l) net" "~(trivial_limit net)" and ev:"eventually (\<lambda>x. f(x)$i = b) net"
1134 using ev[unfolded order_eq_iff eventually_conj_iff] and
1135 Lim_component_ge_cart[OF net, of b i] and
1136 Lim_component_le_cart[OF net, of i b] by auto
1138 lemma connected_ivt_component_cart:
1139 fixes x :: "real^'n"
1140 shows "connected s \<Longrightarrow> x \<in> s \<Longrightarrow> y \<in> s \<Longrightarrow> x$k \<le> a \<Longrightarrow> a \<le> y$k \<Longrightarrow> (\<exists>z\<in>s. z$k = a)"
1141 using connected_ivt_hyperplane[of s x y "axis k 1" a]
1142 by (auto simp add: inner_axis inner_commute)
1144 lemma subspace_substandard_cart: "subspace {x::real^_. (\<forall>i. P i \<longrightarrow> x$i = 0)}"
1145 unfolding subspace_def by auto
1147 lemma closed_substandard_cart:
1148 "closed {x::'a::real_normed_vector ^ 'n. \<forall>i. P i \<longrightarrow> x$i = 0}"
1151 have "closed {x::'a ^ 'n. P i \<longrightarrow> x$i = 0}"
1152 by (cases "P i") (simp_all add: closed_Collect_eq continuous_on_const continuous_on_id continuous_on_component) }
1154 unfolding Collect_all_eq by (simp add: closed_INT)
1157 lemma dim_substandard_cart: "dim {x::real^'n. \<forall>i. i \<notin> d \<longrightarrow> x$i = 0} = card d"
1160 let ?a = "\<lambda>x. axis x 1 :: real^'n"
1161 have *: "{x. \<forall>i\<in>Basis. i \<notin> ?a ` d \<longrightarrow> x \<bullet> i = 0} = ?A"
1162 by (auto simp: image_iff Basis_vec_def axis_eq_axis inner_axis)
1163 have "?a ` d \<subseteq> Basis"
1164 by (auto simp: Basis_vec_def)
1166 using dim_substandard[of "?a ` d"] card_image[of ?a d]
1167 by (auto simp: axis_eq_axis inj_on_def *)
1170 lemma affinity_inverses:
1171 assumes m0: "m \<noteq> (0::'a::field)"
1172 shows "(\<lambda>x. m *s x + c) \<circ> (\<lambda>x. inverse(m) *s x + (-(inverse(m) *s c))) = id"
1173 "(\<lambda>x. inverse(m) *s x + (-(inverse(m) *s c))) \<circ> (\<lambda>x. m *s x + c) = id"
1175 apply (auto simp add: fun_eq_iff vector_add_ldistrib diff_conv_add_uminus simp del: add_uminus_conv_diff)
1176 apply (simp_all add: vector_smult_lneg[symmetric] vector_smult_assoc vector_sneg_minus1 [symmetric])
1179 lemma vector_affinity_eq:
1180 assumes m0: "(m::'a::field) \<noteq> 0"
1181 shows "m *s x + c = y \<longleftrightarrow> x = inverse m *s y + -(inverse m *s c)"
1183 assume h: "m *s x + c = y"
1184 hence "m *s x = y - c" by (simp add: field_simps)
1185 hence "inverse m *s (m *s x) = inverse m *s (y - c)" by simp
1186 then show "x = inverse m *s y + - (inverse m *s c)"
1187 using m0 by (simp add: vector_smult_assoc vector_ssub_ldistrib)
1189 assume h: "x = inverse m *s y + - (inverse m *s c)"
1190 show "m *s x + c = y" unfolding h
1191 using m0 by (simp add: vector_smult_assoc vector_ssub_ldistrib)
1194 lemma vector_eq_affinity:
1195 "(m::'a::field) \<noteq> 0 ==> (y = m *s x + c \<longleftrightarrow> inverse(m) *s y + -(inverse(m) *s c) = x)"
1196 using vector_affinity_eq[where m=m and x=x and y=y and c=c]
1200 fixes f :: "real^'n \<Rightarrow> real"
1201 shows "(\<chi> i. f (axis i 1)) = (\<Sum>i\<in>Basis. f i *\<^sub>R i)"
1202 unfolding euclidean_eq_iff[where 'a="real^'n"]
1203 by simp (simp add: Basis_vec_def inner_axis)
1205 lemma const_vector_cart:"((\<chi> i. d)::real^'n) = (\<Sum>i\<in>Basis. d *\<^sub>R i)"
1206 by (rule vector_cart)
1208 subsection "Convex Euclidean Space"
1210 lemma Cart_1:"(1::real^'n) = \<Sum>Basis"
1211 using const_vector_cart[of 1] by (simp add: one_vec_def)
1213 declare vector_add_ldistrib[simp] vector_ssub_ldistrib[simp] vector_smult_assoc[simp] vector_smult_rneg[simp]
1214 declare vector_sadd_rdistrib[simp] vector_sub_rdistrib[simp]
1216 lemmas vector_component_simps = vector_minus_component vector_smult_component vector_add_component less_eq_vec_def vec_lambda_beta vector_uminus_component
1218 lemma convex_box_cart:
1219 assumes "\<And>i. convex {x. P i x}"
1220 shows "convex {x. \<forall>i. P i (x$i)}"
1221 using assms unfolding convex_def by auto
1223 lemma convex_positive_orthant_cart: "convex {x::real^'n. (\<forall>i. 0 \<le> x$i)}"
1224 by (rule convex_box_cart) (simp add: atLeast_def[symmetric] convex_real_interval)
1226 lemma unit_interval_convex_hull_cart:
1227 "cbox (0::real^'n) 1 = convex hull {x. \<forall>i. (x$i = 0) \<or> (x$i = 1)}"
1228 unfolding Cart_1 unit_interval_convex_hull[where 'a="real^'n"] box_real[symmetric]
1229 by (rule arg_cong[where f="\<lambda>x. convex hull x"]) (simp add: Basis_vec_def inner_axis)
1231 lemma cube_convex_hull_cart:
1233 obtains s::"(real^'n) set"
1234 where "finite s" "cbox (x - (\<chi> i. d)) (x + (\<chi> i. d)) = convex hull s"
1236 from assms obtain s where "finite s"
1237 and "cbox (x - setsum (op *\<^sub>R d) Basis) (x + setsum (op *\<^sub>R d) Basis) = convex hull s"
1238 by (rule cube_convex_hull)
1239 with that[of s] show thesis
1240 by (simp add: const_vector_cart)
1244 subsection "Derivative"
1246 definition "jacobian f net = matrix(frechet_derivative f net)"
1248 lemma jacobian_works:
1249 "(f::(real^'a) \<Rightarrow> (real^'b)) differentiable net \<longleftrightarrow>
1250 (f has_derivative (\<lambda>h. (jacobian f net) *v h)) net"
1252 unfolding jacobian_def
1253 apply (simp only: matrix_works[OF linear_frechet_derivative]) defer
1254 apply (rule differentiableI)
1256 unfolding frechet_derivative_works
1261 subsection \<open>Component of the differential must be zero if it exists at a local
1262 maximum or minimum for that corresponding component.\<close>
1264 lemma differential_zero_maxmin_cart:
1265 fixes f::"real^'a \<Rightarrow> real^'b"
1266 assumes "0 < e" "((\<forall>y \<in> ball x e. (f y)$k \<le> (f x)$k) \<or> (\<forall>y\<in>ball x e. (f x)$k \<le> (f y)$k))"
1267 "f differentiable (at x)"
1268 shows "jacobian f (at x) $ k = 0"
1269 using differential_zero_maxmin_component[of "axis k 1" e x f] assms
1270 vector_cart[of "\<lambda>j. frechet_derivative f (at x) j $ k"]
1271 by (simp add: Basis_vec_def axis_eq_axis inner_axis jacobian_def matrix_def)
1273 subsection \<open>Lemmas for working on @{typ "real^1"}\<close>
1275 lemma forall_1[simp]: "(\<forall>i::1. P i) \<longleftrightarrow> P 1"
1276 by (metis (full_types) num1_eq_iff)
1278 lemma ex_1[simp]: "(\<exists>x::1. P x) \<longleftrightarrow> P 1"
1279 by auto (metis (full_types) num1_eq_iff)
1283 shows "x = 1 \<or> x = 2"
1286 then have "0 <= z" and "z < 2" by simp_all
1287 then have "z = 0 | z = 1" by arith
1288 then show ?case by auto
1291 lemma forall_2: "(\<forall>i::2. P i) \<longleftrightarrow> P 1 \<and> P 2"
1292 by (metis exhaust_2)
1296 shows "x = 1 \<or> x = 2 \<or> x = 3"
1299 then have "0 <= z" and "z < 3" by simp_all
1300 then have "z = 0 \<or> z = 1 \<or> z = 2" by arith
1301 then show ?case by auto
1304 lemma forall_3: "(\<forall>i::3. P i) \<longleftrightarrow> P 1 \<and> P 2 \<and> P 3"
1305 by (metis exhaust_3)
1307 lemma UNIV_1 [simp]: "UNIV = {1::1}"
1308 by (auto simp add: num1_eq_iff)
1310 lemma UNIV_2: "UNIV = {1::2, 2::2}"
1311 using exhaust_2 by auto
1313 lemma UNIV_3: "UNIV = {1::3, 2::3, 3::3}"
1314 using exhaust_3 by auto
1316 lemma setsum_1: "setsum f (UNIV::1 set) = f 1"
1317 unfolding UNIV_1 by simp
1319 lemma setsum_2: "setsum f (UNIV::2 set) = f 1 + f 2"
1320 unfolding UNIV_2 by simp
1322 lemma setsum_3: "setsum f (UNIV::3 set) = f 1 + f 2 + f 3"
1323 unfolding UNIV_3 by (simp add: ac_simps)
1325 instantiation num1 :: cart_one
1330 show "CARD(1) = Suc 0" by auto
1335 subsection\<open>The collapse of the general concepts to dimension one.\<close>
1337 lemma vector_one: "(x::'a ^1) = (\<chi> i. (x$1))"
1338 by (simp add: vec_eq_iff)
1340 lemma forall_one: "(\<forall>(x::'a ^1). P x) \<longleftrightarrow> (\<forall>x. P(\<chi> i. x))"
1342 apply (erule_tac x= "x$1" in allE)
1343 apply (simp only: vector_one[symmetric])
1346 lemma norm_vector_1: "norm (x :: _^1) = norm (x$1)"
1347 by (simp add: norm_vec_def)
1349 lemma norm_real: "norm(x::real ^ 1) = \<bar>x$1\<bar>"
1350 by (simp add: norm_vector_1)
1352 lemma dist_real: "dist(x::real ^ 1) y = \<bar>(x$1) - (y$1)\<bar>"
1353 by (auto simp add: norm_real dist_norm)
1356 subsection\<open>Explicit vector construction from lists.\<close>
1358 definition "vector l = (\<chi> i. foldr (\<lambda>x f n. fun_upd (f (n+1)) n x) l (\<lambda>n x. 0) 1 i)"
1360 lemma vector_1: "(vector[x]) $1 = x"
1361 unfolding vector_def by simp
1364 "(vector[x,y]) $1 = x"
1365 "(vector[x,y] :: 'a^2)$2 = (y::'a::zero)"
1366 unfolding vector_def by simp_all
1369 "(vector [x,y,z] ::('a::zero)^3)$1 = x"
1370 "(vector [x,y,z] ::('a::zero)^3)$2 = y"
1371 "(vector [x,y,z] ::('a::zero)^3)$3 = z"
1372 unfolding vector_def by simp_all
1374 lemma forall_vector_1: "(\<forall>v::'a::zero^1. P v) \<longleftrightarrow> (\<forall>x. P(vector[x]))"
1376 apply (erule_tac x="v$1" in allE)
1377 apply (subgoal_tac "vector [v$1] = v")
1379 apply (vector vector_def)
1383 lemma forall_vector_2: "(\<forall>v::'a::zero^2. P v) \<longleftrightarrow> (\<forall>x y. P(vector[x, y]))"
1385 apply (erule_tac x="v$1" in allE)
1386 apply (erule_tac x="v$2" in allE)
1387 apply (subgoal_tac "vector [v$1, v$2] = v")
1389 apply (vector vector_def)
1390 apply (simp add: forall_2)
1393 lemma forall_vector_3: "(\<forall>v::'a::zero^3. P v) \<longleftrightarrow> (\<forall>x y z. P(vector[x, y, z]))"
1395 apply (erule_tac x="v$1" in allE)
1396 apply (erule_tac x="v$2" in allE)
1397 apply (erule_tac x="v$3" in allE)
1398 apply (subgoal_tac "vector [v$1, v$2, v$3] = v")
1400 apply (vector vector_def)
1401 apply (simp add: forall_3)
1404 lemma bounded_linear_component_cart[intro]: "bounded_linear (\<lambda>x::real^'n. x $ k)"
1405 apply (rule bounded_linearI[where K=1])
1406 using component_le_norm_cart[of _ k] unfolding real_norm_def by auto
1408 lemma integral_component_eq_cart[simp]:
1409 fixes f :: "'n::euclidean_space \<Rightarrow> real^'m"
1410 assumes "f integrable_on s"
1411 shows "integral s (\<lambda>x. f x $ k) = integral s f $ k"
1412 using integral_linear[OF assms(1) bounded_linear_component_cart,unfolded o_def] .
1414 lemma interval_split_cart:
1415 "{a..b::real^'n} \<inter> {x. x$k \<le> c} = {a .. (\<chi> i. if i = k then min (b$k) c else b$i)}"
1416 "cbox a b \<inter> {x. x$k \<ge> c} = {(\<chi> i. if i = k then max (a$k) c else a$i) .. b}"
1417 apply (rule_tac[!] set_eqI)
1418 unfolding Int_iff mem_interval_cart mem_Collect_eq interval_cbox_cart
1419 unfolding vec_lambda_beta