# HG changeset patch # User wenzelm # Date 1348837516 -7200 # Node ID 343bfcbad2ecfb040a4501d48c38c5d424cbadd6 # Parent 9b831f93d4e8e32e4ce25102c7ab2f55dcaa2810 tuned proofs; diff -r 9b831f93d4e8 -r 343bfcbad2ec src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy Fri Sep 28 13:16:10 2012 +0200 +++ b/src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy Fri Sep 28 15:05:16 2012 +0200 @@ -6,7 +6,8 @@ begin lemma delta_mult_idempotent: - "(if k=a then 1 else (0::'a::semiring_1)) * (if k=a then 1 else 0) = (if k=a then 1 else 0)" by (cases "k=a", auto) + "(if k=a then 1 else (0::'a::semiring_1)) * (if k=a then 1 else 0) = (if k=a then 1 else 0)" + by (cases "k=a") auto lemma setsum_Plus: "\finite A; finite B\ \ @@ -25,34 +26,42 @@ "setsum h {..i\{..j\{..j. j + i * B) {..j. j + i * B) {..j. j + i * B) ` {.. {i * B.. (\j. j + i * B) ` {.. (\j. j + i * B) ` {.. (\ x y. (\ i. (x$i) * (y$i)))" - instance .. + +definition "op * \ (\ x y. (\ i. (x$i) * (y$i)))" +instance .. + end instantiation vec :: (one, finite) one begin - definition "1 \ (\ i. 1)" - instance .. + +definition "1 \ (\ i. 1)" +instance .. + end instantiation vec :: (ord, finite) ord begin - definition "x \ y \ (\i. x$i \ y$i)" - definition "x < y \ (\i. x$i < y$i)" - instance .. + +definition "x \ y \ (\i. x$i \ y$i)" +definition "x < y \ (\i. x$i < y$i)" +instance .. + end text{* The ordering on one-dimensional vectors is linear. *} @@ -100,29 +109,30 @@ definition vector_scalar_mult:: "'a::times \ 'a ^ 'n \ 'a ^ 'n" (infixl "*s" 70) where "c *s x = (\ i. c * (x$i))" + subsection {* A naive proof procedure to lift really trivial arithmetic stuff from the basis of the vector space. *} method_setup vector = {* let val ss1 = HOL_basic_ss addsimps [@{thm setsum_addf} RS sym, - @{thm setsum_subtractf} RS sym, @{thm setsum_right_distrib}, - @{thm setsum_left_distrib}, @{thm setsum_negf} RS sym] + @{thm setsum_subtractf} RS sym, @{thm setsum_right_distrib}, + @{thm setsum_left_distrib}, @{thm setsum_negf} RS sym] val ss2 = @{simpset} addsimps [@{thm plus_vec_def}, @{thm times_vec_def}, @{thm minus_vec_def}, @{thm uminus_vec_def}, @{thm one_vec_def}, @{thm zero_vec_def}, @{thm vec_def}, @{thm scaleR_vec_def}, @{thm vec_lambda_beta}, @{thm vector_scalar_mult_def}] - fun vector_arith_tac ths = - simp_tac ss1 - THEN' (fn i => rtac @{thm setsum_cong2} i + fun vector_arith_tac ths = + simp_tac ss1 + THEN' (fn i => rtac @{thm setsum_cong2} i ORELSE rtac @{thm setsum_0'} i ORELSE simp_tac (HOL_basic_ss addsimps [@{thm vec_eq_iff}]) i) - (* THEN' TRY o clarify_tac HOL_cs THEN' (TRY o rtac @{thm iffI}) *) - THEN' asm_full_simp_tac (ss2 addsimps ths) - in + (* THEN' TRY o clarify_tac HOL_cs THEN' (TRY o rtac @{thm iffI}) *) + THEN' asm_full_simp_tac (ss2 addsimps ths) +in Attrib.thms >> (fn ths => K (SIMPLE_METHOD' (vector_arith_tac ths))) - end +end *} "lift trivial vector statements to real arith statements" lemma vec_0[simp]: "vec 0 = 0" by (vector zero_vec_def) @@ -137,12 +147,17 @@ lemma vec_cmul: "vec(c * x) = c *s vec x " by (vector vec_def) lemma vec_neg: "vec(- x) = - vec x " by (vector vec_def) -lemma vec_setsum: assumes fS: "finite S" +lemma vec_setsum: + assumes "finite S" shows "vec(setsum f S) = setsum (vec o f) S" - apply (induct rule: finite_induct[OF fS]) - apply (simp) - apply (auto simp add: vec_add) - done + using assms +proof induct + case empty + then show ?case by simp +next + case insert + then show ?case by (auto simp add: vec_add) +qed text{* Obvious "component-pushing". *} @@ -162,6 +177,7 @@ vector_smult_component vector_minus_component vector_uminus_component vector_scaleR_component cond_component + subsection {* Some frequently useful arithmetic lemmas over vectors. *} instance vec :: (semigroup_mult, finite) semigroup_mult @@ -200,21 +216,21 @@ instance vec :: (ring_1, finite) ring_1 .. instance vec :: (real_algebra, finite) real_algebra - apply intro_classes - apply (simp_all add: vec_eq_iff) - done + by default (simp_all add: vec_eq_iff) instance vec :: (real_algebra_1, finite) real_algebra_1 .. -lemma of_nat_index: - "(of_nat n :: 'a::semiring_1 ^'n)$i = of_nat n" - apply (induct n) - apply vector - apply vector - done +lemma of_nat_index: "(of_nat n :: 'a::semiring_1 ^'n)$i = of_nat n" +proof (induct n) + case 0 + then show ?case by vector +next + case Suc + then show ?case by vector +qed -lemma one_index[simp]: - "(1 :: 'a::one ^'n)$i = 1" by vector +lemma one_index[simp]: "(1 :: 'a::one ^'n)$i = 1" + by vector instance vec :: (semiring_char_0, finite) semiring_char_0 proof @@ -227,7 +243,7 @@ instance vec :: (semiring_numeral, finite) semiring_numeral .. lemma numeral_index [simp]: "numeral w $ i = numeral w" - by (induct w, simp_all only: numeral.simps vector_add_component one_index) + by (induct w) (simp_all only: numeral.simps vector_add_component one_index) lemma neg_numeral_index [simp]: "neg_numeral w $ i = neg_numeral w" by (simp only: neg_numeral_def vector_uminus_component numeral_index) @@ -291,7 +307,13 @@ lemma setsum_component [simp]: fixes f:: " 'a \ ('b::comm_monoid_add) ^'n" shows "(setsum f S)$i = setsum (\x. (f x)$i) S" - by (cases "finite S", induct S set: finite, simp_all) +proof (cases "finite S") + case True + then show ?thesis by induct simp_all +next + case False + then show ?thesis by simp +qed lemma setsum_eq: "setsum f S = (\ i. setsum (\x. (f x)$i ) S)" by (simp add: vec_eq_iff) @@ -306,7 +328,7 @@ fixes f:: "'a \ real ^'n" assumes fP: "finite P" and fPs: "\Q. Q \ P \ norm (setsum f Q) \ e" shows "setsum (\x. norm (f x)) P \ 2 * real CARD('n) * e" -proof- +proof - let ?d = "real CARD('n)" let ?nf = "\x. norm (f x)" let ?U = "UNIV :: 'n set" @@ -314,7 +336,9 @@ by (rule setsum_commute) have th1: "2 * ?d * e = of_nat (card ?U) * (2 * e)" by (simp add: real_of_nat_def) have "setsum ?nf P \ setsum (\x. setsum (\i. \f x $ i\) ?U) P" - apply (rule setsum_mono) by (rule norm_le_l1_cart) + apply (rule setsum_mono) + apply (rule norm_le_l1_cart) + done also have "\ \ 2 * ?d * e" unfolding th0 th1 proof(rule setsum_bounded) @@ -333,7 +357,8 @@ have "setsum (\x. \f x $ i\) P = setsum (\x. \f x $ i\) ?Pp + setsum (\x. \f x $ i\) ?Pn" apply (subst thp) apply (rule setsum_Un_zero) - using fP thp0 by auto + using fP thp0 apply auto + done also have "\ \ 2*e" using Pne Ppe by arith finally show "setsum (\x. \f x $ i\) P \ 2*e" . qed @@ -344,8 +369,10 @@ lemma split_dimensions'[consumes 1]: assumes "k < DIM('a::euclidean_space^'b)" - obtains i j where "i < CARD('b::finite)" and "j < DIM('a::euclidean_space)" and "k = j + i * DIM('a::euclidean_space)" -using split_times_into_modulo[OF assms[simplified]] . + obtains i j where "i < CARD('b::finite)" + and "j < DIM('a::euclidean_space)" + and "k = j + i * DIM('a::euclidean_space)" + using split_times_into_modulo[OF assms[simplified]] . lemma cart_euclidean_bound[intro]: assumes j:"j < DIM('a::euclidean_space)" @@ -356,12 +383,14 @@ "(\i (\(i::'b::finite) j. j P (j + \' i * DIM('a)))" (is "?l \ ?r") proof (safe elim!: split_times_into_modulo) - fix i :: 'b and j assume "j < DIM('a)" + fix i :: 'b and j + assume "j < DIM('a)" note linear_less_than_times[OF pi'_range[of i] this] moreover assume "?l" ultimately show "P (j + \' i * DIM('a))" by auto next - fix i j assume "i < CARD('b)" "j < DIM('a)" and "?r" + fix i j + assume "i < CARD('b)" "j < DIM('a)" and "?r" from `?r`[rule_format, OF `j < DIM('a)`, of "\ i"] `i < CARD('b)` show "P (j + i * DIM('a))" by simp qed @@ -408,9 +437,11 @@ have "i = j" proof (cases rule: linorder_cases) - assume "i < j" from eq[OF this `x < A` *] show "i = j" by simp + assume "i < j" + from eq[OF this `x < A` *] show "i = j" by simp next - assume "j < i" from eq[OF this `y < A` *[symmetric]] show "i = j" by simp + assume "j < i" + from eq[OF this `y < A` *[symmetric]] show "i = j" by simp qed simp thus "x = y \ i = j" using * by simp qed simp @@ -424,6 +455,7 @@ unfolding less_vec_def apply(subst eucl_less) by (simp add: cart_simps) qed + subsection{* Basis vectors in coordinate directions. *} definition "cart_basis k = (\ i. if i = k then 1 else 0)" @@ -435,7 +467,8 @@ shows "norm (cart_basis k :: real ^'n) = 1" apply (simp add: cart_basis_def norm_eq_sqrt_inner) unfolding inner_vec_def apply (vector delta_mult_idempotent) - using setsum_delta[of "UNIV :: 'n set" "k" "\k. 1::real"] by auto + using setsum_delta[of "UNIV :: 'n set" "k" "\k. 1::real"] apply auto + done lemma norm_basis_1: "norm(cart_basis 1 :: real ^'n::{finite,one}) = 1" by (rule norm_basis) @@ -443,10 +476,11 @@ lemma vector_choose_size: "0 <= c ==> \(x::real^'n). norm x = c" by (rule exI[where x="c *\<^sub>R cart_basis arbitrary"]) simp -lemma vector_choose_dist: assumes e: "0 <= e" +lemma vector_choose_dist: + assumes e: "0 <= e" shows "\(y::real^'n). dist x y = e" -proof- - from vector_choose_size[OF e] obtain c:: "real ^'n" where "norm c = e" +proof - + from vector_choose_size[OF e] obtain c:: "real ^'n" where "norm c = e" by blast then have "dist x (x - c) = e" by (simp add: dist_norm) then show ?thesis by blast @@ -455,23 +489,22 @@ lemma basis_inj[intro]: "inj (cart_basis :: 'n \ real ^'n)" by (simp add: inj_on_def vec_eq_iff) -lemma basis_expansion: - "setsum (\i. (x$i) *s cart_basis i) UNIV = (x::('a::ring_1) ^'n)" (is "?lhs = ?rhs" is "setsum ?f ?S = _") - by (auto simp add: vec_eq_iff if_distrib setsum_delta[of "?S", where ?'b = "'a", simplified] cong del: if_weak_cong) +lemma basis_expansion: "setsum (\i. (x$i) *s cart_basis i) UNIV = (x::('a::ring_1) ^'n)" + (is "?lhs = ?rhs" is "setsum ?f ?S = _") + by (auto simp add: vec_eq_iff + if_distrib setsum_delta[of "?S", where ?'b = "'a", simplified] cong del: if_weak_cong) lemma smult_conv_scaleR: "c *s x = scaleR c x" unfolding vector_scalar_mult_def scaleR_vec_def by simp -lemma basis_expansion': - "setsum (\i. (x$i) *\<^sub>R cart_basis i) UNIV = x" +lemma basis_expansion': "setsum (\i. (x$i) *\<^sub>R cart_basis i) UNIV = x" by (rule basis_expansion [where 'a=real, unfolded smult_conv_scaleR]) lemma basis_expansion_unique: "setsum (\i. f i *s cart_basis i) UNIV = (x::('a::comm_ring_1) ^'n) \ (\i. f i = x$i)" by (simp add: vec_eq_iff setsum_delta if_distrib cong del: if_weak_cong) -lemma dot_basis: - shows "cart_basis i \ x = x$i" "x \ (cart_basis i) = (x$i)" +lemma dot_basis: "cart_basis i \ x = x$i" "x \ (cart_basis i) = (x$i)" by (auto simp add: inner_vec_def cart_basis_def cond_application_beta if_distrib setsum_delta cong del: if_weak_cong) @@ -485,8 +518,7 @@ lemma basis_eq_0: "cart_basis i = (0::'a::semiring_1^'n) \ False" by (auto simp add: vec_eq_iff) -lemma basis_nonzero: - shows "cart_basis k \ (0:: 'a::semiring_1 ^'n)" +lemma basis_nonzero: "cart_basis k \ (0:: 'a::semiring_1 ^'n)" by (simp add: basis_eq_0) text {* some lemmas to map between Eucl and Cart *} @@ -496,25 +528,22 @@ subsection {* Orthogonality on cartesian products *} -lemma orthogonal_basis: - shows "orthogonal (cart_basis i) x \ x$i = (0::real)" +lemma orthogonal_basis: "orthogonal (cart_basis i) x \ x$i = (0::real)" by (auto simp add: orthogonal_def inner_vec_def cart_basis_def if_distrib cond_application_beta setsum_delta cong del: if_weak_cong) -lemma orthogonal_basis_basis: - shows "orthogonal (cart_basis i :: real^'n) (cart_basis j) \ i \ j" +lemma orthogonal_basis_basis: "orthogonal (cart_basis i :: real^'n) (cart_basis j) \ i \ j" unfolding orthogonal_basis[of i] basis_component[of j] by simp subsection {* Linearity on cartesian products *} lemma linear_vmul_component: - assumes lf: "linear f" + assumes "linear f" shows "linear (\x. f x $ k *\<^sub>R v)" - using lf - by (auto simp add: linear_def algebra_simps) + using assms by (auto simp add: linear_def algebra_simps) -subsection{* Adjoints on cartesian products *} +subsection {* Adjoints on cartesian products *} text {* TODO: The following lemmas about adjoints should hold for any Hilbert space (i.e. complete inner product space). @@ -525,26 +554,26 @@ fixes f:: "real ^'n \ real ^'m" assumes lf: "linear f" shows "\x y. f x \ y = x \ adjoint f y" -proof- +proof - let ?N = "UNIV :: 'n set" let ?M = "UNIV :: 'm set" have fN: "finite ?N" by simp have fM: "finite ?M" by simp - {fix y:: "real ^ 'm" + { fix y:: "real ^ 'm" let ?w = "(\ i. (f (cart_basis i) \ y)) :: real ^ 'n" - {fix x + { fix x have "f x \ y = f (setsum (\i. (x$i) *\<^sub>R cart_basis i) ?N) \ y" by (simp only: basis_expansion') also have "\ = (setsum (\i. (x$i) *\<^sub>R f (cart_basis i)) ?N) \ y" unfolding linear_setsum[OF lf fN] by (simp add: linear_cmul[OF lf]) finally have "f x \ y = x \ ?w" - apply (simp only: ) - apply (simp add: inner_vec_def setsum_left_distrib setsum_right_distrib setsum_commute[of _ ?M ?N] field_simps) - done} + by (simp add: inner_vec_def setsum_left_distrib + setsum_right_distrib setsum_commute[of _ ?M ?N] field_simps) + } } - then show ?thesis unfolding adjoint_def - some_eq_ex[of "\f'. \x y. f x \ y = x \ f' y"] + then show ?thesis + unfolding adjoint_def some_eq_ex[of "\f'. \x y. f x \ y = x \ f' y"] using choice_iff[of "\a b. \x. f x \ a = x \ b "] by metis qed @@ -566,7 +595,7 @@ fixes f:: "real ^'n \ real ^'m" assumes lf: "linear f" shows "x \ adjoint f y = f x \ y" - and "adjoint f y \ x = y \ f x" + and "adjoint f y \ x = y \ f x" by (simp_all add: adjoint_works[OF lf] inner_commute) lemma adjoint_adjoint: @@ -580,13 +609,16 @@ text{* Matrix notation. NB: an MxN matrix is of type @{typ "'a^'n^'m"}, not @{typ "'a^'m^'n"} *} -definition matrix_matrix_mult :: "('a::semiring_1) ^'n^'m \ 'a ^'p^'n \ 'a ^ 'p ^'m" (infixl "**" 70) +definition matrix_matrix_mult :: "('a::semiring_1) ^'n^'m \ 'a ^'p^'n \ 'a ^ 'p ^'m" + (infixl "**" 70) where "m ** m' == (\ i j. setsum (\k. ((m$i)$k) * ((m'$k)$j)) (UNIV :: 'n set)) ::'a ^ 'p ^'m" -definition matrix_vector_mult :: "('a::semiring_1) ^'n^'m \ 'a ^'n \ 'a ^ 'm" (infixl "*v" 70) +definition matrix_vector_mult :: "('a::semiring_1) ^'n^'m \ 'a ^'n \ 'a ^ 'm" + (infixl "*v" 70) where "m *v x \ (\ i. setsum (\j. ((m$i)$j) * (x$j)) (UNIV ::'n set)) :: 'a^'m" -definition vector_matrix_mult :: "'a ^ 'm \ ('a::semiring_1) ^'n^'m \ 'a ^'n " (infixl "v*" 70) +definition vector_matrix_mult :: "'a ^ 'm \ ('a::semiring_1) ^'n^'m \ 'a ^'n " + (infixl "v*" 70) where "v v* m == (\ j. setsum (\i. ((m$i)$j) * (v$i)) (UNIV :: 'm set)) :: 'a^'n" definition "(mat::'a::zero => 'a ^'n^'n) k = (\ i j. if i = j then k else 0)" @@ -606,7 +638,9 @@ shows "mat 1 ** A = A" apply (simp add: matrix_matrix_mult_def mat_def) apply vector - by (auto simp only: if_distrib cond_application_beta setsum_delta'[OF finite] mult_1_left mult_zero_left if_True UNIV_I) + apply (auto simp only: if_distrib cond_application_beta setsum_delta'[OF finite] + mult_1_left mult_zero_left if_True UNIV_I) + done lemma matrix_mul_rid: @@ -614,7 +648,9 @@ shows "A ** mat 1 = A" apply (simp add: matrix_matrix_mult_def mat_def) apply vector - by (auto simp only: if_distrib cond_application_beta setsum_delta[OF finite] mult_1_right mult_zero_right if_True UNIV_I cong: if_cong) + apply (auto simp only: if_distrib cond_application_beta setsum_delta[OF finite] + mult_1_right mult_zero_right if_True UNIV_I cong: if_cong) + done lemma matrix_mul_assoc: "A ** (B ** C) = (A ** B) ** C" apply (vector matrix_matrix_mult_def setsum_right_distrib setsum_left_distrib mult_assoc) @@ -623,17 +659,19 @@ done lemma matrix_vector_mul_assoc: "A *v (B *v x) = (A ** B) *v x" - apply (vector matrix_matrix_mult_def matrix_vector_mult_def setsum_right_distrib setsum_left_distrib mult_assoc) + apply (vector matrix_matrix_mult_def matrix_vector_mult_def + setsum_right_distrib setsum_left_distrib mult_assoc) apply (subst setsum_commute) apply simp done lemma matrix_vector_mul_lid: "mat 1 *v x = (x::'a::semiring_1 ^ 'n)" apply (vector matrix_vector_mult_def mat_def) - by (simp add: if_distrib cond_application_beta - setsum_delta' cong del: if_weak_cong) + apply (simp add: if_distrib cond_application_beta setsum_delta' cong del: if_weak_cong) + done -lemma matrix_transpose_mul: "transpose(A ** B) = transpose B ** transpose (A::'a::comm_semiring_1^_^_)" +lemma matrix_transpose_mul: + "transpose(A ** B) = transpose B ** transpose (A::'a::comm_semiring_1^_^_)" by (simp add: matrix_matrix_mult_def transpose_def vec_eq_iff mult_commute) lemma matrix_eq: @@ -645,16 +683,18 @@ apply (clarsimp simp add: matrix_vector_mult_def cart_basis_def if_distrib cond_application_beta vec_eq_iff cong del: if_weak_cong) apply (erule_tac x="cart_basis ia" in allE) apply (erule_tac x="i" in allE) - by (auto simp add: cart_basis_def if_distrib cond_application_beta setsum_delta[OF finite] cong del: if_weak_cong) + apply (auto simp add: cart_basis_def if_distrib cond_application_beta + setsum_delta[OF finite] cong del: if_weak_cong) + done -lemma matrix_vector_mul_component: - shows "((A::real^_^_) *v x)$k = (A$k) \ x" +lemma matrix_vector_mul_component: "((A::real^_^_) *v x)$k = (A$k) \ x" by (simp add: matrix_vector_mult_def inner_vec_def) lemma dot_lmul_matrix: "((x::real ^_) v* A) \ y = x \ (A *v y)" apply (simp add: inner_vec_def matrix_vector_mult_def vector_matrix_mult_def setsum_left_distrib setsum_right_distrib mult_ac) apply (subst setsum_commute) - by simp + apply simp + done lemma transpose_mat: "transpose (mat n) = mat n" by (vector transpose_def mat_def) @@ -673,28 +713,31 @@ by (simp add: row_def column_def transpose_def vec_eq_iff) lemma rows_transpose: "rows(transpose (A::'a::semiring_1^_^_)) = columns A" -by (auto simp add: rows_def columns_def row_transpose intro: set_eqI) + by (auto simp add: rows_def columns_def row_transpose intro: set_eqI) -lemma columns_transpose: "columns(transpose (A::'a::semiring_1^_^_)) = rows A" by (metis transpose_transpose rows_transpose) +lemma columns_transpose: "columns(transpose (A::'a::semiring_1^_^_)) = rows A" + by (metis transpose_transpose rows_transpose) text{* Two sometimes fruitful ways of looking at matrix-vector multiplication. *} lemma matrix_mult_dot: "A *v x = (\ i. A$i \ x)" by (simp add: matrix_vector_mult_def inner_vec_def) -lemma matrix_mult_vsum: "(A::'a::comm_semiring_1^'n^'m) *v x = setsum (\i. (x$i) *s column i A) (UNIV:: 'n set)" +lemma matrix_mult_vsum: + "(A::'a::comm_semiring_1^'n^'m) *v x = setsum (\i. (x$i) *s column i A) (UNIV:: 'n set)" by (simp add: matrix_vector_mult_def vec_eq_iff column_def mult_commute) lemma vector_componentwise: "(x::'a::ring_1^'n) = (\ j. setsum (\i. (x$i) * (cart_basis i :: 'a^'n)$j) (UNIV :: 'n set))" apply (subst basis_expansion[symmetric]) - by (vector vec_eq_iff setsum_component) + apply (vector vec_eq_iff setsum_component) + done lemma linear_componentwise: fixes f:: "real ^'m \ real ^ _" assumes lf: "linear f" shows "(f x)$j = setsum (\i. (x$i) * (f (cart_basis i)$j)) (UNIV :: 'm set)" (is "?lhs = ?rhs") -proof- +proof - let ?M = "(UNIV :: 'm set)" let ?N = "(UNIV :: 'n set)" have fM: "finite ?M" by simp @@ -702,48 +745,57 @@ unfolding vector_smult_component[symmetric] smult_conv_scaleR unfolding setsum_component[of "(\i.(x$i) *\<^sub>R f (cart_basis i :: real^'m))" ?M] .. - then show ?thesis unfolding linear_setsum_mul[OF lf fM, symmetric] basis_expansion' .. + then show ?thesis + unfolding linear_setsum_mul[OF lf fM, symmetric] basis_expansion' .. qed text{* Inverse matrices (not necessarily square) *} -definition "invertible(A::'a::semiring_1^'n^'m) \ (\A'::'a^'m^'n. A ** A' = mat 1 \ A' ** A = mat 1)" +definition + "invertible(A::'a::semiring_1^'n^'m) \ (\A'::'a^'m^'n. A ** A' = mat 1 \ A' ** A = mat 1)" -definition "matrix_inv(A:: 'a::semiring_1^'n^'m) = - (SOME A'::'a^'m^'n. A ** A' = mat 1 \ A' ** A = mat 1)" +definition + "matrix_inv(A:: 'a::semiring_1^'n^'m) = + (SOME A'::'a^'m^'n. A ** A' = mat 1 \ A' ** A = mat 1)" text{* Correspondence between matrices and linear operators. *} -definition matrix:: "('a::{plus,times, one, zero}^'m \ 'a ^ 'n) \ 'a^'m^'n" -where "matrix f = (\ i j. (f(cart_basis j))$i)" +definition matrix :: "('a::{plus,times, one, zero}^'m \ 'a ^ 'n) \ 'a^'m^'n" + where "matrix f = (\ i j. (f(cart_basis j))$i)" lemma matrix_vector_mul_linear: "linear(\x. A *v (x::real ^ _))" - by (simp add: linear_def matrix_vector_mult_def vec_eq_iff field_simps setsum_right_distrib setsum_addf) + by (simp add: linear_def matrix_vector_mult_def vec_eq_iff + field_simps setsum_right_distrib setsum_addf) -lemma matrix_works: assumes lf: "linear f" shows "matrix f *v x = f (x::real ^ 'n)" -apply (simp add: matrix_def matrix_vector_mult_def vec_eq_iff mult_commute) -apply clarify -apply (rule linear_componentwise[OF lf, symmetric]) -done +lemma matrix_works: + assumes lf: "linear f" + shows "matrix f *v x = f (x::real ^ 'n)" + apply (simp add: matrix_def matrix_vector_mult_def vec_eq_iff mult_commute) + apply clarify + apply (rule linear_componentwise[OF lf, symmetric]) + done -lemma matrix_vector_mul: "linear f ==> f = (\x. matrix f *v (x::real ^ 'n))" by (simp add: ext matrix_works) +lemma matrix_vector_mul: "linear f ==> f = (\x. matrix f *v (x::real ^ 'n))" + by (simp add: ext matrix_works) lemma matrix_of_matrix_vector_mul: "matrix(\x. A *v (x :: real ^ 'n)) = A" by (simp add: matrix_eq matrix_vector_mul_linear matrix_works) lemma matrix_compose: assumes lf: "linear (f::real^'n \ real^'m)" - and lg: "linear (g::real^'m \ real^_)" + and lg: "linear (g::real^'m \ real^_)" shows "matrix (g o f) = matrix g ** matrix f" using lf lg linear_compose[OF lf lg] matrix_works[OF linear_compose[OF lf lg]] - by (simp add: matrix_eq matrix_works matrix_vector_mul_assoc[symmetric] o_def) + by (simp add: matrix_eq matrix_works matrix_vector_mul_assoc[symmetric] o_def) -lemma matrix_vector_column:"(A::'a::comm_semiring_1^'n^_) *v x = setsum (\i. (x$i) *s ((transpose A)$i)) (UNIV:: 'n set)" +lemma matrix_vector_column: + "(A::'a::comm_semiring_1^'n^_) *v x = setsum (\i. (x$i) *s ((transpose A)$i)) (UNIV:: 'n set)" by (simp add: matrix_vector_mult_def transpose_def vec_eq_iff mult_commute) lemma adjoint_matrix: "adjoint(\x. (A::real^'n^'m) *v x) = (\x. transpose A *v x)" apply (rule adjoint_unique) - apply (simp add: transpose_def inner_vec_def matrix_vector_mult_def setsum_left_distrib setsum_right_distrib) + apply (simp add: transpose_def inner_vec_def matrix_vector_mult_def + setsum_left_distrib setsum_right_distrib) apply (subst setsum_commute) apply (auto simp add: mult_ac) done @@ -751,7 +803,10 @@ lemma matrix_adjoint: assumes lf: "linear (f :: real^'n \ real ^'m)" shows "matrix(adjoint f) = transpose(matrix f)" apply (subst matrix_vector_mul[OF lf]) - unfolding adjoint_matrix matrix_of_matrix_vector_mul .. + unfolding adjoint_matrix matrix_of_matrix_vector_mul + apply rule + done + subsection {* lambda skolemization on cartesian products *} @@ -759,15 +814,15 @@ lemma lambda_skolem: "(\i. \x. P i x) \ (\x::'a ^ 'n. \i. P i (x $ i))" (is "?lhs \ ?rhs") -proof- +proof - let ?S = "(UNIV :: 'n set)" - {assume H: "?rhs" - then have ?lhs by auto} + { assume H: "?rhs" + then have ?lhs by auto } moreover - {assume H: "?lhs" + { assume H: "?lhs" then obtain f where f:"\i. P i (f i)" unfolding choice_iff by metis let ?x = "(\ i. (f i)) :: 'a ^ 'n" - {fix i + { fix i from f have "P i (f i)" by metis then have "P i (?x $ i)" by auto } @@ -776,68 +831,69 @@ ultimately show ?thesis by metis qed + subsection {* Standard bases are a spanning set, and obviously finite. *} lemma span_stdbasis:"span {cart_basis i :: real^'n | i. i \ (UNIV :: 'n set)} = UNIV" -apply (rule set_eqI) -apply auto -apply (subst basis_expansion'[symmetric]) -apply (rule span_setsum) -apply simp -apply auto -apply (rule span_mul) -apply (rule span_superset) -apply auto -done + apply (rule set_eqI) + apply auto + apply (subst basis_expansion'[symmetric]) + apply (rule span_setsum) + apply simp + apply auto + apply (rule span_mul) + apply (rule span_superset) + apply auto + done lemma finite_stdbasis: "finite {cart_basis i ::real^'n |i. i\ (UNIV:: 'n set)}" (is "finite ?S") -proof- - have eq: "?S = cart_basis ` UNIV" by blast - show ?thesis unfolding eq by auto +proof - + have "?S = cart_basis ` UNIV" by blast + then show ?thesis by auto qed lemma card_stdbasis: "card {cart_basis i ::real^'n |i. i\ (UNIV :: 'n set)} = CARD('n)" (is "card ?S = _") -proof- - have eq: "?S = cart_basis ` UNIV" by blast - show ?thesis unfolding eq using card_image[OF basis_inj] by simp +proof - + have "?S = cart_basis ` UNIV" by blast + then show ?thesis using card_image[OF basis_inj] by simp qed - lemma independent_stdbasis_lemma: assumes x: "(x::real ^ 'n) \ span (cart_basis ` S)" - and iS: "i \ S" + and iS: "i \ S" shows "(x$i) = 0" -proof- +proof - let ?U = "UNIV :: 'n set" let ?B = "cart_basis ` S" let ?P = "{(x::real^_). \i\ ?U. i \ S \ x$i =0}" - {fix x::"real^_" assume xS: "x\ ?B" - from xS have "x \ ?P" by auto} - moreover - have "subspace ?P" - by (auto simp add: subspace_def) - ultimately show ?thesis - using x span_induct[of x ?B ?P] iS by blast + { fix x::"real^_" assume xS: "x\ ?B" + from xS have "x \ ?P" by auto } + moreover + have "subspace ?P" + by (auto simp add: subspace_def) + ultimately show ?thesis + using x span_induct[of x ?B ?P] iS by blast qed lemma independent_stdbasis: "independent {cart_basis i ::real^'n |i. i\ (UNIV :: 'n set)}" -proof- +proof - let ?I = "UNIV :: 'n set" let ?b = "cart_basis :: _ \ real ^'n" let ?B = "?b ` ?I" - have eq: "{?b i|i. i \ ?I} = ?B" - by auto - {assume d: "dependent ?B" + have eq: "{?b i|i. i \ ?I} = ?B" by auto + { assume d: "dependent ?B" then obtain k where k: "k \ ?I" "?b k \ span (?B - {?b k})" unfolding dependent_def by auto have eq1: "?B - {?b k} = ?B - ?b ` {k}" by simp have eq2: "?B - {?b k} = ?b ` (?I - {k})" unfolding eq1 apply (rule inj_on_image_set_diff[symmetric]) - apply (rule basis_inj) using k(1) by auto + apply (rule basis_inj) using k(1) + apply auto + done from k(2) have th0: "?b k \ span (?b ` (?I - {k}))" unfolding eq2 . from independent_stdbasis_lemma[OF th0, of k, simplified] - have False by simp} + have False by simp } then show ?thesis unfolding eq dependent_def .. qed @@ -846,27 +902,31 @@ lemma linear_eq_stdbasis_cart: assumes lf: "linear (f::real^'m \ _)" and lg: "linear g" - and fg: "\i. f (cart_basis i) = g(cart_basis i)" + and fg: "\i. f (cart_basis i) = g(cart_basis i)" shows "f = g" -proof- +proof - let ?U = "UNIV :: 'm set" let ?I = "{cart_basis i:: real^'m|i. i \ ?U}" - {fix x assume x: "x \ (UNIV :: (real^'m) set)" + { fix x assume x: "x \ (UNIV :: (real^'m) set)" from equalityD2[OF span_stdbasis] have IU: " (UNIV :: (real^'m) set) \ span ?I" by blast from linear_eq[OF lf lg IU] fg x - have "f x = g x" unfolding Ball_def mem_Collect_eq by metis} + have "f x = g x" unfolding Ball_def mem_Collect_eq by metis + } then show ?thesis by auto qed lemma bilinear_eq_stdbasis_cart: assumes bf: "bilinear (f:: real^'m \ real^'n \ _)" - and bg: "bilinear g" - and fg: "\i j. f (cart_basis i) (cart_basis j) = g (cart_basis i) (cart_basis j)" + and bg: "bilinear g" + and fg: "\i j. f (cart_basis i) (cart_basis j) = g (cart_basis i) (cart_basis j)" shows "f = g" -proof- - from fg have th: "\x \ {cart_basis i| i. i\ (UNIV :: 'm set)}. \y\ {cart_basis j |j. j \ (UNIV :: 'n set)}. f x y = g x y" by blast - from bilinear_eq[OF bf bg equalityD2[OF span_stdbasis] equalityD2[OF span_stdbasis] th] show ?thesis by blast +proof - + from fg have th: "\x \ {cart_basis i| i. i\ (UNIV :: 'm set)}. + \y\ {cart_basis j |j. j \ (UNIV :: 'n set)}. f x y = g x y" + by blast + from bilinear_eq[OF bf bg equalityD2[OF span_stdbasis] equalityD2[OF span_stdbasis] th] + show ?thesis by blast qed lemma left_invertible_transpose: @@ -878,21 +938,21 @@ by (metis matrix_transpose_mul transpose_mat transpose_transpose) lemma matrix_left_invertible_injective: -"(\B. (B::real^'m^'n) ** (A::real^'n^'m) = mat 1) \ (\x y. A *v x = A *v y \ x = y)" -proof- - {fix B:: "real^'m^'n" and x y assume B: "B ** A = mat 1" and xy: "A *v x = A*v y" + "(\B. (B::real^'m^'n) ** (A::real^'n^'m) = mat 1) \ (\x y. A *v x = A *v y \ x = y)" +proof - + { fix B:: "real^'m^'n" and x y assume B: "B ** A = mat 1" and xy: "A *v x = A*v y" from xy have "B*v (A *v x) = B *v (A*v y)" by simp hence "x = y" - unfolding matrix_vector_mul_assoc B matrix_vector_mul_lid .} + unfolding matrix_vector_mul_assoc B matrix_vector_mul_lid . } moreover - {assume A: "\x y. A *v x = A *v y \ x = y" + { assume A: "\x y. A *v x = A *v y \ x = y" hence i: "inj (op *v A)" unfolding inj_on_def by auto from linear_injective_left_inverse[OF matrix_vector_mul_linear i] obtain g where g: "linear g" "g o op *v A = id" by blast have "matrix g ** A = mat 1" unfolding matrix_eq matrix_vector_mul_lid matrix_vector_mul_assoc[symmetric] matrix_works[OF g(1)] using g(2) by (simp add: fun_eq_iff) - then have "\B. (B::real ^'m^'n) ** A = mat 1" by blast} + then have "\B. (B::real ^'m^'n) ** A = mat 1" by blast } ultimately show ?thesis by blast qed @@ -903,15 +963,16 @@ by (simp add: inj_on_def) lemma matrix_right_invertible_surjective: -"(\B. (A::real^'n^'m) ** (B::real^'m^'n) = mat 1) \ surj (\x. A *v x)" -proof- - {fix B :: "real ^'m^'n" assume AB: "A ** B = mat 1" - {fix x :: "real ^ 'm" + "(\B. (A::real^'n^'m) ** (B::real^'m^'n) = mat 1) \ surj (\x. A *v x)" +proof - + { fix B :: "real ^'m^'n" + assume AB: "A ** B = mat 1" + { fix x :: "real ^ 'm" have "A *v (B *v x) = x" - by (simp add: matrix_vector_mul_lid matrix_vector_mul_assoc AB)} + by (simp add: matrix_vector_mul_lid matrix_vector_mul_assoc AB) } hence "surj (op *v A)" unfolding surj_def by metis } moreover - {assume sf: "surj (op *v A)" + { assume sf: "surj (op *v A)" from linear_surjective_right_inverse[OF matrix_vector_mul_linear sf] obtain g:: "real ^'m \ real ^'n" where g: "linear g" "op *v A o g = id" by blast @@ -928,13 +989,14 @@ lemma matrix_left_invertible_independent_columns: fixes A :: "real^'n^'m" - shows "(\(B::real ^'m^'n). B ** A = mat 1) \ (\c. setsum (\i. c i *s column i A) (UNIV :: 'n set) = 0 \ (\i. c i = 0))" - (is "?lhs \ ?rhs") -proof- + shows "(\(B::real ^'m^'n). B ** A = mat 1) \ + (\c. setsum (\i. c i *s column i A) (UNIV :: 'n set) = 0 \ (\i. c i = 0))" + (is "?lhs \ ?rhs") +proof - let ?U = "UNIV :: 'n set" - {assume k: "\x. A *v x = 0 \ x = 0" - {fix c i assume c: "setsum (\i. c i *s column i A) ?U = 0" - and i: "i \ ?U" + { assume k: "\x. A *v x = 0 \ x = 0" + { fix c i + assume c: "setsum (\i. c i *s column i A) ?U = 0" and i: "i \ ?U" let ?x = "\ i. c i" have th0:"A *v ?x = 0" using c @@ -942,82 +1004,95 @@ by auto from k[rule_format, OF th0] i have "c i = 0" by (vector vec_eq_iff)} - hence ?rhs by blast} + hence ?rhs by blast } moreover - {assume H: ?rhs - {fix x assume x: "A *v x = 0" + { assume H: ?rhs + { fix x assume x: "A *v x = 0" let ?c = "\i. ((x$i ):: real)" from H[rule_format, of ?c, unfolded matrix_mult_vsum[symmetric], OF x] - have "x = 0" by vector}} + have "x = 0" by vector } + } ultimately show ?thesis unfolding matrix_left_invertible_ker by blast qed lemma matrix_right_invertible_independent_rows: fixes A :: "real^'n^'m" - shows "(\(B::real^'m^'n). A ** B = mat 1) \ (\c. setsum (\i. c i *s row i A) (UNIV :: 'm set) = 0 \ (\i. c i = 0))" + shows "(\(B::real^'m^'n). A ** B = mat 1) \ + (\c. setsum (\i. c i *s row i A) (UNIV :: 'm set) = 0 \ (\i. c i = 0))" unfolding left_invertible_transpose[symmetric] matrix_left_invertible_independent_columns by (simp add: column_transpose) lemma matrix_right_invertible_span_columns: - "(\(B::real ^'n^'m). (A::real ^'m^'n) ** B = mat 1) \ span (columns A) = UNIV" (is "?lhs = ?rhs") -proof- + "(\(B::real ^'n^'m). (A::real ^'m^'n) ** B = mat 1) \ + span (columns A) = UNIV" (is "?lhs = ?rhs") +proof - let ?U = "UNIV :: 'm set" have fU: "finite ?U" by simp have lhseq: "?lhs \ (\y. \(x::real^'m). setsum (\i. (x$i) *s column i A) ?U = y)" unfolding matrix_right_invertible_surjective matrix_mult_vsum surj_def - apply (subst eq_commute) .. + apply (subst eq_commute) + apply rule + done have rhseq: "?rhs \ (\x. x \ span (columns A))" by blast - {assume h: ?lhs - {fix x:: "real ^'n" - from h[unfolded lhseq, rule_format, of x] obtain y:: "real ^'m" - where y: "setsum (\i. (y$i) *s column i A) ?U = x" by blast - have "x \ span (columns A)" - unfolding y[symmetric] - apply (rule span_setsum[OF fU]) - apply clarify - unfolding smult_conv_scaleR - apply (rule span_mul) - apply (rule span_superset) - unfolding columns_def - by blast} - then have ?rhs unfolding rhseq by blast} + { assume h: ?lhs + { fix x:: "real ^'n" + from h[unfolded lhseq, rule_format, of x] obtain y :: "real ^'m" + where y: "setsum (\i. (y$i) *s column i A) ?U = x" by blast + have "x \ span (columns A)" + unfolding y[symmetric] + apply (rule span_setsum[OF fU]) + apply clarify + unfolding smult_conv_scaleR + apply (rule span_mul) + apply (rule span_superset) + unfolding columns_def + apply blast + done + } + then have ?rhs unfolding rhseq by blast } moreover - {assume h:?rhs + { assume h:?rhs let ?P = "\(y::real ^'n). \(x::real^'m). setsum (\i. (x$i) *s column i A) ?U = y" - {fix y have "?P y" - proof(rule span_induct_alt[of ?P "columns A", folded smult_conv_scaleR]) + { fix y + have "?P y" + proof (rule span_induct_alt[of ?P "columns A", folded smult_conv_scaleR]) show "\x\real ^ 'm. setsum (\i. (x$i) *s column i A) ?U = 0" by (rule exI[where x=0], simp) next - fix c y1 y2 assume y1: "y1 \ columns A" and y2: "?P y2" + fix c y1 y2 + assume y1: "y1 \ columns A" and y2: "?P y2" from y1 obtain i where i: "i \ ?U" "y1 = column i A" unfolding columns_def by blast from y2 obtain x:: "real ^'m" where x: "setsum (\i. (x$i) *s column i A) ?U = y2" by blast let ?x = "(\ j. if j = i then c + (x$i) else (x$j))::real^'m" show "?P (c*s y1 + y2)" - proof(rule exI[where x= "?x"], vector, auto simp add: i x[symmetric] if_distrib right_distrib cond_application_beta cong del: if_weak_cong) - fix j - have th: "\xa \ ?U. (if xa = i then (c + (x$i)) * ((column xa A)$j) - else (x$xa) * ((column xa A$j))) = (if xa = i then c * ((column i A)$j) else 0) + ((x$xa) * ((column xa A)$j))" using i(1) - by (simp add: field_simps) - have "setsum (\xa. if xa = i then (c + (x$i)) * ((column xa A)$j) - else (x$xa) * ((column xa A$j))) ?U = setsum (\xa. (if xa = i then c * ((column i A)$j) else 0) + ((x$xa) * ((column xa A)$j))) ?U" - apply (rule setsum_cong[OF refl]) - using th by blast - also have "\ = setsum (\xa. if xa = i then c * ((column i A)$j) else 0) ?U + setsum (\xa. ((x$xa) * ((column xa A)$j))) ?U" - by (simp add: setsum_addf) - also have "\ = c * ((column i A)$j) + setsum (\xa. ((x$xa) * ((column xa A)$j))) ?U" - unfolding setsum_delta[OF fU] - using i(1) by simp - finally show "setsum (\xa. if xa = i then (c + (x$i)) * ((column xa A)$j) - else (x$xa) * ((column xa A$j))) ?U = c * ((column i A)$j) + setsum (\xa. ((x$xa) * ((column xa A)$j))) ?U" . - qed - next - show "y \ span (columns A)" unfolding h by blast - qed} - then have ?lhs unfolding lhseq ..} + proof (rule exI[where x= "?x"], vector, auto simp add: i x[symmetric] if_distrib right_distrib cond_application_beta cong del: if_weak_cong) + fix j + have th: "\xa \ ?U. (if xa = i then (c + (x$i)) * ((column xa A)$j) + else (x$xa) * ((column xa A$j))) = (if xa = i then c * ((column i A)$j) else 0) + ((x$xa) * ((column xa A)$j))" + using i(1) by (simp add: field_simps) + have "setsum (\xa. if xa = i then (c + (x$i)) * ((column xa A)$j) + else (x$xa) * ((column xa A$j))) ?U = setsum (\xa. (if xa = i then c * ((column i A)$j) else 0) + ((x$xa) * ((column xa A)$j))) ?U" + apply (rule setsum_cong[OF refl]) + using th apply blast + done + also have "\ = setsum (\xa. if xa = i then c * ((column i A)$j) else 0) ?U + setsum (\xa. ((x$xa) * ((column xa A)$j))) ?U" + by (simp add: setsum_addf) + also have "\ = c * ((column i A)$j) + setsum (\xa. ((x$xa) * ((column xa A)$j))) ?U" + unfolding setsum_delta[OF fU] + using i(1) by simp + finally show "setsum (\xa. if xa = i then (c + (x$i)) * ((column xa A)$j) + else (x$xa) * ((column xa A$j))) ?U = c * ((column i A)$j) + setsum (\xa. ((x$xa) * ((column xa A)$j))) ?U" . + qed + next + show "y \ span (columns A)" + unfolding h by blast + qed + } + then have ?lhs unfolding lhseq .. + } ultimately show ?thesis by blast qed @@ -1026,29 +1101,34 @@ unfolding right_invertible_transpose[symmetric] unfolding columns_transpose[symmetric] unfolding matrix_right_invertible_span_columns - .. + .. text {* The same result in terms of square matrices. *} lemma matrix_left_right_inverse: fixes A A' :: "real ^'n^'n" shows "A ** A' = mat 1 \ A' ** A = mat 1" -proof- - {fix A A' :: "real ^'n^'n" assume AA': "A ** A' = mat 1" +proof - + { fix A A' :: "real ^'n^'n" + assume AA': "A ** A' = mat 1" have sA: "surj (op *v A)" unfolding surj_def apply clarify apply (rule_tac x="(A' *v y)" in exI) - by (simp add: matrix_vector_mul_assoc AA' matrix_vector_mul_lid) + apply (simp add: matrix_vector_mul_assoc AA' matrix_vector_mul_lid) + done from linear_surjective_isomorphism[OF matrix_vector_mul_linear sA] obtain f' :: "real ^'n \ real ^'n" where f': "linear f'" "\x. f' (A *v x) = x" "\x. A *v f' x = x" by blast have th: "matrix f' ** A = mat 1" - by (simp add: matrix_eq matrix_works[OF f'(1)] matrix_vector_mul_assoc[symmetric] matrix_vector_mul_lid f'(2)[rule_format]) + by (simp add: matrix_eq matrix_works[OF f'(1)] + matrix_vector_mul_assoc[symmetric] matrix_vector_mul_lid f'(2)[rule_format]) hence "(matrix f' ** A) ** A' = mat 1 ** A'" by simp - hence "matrix f' = A'" by (simp add: matrix_mul_assoc[symmetric] AA' matrix_mul_rid matrix_mul_lid) + hence "matrix f' = A'" + by (simp add: matrix_mul_assoc[symmetric] AA' matrix_mul_rid matrix_mul_lid) hence "matrix f' ** A = A' ** A" by simp - hence "A' ** A = mat 1" by (simp add: th)} + hence "A' ** A = mat 1" by (simp add: th) + } then show ?thesis by blast qed @@ -1058,67 +1138,64 @@ definition "columnvector v = (\ i j. (v$i))" -lemma transpose_columnvector: - "transpose(columnvector v) = rowvector v" +lemma transpose_columnvector: "transpose(columnvector v) = rowvector v" by (simp add: transpose_def rowvector_def columnvector_def vec_eq_iff) lemma transpose_rowvector: "transpose(rowvector v) = columnvector v" by (simp add: transpose_def columnvector_def rowvector_def vec_eq_iff) -lemma dot_rowvector_columnvector: - "columnvector (A *v v) = A ** columnvector v" +lemma dot_rowvector_columnvector: "columnvector (A *v v) = A ** columnvector v" by (vector columnvector_def matrix_matrix_mult_def matrix_vector_mult_def) -lemma dot_matrix_product: "(x::real^'n) \ y = (((rowvector x ::real^'n^1) ** (columnvector y :: real^1^'n))$1)$1" +lemma dot_matrix_product: + "(x::real^'n) \ y = (((rowvector x ::real^'n^1) ** (columnvector y :: real^1^'n))$1)$1" by (vector matrix_matrix_mult_def rowvector_def columnvector_def inner_vec_def) lemma dot_matrix_vector_mul: fixes A B :: "real ^'n ^'n" and x y :: "real ^'n" shows "(A *v x) \ (B *v y) = (((rowvector x :: real^'n^1) ** ((transpose A ** B) ** (columnvector y :: real ^1^'n)))$1)$1" -unfolding dot_matrix_product transpose_columnvector[symmetric] - dot_rowvector_columnvector matrix_transpose_mul matrix_mul_assoc .. + unfolding dot_matrix_product transpose_columnvector[symmetric] + dot_rowvector_columnvector matrix_transpose_mul matrix_mul_assoc .. lemma infnorm_cart:"infnorm (x::real^'n) = Sup {abs(x$i) |i. i\ (UNIV :: 'n set)}" unfolding infnorm_def apply(rule arg_cong[where f=Sup]) apply safe apply(rule_tac x="\ i" in exI) defer - apply(rule_tac x="\' i" in exI) unfolding nth_conv_component using pi'_range by auto + apply(rule_tac x="\' i" in exI) + unfolding nth_conv_component + using pi'_range apply auto + done -lemma infnorm_set_image_cart: - "{abs(x$i) |i. i\ (UNIV :: _ set)} = +lemma infnorm_set_image_cart: "{abs(x$i) |i. i\ (UNIV :: _ set)} = (\i. abs(x$i)) ` (UNIV)" by blast lemma infnorm_set_lemma_cart: - shows "finite {abs((x::'a::abs ^'n)$i) |i. i\ (UNIV :: 'n set)}" - and "{abs(x$i) |i. i\ (UNIV :: 'n::finite set)} \ {}" - unfolding infnorm_set_image_cart - by auto + "finite {abs((x::'a::abs ^'n)$i) |i. i\ (UNIV :: 'n set)}" + "{abs(x$i) |i. i\ (UNIV :: 'n::finite set)} \ {}" + unfolding infnorm_set_image_cart by auto -lemma component_le_infnorm_cart: - shows "\x$i\ \ infnorm (x::real^'n)" +lemma component_le_infnorm_cart: "\x$i\ \ infnorm (x::real^'n)" unfolding nth_conv_component using component_le_infnorm[of x] . -lemma continuous_component: - shows "continuous F f \ continuous F (\x. f x $ i)" +lemma continuous_component: "continuous F f \ continuous F (\x. f x $ i)" unfolding continuous_def by (rule tendsto_vec_nth) -lemma continuous_on_component: - shows "continuous_on s f \ continuous_on s (\x. f x $ i)" +lemma continuous_on_component: "continuous_on s f \ continuous_on s (\x. f x $ i)" unfolding continuous_on_def by (fast intro: tendsto_vec_nth) lemma closed_positive_orthant: "closed {x::real^'n. \i. 0 \x$i}" by (simp add: Collect_all_eq closed_INT closed_Collect_le) lemma bounded_component_cart: "bounded s \ bounded ((\x. x $ i) ` s)" -unfolding bounded_def -apply clarify -apply (rule_tac x="x $ i" in exI) -apply (rule_tac x="e" in exI) -apply clarify -apply (rule order_trans [OF dist_vec_nth_le], simp) -done + unfolding bounded_def + apply clarify + apply (rule_tac x="x $ i" in exI) + apply (rule_tac x="e" in exI) + apply clarify + apply (rule order_trans [OF dist_vec_nth_le], simp) + done lemma compact_lemma_cart: fixes f :: "nat \ 'a::heine_borel ^ 'n" @@ -1127,24 +1204,35 @@ \l r. subseq r \ (\e>0. eventually (\n. \i\d. dist (f (r n) $ i) (l $ i) < e) sequentially)" proof - fix d::"'n set" have "finite d" by simp + fix d :: "'n set" + have "finite d" by simp thus "\l::'a ^ 'n. \r. subseq r \ (\e>0. eventually (\n. \i\d. dist (f (r n) $ i) (l $ i) < e) sequentially)" - proof(induct d) case empty thus ?case unfolding subseq_def by auto - next case (insert k d) - have s': "bounded ((\x. x $ k) ` s)" using `bounded s` by (rule bounded_component_cart) - obtain l1::"'a^'n" and r1 where r1:"subseq r1" and lr1:"\e>0. eventually (\n. \i\d. dist (f (r1 n) $ i) (l1 $ i) < e) sequentially" + proof (induct d) + case empty + thus ?case unfolding subseq_def by auto + next + case (insert k d) + have s': "bounded ((\x. x $ k) ` s)" + using `bounded s` by (rule bounded_component_cart) + obtain l1::"'a^'n" and r1 where r1:"subseq r1" + and lr1:"\e>0. eventually (\n. \i\d. dist (f (r1 n) $ i) (l1 $ i) < e) sequentially" using insert(3) by auto - have f': "\n. f (r1 n) $ k \ (\x. x $ k) ` s" using `\n. f n \ s` by simp - obtain l2 r2 where r2:"subseq r2" and lr2:"((\i. f (r1 (r2 i)) $ k) ---> l2) sequentially" + have f': "\n. f (r1 n) $ k \ (\x. x $ k) ` s" + using `\n. f n \ s` by simp + obtain l2 r2 where r2: "subseq r2" + and lr2: "((\i. f (r1 (r2 i)) $ k) ---> l2) sequentially" using bounded_imp_convergent_subsequence[OF s' f'] unfolding o_def by auto - def r \ "r1 \ r2" have r:"subseq r" + def r \ "r1 \ r2" + have r: "subseq r" using r1 and r2 unfolding r_def o_def subseq_def by auto moreover def l \ "(\ i. if i = k then l2 else l1$i)::'a^'n" - { fix e::real assume "e>0" - from lr1 `e>0` have N1:"eventually (\n. \i\d. dist (f (r1 n) $ i) (l1 $ i) < e) sequentially" by blast - from lr2 `e>0` have N2:"eventually (\n. dist (f (r1 (r2 n)) $ k) l2 < e) sequentially" by (rule tendstoD) + { fix e :: real assume "e > 0" + from lr1 `e>0` have N1:"eventually (\n. \i\d. dist (f (r1 n) $ i) (l1 $ i) < e) sequentially" + by blast + from lr2 `e>0` have N2:"eventually (\n. dist (f (r1 (r2 n)) $ k) l2 < e) sequentially" + by (rule tendstoD) from r2 N1 have N1': "eventually (\n. \i\d. dist (f (r1 (r2 n)) $ i) (l1 $ i) < e) sequentially" by (rule eventually_subseq) have "eventually (\n. \i\(insert k d). dist (f (r n) $ i) (l $ i) < e) sequentially" @@ -1159,16 +1247,17 @@ fix s :: "('a ^ 'b) set" and f :: "nat \ 'a ^ 'b" assume s: "bounded s" and f: "\n. f n \ s" then obtain l r where r: "subseq r" - and l: "\e>0. eventually (\n. \i\UNIV. dist (f (r n) $ i) (l $ i) < e) sequentially" + and l: "\e>0. eventually (\n. \i\UNIV. dist (f (r n) $ i) (l $ i) < e) sequentially" using compact_lemma_cart [OF s f] by blast let ?d = "UNIV::'b set" { fix e::real assume "e>0" hence "0 < e / (real_of_nat (card ?d))" - using zero_less_card_finite using divide_pos_pos[of e, of "real_of_nat (card ?d)"] by auto + using zero_less_card_finite divide_pos_pos[of e, of "real_of_nat (card ?d)"] by auto with l have "eventually (\n. \i. dist (f (r n) $ i) (l $ i) < e / (real_of_nat (card ?d))) sequentially" by simp moreover - { fix n assume n: "\i. dist (f (r n) $ i) (l $ i) < e / (real_of_nat (card ?d))" + { fix n + assume n: "\i. dist (f (r n) $ i) (l $ i) < e / (real_of_nat (card ?d))" have "dist (f (r n)) l \ (\i\?d. dist (f (r n) $ i) (l $ i))" unfolding dist_vec_def using zero_le_dist by (rule setL2_le_setsum) also have "\ < (\i\?d. e / (real_of_nat (card ?d)))" @@ -1178,28 +1267,31 @@ ultimately have "eventually (\n. dist (f (r n)) l < e) sequentially" by (rule eventually_elim1) } - hence *:"((f \ r) ---> l) sequentially" unfolding o_def tendsto_iff by simp + hence "((f \ r) ---> l) sequentially" unfolding o_def tendsto_iff by simp with r show "\l r. subseq r \ ((f \ r) ---> l) sequentially" by auto qed -lemma interval_cart: fixes a :: "'a::ord^'n" shows - "{a <..< b} = {x::'a^'n. \i. a$i < x$i \ x$i < b$i}" and - "{a .. b} = {x::'a^'n. \i. a$i \ x$i \ x$i \ b$i}" +lemma interval_cart: + fixes a :: "'a::ord^'n" + shows "{a <..< b} = {x::'a^'n. \i. a$i < x$i \ x$i < b$i}" + and "{a .. b} = {x::'a^'n. \i. a$i \ x$i \ x$i \ b$i}" by (auto simp add: set_eq_iff less_vec_def less_eq_vec_def) -lemma mem_interval_cart: fixes a :: "'a::ord^'n" shows - "x \ {a<.. (\i. a$i < x$i \ x$i < b$i)" - "x \ {a .. b} \ (\i. a$i \ x$i \ x$i \ b$i)" - using interval_cart[of a b] by(auto simp add: set_eq_iff less_vec_def less_eq_vec_def) +lemma mem_interval_cart: + fixes a :: "'a::ord^'n" + shows "x \ {a<.. (\i. a$i < x$i \ x$i < b$i)" + and "x \ {a .. b} \ (\i. a$i \ x$i \ x$i \ b$i)" + using interval_cart[of a b] by (auto simp add: set_eq_iff less_vec_def less_eq_vec_def) -lemma interval_eq_empty_cart: fixes a :: "real^'n" shows - "({a <..< b} = {} \ (\i. b$i \ a$i))" (is ?th1) and - "({a .. b} = {} \ (\i. b$i < a$i))" (is ?th2) -proof- +lemma interval_eq_empty_cart: + fixes a :: "real^'n" + shows "({a <..< b} = {} \ (\i. b$i \ a$i))" (is ?th1) + and "({a .. b} = {} \ (\i. b$i < a$i))" (is ?th2) +proof - { fix i x assume as:"b$i \ a$i" and x:"x\{a <..< b}" hence "a $ i < x $ i \ x $ i < b $ i" unfolding mem_interval_cart by auto hence "a$i < b$i" by auto - hence False using as by auto } + hence False using as by auto } moreover { assume as:"\i. \ (b$i \ a$i)" let ?x = "(1/2) *\<^sub>R (a + b)" @@ -1207,14 +1299,14 @@ have "a$i < b$i" using as[THEN spec[where x=i]] by auto hence "a$i < ((1/2) *\<^sub>R (a+b)) $ i" "((1/2) *\<^sub>R (a+b)) $ i < b$i" unfolding vector_smult_component and vector_add_component - by auto } - hence "{a <..< b} \ {}" using mem_interval_cart(1)[of "?x" a b] by auto } + by auto } + hence "{a <..< b} \ {}" using mem_interval_cart(1)[of "?x" a b] by auto } ultimately show ?th1 by blast { fix i x assume as:"b$i < a$i" and x:"x\{a .. b}" hence "a $ i \ x $ i \ x $ i \ b $ i" unfolding mem_interval_cart by auto hence "a$i \ b$i" by auto - hence False using as by auto } + hence False using as by auto } moreover { assume as:"\i. \ (b$i < a$i)" let ?x = "(1/2) *\<^sub>R (a + b)" @@ -1222,37 +1314,41 @@ have "a$i \ b$i" using as[THEN spec[where x=i]] by auto hence "a$i \ ((1/2) *\<^sub>R (a+b)) $ i" "((1/2) *\<^sub>R (a+b)) $ i \ b$i" unfolding vector_smult_component and vector_add_component - by auto } + by auto } hence "{a .. b} \ {}" using mem_interval_cart(2)[of "?x" a b] by auto } ultimately show ?th2 by blast qed -lemma interval_ne_empty_cart: fixes a :: "real^'n" shows - "{a .. b} \ {} \ (\i. a$i \ b$i)" and - "{a <..< b} \ {} \ (\i. a$i < b$i)" +lemma interval_ne_empty_cart: + fixes a :: "real^'n" + shows "{a .. b} \ {} \ (\i. a$i \ b$i)" + and "{a <..< b} \ {} \ (\i. a$i < b$i)" unfolding interval_eq_empty_cart[of a b] by (auto simp add: not_less not_le) (* BH: Why doesn't just "auto" work here? *) -lemma subset_interval_imp_cart: fixes a :: "real^'n" shows - "(\i. a$i \ c$i \ d$i \ b$i) \ {c .. d} \ {a .. b}" and - "(\i. a$i < c$i \ d$i < b$i) \ {c .. d} \ {a<..i. a$i \ c$i \ d$i \ b$i) \ {c<.. {a .. b}" and - "(\i. a$i \ c$i \ d$i \ b$i) \ {c<.. {a<..i. a$i \ c$i \ d$i \ b$i) \ {c .. d} \ {a .. b}" + and "(\i. a$i < c$i \ d$i < b$i) \ {c .. d} \ {a<..i. a$i \ c$i \ d$i \ b$i) \ {c<.. {a .. b}" + and "(\i. a$i \ c$i \ d$i \ b$i) \ {c<.. {a<.. {a<.. {a<.. {a .. b}" -proof(simp add: subset_eq, rule) +lemma interval_open_subset_closed_cart: + fixes a :: "'a::preorder^'n" + shows "{a<.. {a .. b}" +proof (simp add: subset_eq, rule) fix x - assume x:"x \{a<..{a<.. x $ i" using x order_less_imp_le[of "a$i" "x$i"] @@ -1269,105 +1365,123 @@ by(simp add: set_eq_iff less_vec_def less_eq_vec_def vec_eq_iff) qed -lemma subset_interval_cart: fixes a :: "real^'n" shows - "{c .. d} \ {a .. b} \ (\i. c$i \ d$i) --> (\i. a$i \ c$i \ d$i \ b$i)" (is ?th1) and - "{c .. d} \ {a<.. (\i. c$i \ d$i) --> (\i. a$i < c$i \ d$i < b$i)" (is ?th2) and - "{c<.. {a .. b} \ (\i. c$i < d$i) --> (\i. a$i \ c$i \ d$i \ b$i)" (is ?th3) and - "{c<.. {a<.. (\i. c$i < d$i) --> (\i. a$i \ c$i \ d$i \ b$i)" (is ?th4) +lemma subset_interval_cart: + fixes a :: "real^'n" + shows "{c .. d} \ {a .. b} \ (\i. c$i \ d$i) --> (\i. a$i \ c$i \ d$i \ b$i)" (is ?th1) + and "{c .. d} \ {a<.. (\i. c$i \ d$i) --> (\i. a$i < c$i \ d$i < b$i)" (is ?th2) + and "{c<.. {a .. b} \ (\i. c$i < d$i) --> (\i. a$i \ c$i \ d$i \ b$i)" (is ?th3) + and "{c<.. {a<.. (\i. c$i < d$i) --> (\i. a$i \ c$i \ d$i \ b$i)" (is ?th4) using subset_interval[of c d a b] by (simp_all add: cart_simps real_euclidean_nth) -lemma disjoint_interval_cart: fixes a::"real^'n" shows - "{a .. b} \ {c .. d} = {} \ (\i. (b$i < a$i \ d$i < c$i \ b$i < c$i \ d$i < a$i))" (is ?th1) and - "{a .. b} \ {c<.. (\i. (b$i < a$i \ d$i \ c$i \ b$i \ c$i \ d$i \ a$i))" (is ?th2) and - "{a<.. {c .. d} = {} \ (\i. (b$i \ a$i \ d$i < c$i \ b$i \ c$i \ d$i \ a$i))" (is ?th3) and - "{a<.. {c<.. (\i. (b$i \ a$i \ d$i \ c$i \ b$i \ c$i \ d$i \ a$i))" (is ?th4) +lemma disjoint_interval_cart: + fixes a::"real^'n" + shows "{a .. b} \ {c .. d} = {} \ (\i. (b$i < a$i \ d$i < c$i \ b$i < c$i \ d$i < a$i))" (is ?th1) + and "{a .. b} \ {c<.. (\i. (b$i < a$i \ d$i \ c$i \ b$i \ c$i \ d$i \ a$i))" (is ?th2) + and "{a<.. {c .. d} = {} \ (\i. (b$i \ a$i \ d$i < c$i \ b$i \ c$i \ d$i \ a$i))" (is ?th3) + and "{a<.. {c<.. (\i. (b$i \ a$i \ d$i \ c$i \ b$i \ c$i \ d$i \ a$i))" (is ?th4) using disjoint_interval[of a b c d] by (simp_all add: cart_simps real_euclidean_nth) -lemma inter_interval_cart: fixes a :: "'a::linorder^'n" shows - "{a .. b} \ {c .. d} = {(\ i. max (a$i) (c$i)) .. (\ i. min (b$i) (d$i))}" +lemma inter_interval_cart: + fixes a :: "'a::linorder^'n" + shows "{a .. b} \ {c .. d} = {(\ i. max (a$i) (c$i)) .. (\ i. min (b$i) (d$i))}" unfolding set_eq_iff and Int_iff and mem_interval_cart by auto -lemma closed_interval_left_cart: fixes b::"real^'n" +lemma closed_interval_left_cart: + fixes b :: "real^'n" shows "closed {x::real^'n. \i. x$i \ b$i}" by (simp add: Collect_all_eq closed_INT closed_Collect_le) -lemma closed_interval_right_cart: fixes a::"real^'n" +lemma closed_interval_right_cart: + fixes a::"real^'n" shows "closed {x::real^'n. \i. a$i \ x$i}" by (simp add: Collect_all_eq closed_INT closed_Collect_le) -lemma is_interval_cart:"is_interval (s::(real^'n) set) \ - (\a\s. \b\s. \x. (\i. ((a$i \ x$i \ x$i \ b$i) \ (b$i \ x$i \ x$i \ a$i))) \ x \ s)" - unfolding is_interval_def Ball_def by (simp add: cart_simps real_euclidean_nth) +lemma is_interval_cart: + "is_interval (s::(real^'n) set) \ + (\a\s. \b\s. \x. (\i. ((a$i \ x$i \ x$i \ b$i) \ (b$i \ x$i \ x$i \ a$i))) \ x \ s)" + by (simp add: is_interval_def Ball_def cart_simps real_euclidean_nth) -lemma closed_halfspace_component_le_cart: - shows "closed {x::real^'n. x$i \ a}" +lemma closed_halfspace_component_le_cart: "closed {x::real^'n. x$i \ a}" by (simp add: closed_Collect_le) -lemma closed_halfspace_component_ge_cart: - shows "closed {x::real^'n. x$i \ a}" +lemma closed_halfspace_component_ge_cart: "closed {x::real^'n. x$i \ a}" by (simp add: closed_Collect_le) -lemma open_halfspace_component_lt_cart: - shows "open {x::real^'n. x$i < a}" +lemma open_halfspace_component_lt_cart: "open {x::real^'n. x$i < a}" + by (simp add: open_Collect_less) + +lemma open_halfspace_component_gt_cart: "open {x::real^'n. x$i > a}" by (simp add: open_Collect_less) -lemma open_halfspace_component_gt_cart: - shows "open {x::real^'n. x$i > a}" - by (simp add: open_Collect_less) - -lemma Lim_component_le_cart: fixes f :: "'a \ real^'n" +lemma Lim_component_le_cart: + fixes f :: "'a \ real^'n" assumes "(f ---> l) net" "\ (trivial_limit net)" "eventually (\x. f(x)$i \ b) net" shows "l$i \ b" -proof- - { fix x have "x \ {x::real^'n. inner (cart_basis i) x \ b} \ x$i \ b" unfolding inner_basis by auto } note * = this - show ?thesis using Lim_in_closed_set[of "{x. inner (cart_basis i) x \ b}" f net l] unfolding * +proof - + { fix x + have "x \ {x::real^'n. inner (cart_basis i) x \ b} \ x$i \ b" + unfolding inner_basis by auto } + then show ?thesis using Lim_in_closed_set[of "{x. inner (cart_basis i) x \ b}" f net l] using closed_halfspace_le[of "(cart_basis i)::real^'n" b] and assms(1,2,3) by auto qed -lemma Lim_component_ge_cart: fixes f :: "'a \ real^'n" +lemma Lim_component_ge_cart: + fixes f :: "'a \ real^'n" assumes "(f ---> l) net" "\ (trivial_limit net)" "eventually (\x. b \ (f x)$i) net" shows "b \ l$i" -proof- - { fix x have "x \ {x::real^'n. inner (cart_basis i) x \ b} \ x$i \ b" unfolding inner_basis by auto } note * = this - show ?thesis using Lim_in_closed_set[of "{x. inner (cart_basis i) x \ b}" f net l] unfolding * +proof - + { fix x + have "x \ {x::real^'n. inner (cart_basis i) x \ b} \ x$i \ b" + unfolding inner_basis by auto } + then show ?thesis using Lim_in_closed_set[of "{x. inner (cart_basis i) x \ b}" f net l] using closed_halfspace_ge[of b "(cart_basis i)::real^'n"] and assms(1,2,3) by auto qed -lemma Lim_component_eq_cart: fixes f :: "'a \ real^'n" - assumes net:"(f ---> l) net" "~(trivial_limit net)" and ev:"eventually (\x. f(x)$i = b) net" +lemma Lim_component_eq_cart: + fixes f :: "'a \ real^'n" + assumes net: "(f ---> l) net" "~(trivial_limit net)" and ev:"eventually (\x. f(x)$i = b) net" shows "l$i = b" - using ev[unfolded order_eq_iff eventually_conj_iff] using Lim_component_ge_cart[OF net, of b i] and + using ev[unfolded order_eq_iff eventually_conj_iff] and + Lim_component_ge_cart[OF net, of b i] and Lim_component_le_cart[OF net, of i b] by auto -lemma connected_ivt_component_cart: fixes x::"real^'n" shows - "connected s \ x \ s \ y \ s \ x$k \ a \ a \ y$k \ (\z\s. z$k = a)" - using connected_ivt_hyperplane[of s x y "(cart_basis k)::real^'n" a] by (auto simp add: inner_basis) +lemma connected_ivt_component_cart: + fixes x :: "real^'n" + shows "connected s \ x \ s \ y \ s \ x$k \ a \ a \ y$k \ (\z\s. z$k = a)" + using connected_ivt_hyperplane[of s x y "(cart_basis k)::real^'n" a] + by (auto simp add: inner_basis) -lemma subspace_substandard_cart: - "subspace {x::real^_. (\i. P i \ x$i = 0)}" +lemma subspace_substandard_cart: "subspace {x::real^_. (\i. P i \ x$i = 0)}" unfolding subspace_def by auto lemma closed_substandard_cart: "closed {x::'a::real_normed_vector ^ 'n. \i. P i \ x$i = 0}" -proof- +proof - { fix i::'n have "closed {x::'a ^ 'n. P i \ x$i = 0}" - by (cases "P i", simp_all add: closed_Collect_eq) } + by (cases "P i") (simp_all add: closed_Collect_eq) } thus ?thesis unfolding Collect_all_eq by (simp add: closed_INT) qed -lemma dim_substandard_cart: - shows "dim {x::real^'n. \i. i \ d \ x$i = 0} = card d" (is "dim ?A = _") -proof- have *:"{x. \i \' ` d \ x $$ i = 0} = - {x::real^'n. \i. i \ d \ x$i = 0}"apply safe - apply(erule_tac x="\' i" in allE) defer - apply(erule_tac x="\ i" in allE) - unfolding image_iff real_euclidean_nth[symmetric] by (auto simp: pi'_inj[THEN inj_eq]) - have " \' ` d \ {..'" d] using pi'_inj unfolding inj_on_def by auto +lemma dim_substandard_cart: "dim {x::real^'n. \i. i \ d \ x$i = 0} = card d" + (is "dim ?A = _") +proof - + have *: "{x. \i \' ` d \ x $$ i = 0} = + {x::real^'n. \i. i \ d \ x$i = 0}" + apply safe + apply (erule_tac x="\' i" in allE) defer + apply (erule_tac x="\ i" in allE) + unfolding image_iff real_euclidean_nth[symmetric] + apply (auto simp: pi'_inj[THEN inj_eq]) + done + have " \' ` d \ {..'" d] using pi'_inj unfolding inj_on_def + by auto qed lemma affinity_inverses: @@ -1375,8 +1489,9 @@ shows "(\x. m *s x + c) o (\x. inverse(m) *s x + (-(inverse(m) *s c))) = id" "(\x. inverse(m) *s x + (-(inverse(m) *s c))) o (\x. m *s x + c) = id" using m0 -apply (auto simp add: fun_eq_iff vector_add_ldistrib) -by (simp_all add: vector_smult_lneg[symmetric] vector_smult_assoc vector_sneg_minus1[symmetric]) + apply (auto simp add: fun_eq_iff vector_add_ldistrib) + apply (simp_all add: vector_smult_lneg[symmetric] vector_smult_assoc vector_sneg_minus1[symmetric]) + done lemma vector_affinity_eq: assumes m0: "(m::'a::field) \ 0" @@ -1394,25 +1509,30 @@ qed lemma vector_eq_affinity: - "(m::'a::field) \ 0 ==> (y = m *s x + c \ inverse(m) *s y + -(inverse(m) *s c) = x)" + "(m::'a::field) \ 0 ==> (y = m *s x + c \ inverse(m) *s y + -(inverse(m) *s c) = x)" using vector_affinity_eq[where m=m and x=x and y=y and c=c] by metis lemma const_vector_cart:"((\ i. d)::real^'n) = (\\ i. d)" apply(subst euclidean_eq) -proof safe case goal1 - hence *:"(basis i::real^'n) = cart_basis (\ i)" - unfolding basis_real_n[THEN sym] by auto +proof safe + case goal1 + hence *: "(basis i::real^'n) = cart_basis (\ i)" + unfolding basis_real_n[symmetric] by auto have "((\ i. d)::real^'n) $$ i = d" unfolding euclidean_component_def * unfolding dot_basis by auto thus ?case using goal1 by auto qed + subsection "Convex Euclidean Space" lemma Cart_1:"(1::real^'n) = (\\ i. 1)" apply(subst euclidean_eq) -proof safe case goal1 thus ?case using nth_conv_component[THEN sym,where i1="\ i" and x1="1::real^'n"] by auto +proof safe + case goal1 + thus ?case + using nth_conv_component[THEN sym,where i1="\ i" and x1="1::real^'n"] by auto qed declare vector_add_ldistrib[simp] vector_ssub_ldistrib[simp] vector_smult_assoc[simp] vector_smult_rneg[simp] @@ -1433,21 +1553,32 @@ unfolding Cart_1 unit_interval_convex_hull[where 'a="real^'n"] apply(rule arg_cong[where f="\x. convex hull x"]) apply(rule set_eqI) unfolding mem_Collect_eq apply safe apply(erule_tac x="\' i" in allE) unfolding nth_conv_component defer - apply(erule_tac x="\ i" in allE) by auto + apply(erule_tac x="\ i" in allE) + apply auto + done lemma cube_convex_hull_cart: - assumes "0 < d" obtains s::"(real^'n) set" where "finite s" "{x - (\ i. d) .. x + (\ i. d)} = convex hull s" -proof- from cube_convex_hull[OF assms, where 'a="real^'n" and x=x] guess s . note s=this - show thesis apply(rule that[OF s(1)]) unfolding s(2)[THEN sym] const_vector_cart .. + assumes "0 < d" + obtains s::"(real^'n) set" + where "finite s" "{x - (\ i. d) .. x + (\ i. d)} = convex hull s" +proof - + obtain s where s: "finite s" "{x - (\\ i. d)..x + (\\ i. d)} = convex hull s" + by (rule cube_convex_hull [OF assms]) + show thesis + apply(rule that[OF s(1)]) unfolding s(2)[symmetric] const_vector_cart .. qed lemma std_simplex_cart: "(insert (0::real^'n) { cart_basis i | i. i\UNIV}) = - (insert 0 { basis i | i. i i. u$i + (x$i - a$i) / (b$i - a$i) * (v$i - u$i))::real^'n)" unfolding interval_bij_def apply(rule ext)+ apply safe unfolding vec_eq_iff vec_lambda_beta unfolding nth_conv_component - apply rule apply(subst euclidean_lambda_beta) using pi'_range by auto + apply rule + apply (subst euclidean_lambda_beta) + using pi'_range apply auto + done lemma interval_bij_affine_cart: "interval_bij (a,b) (u,v) = (\x. (\ i. (v$i - u$i) / (b$i - a$i) * x$i) + (\ i. u$i - (v$i - u$i) / (b$i - a$i) * a$i)::real^'n)" - apply rule unfolding vec_eq_iff interval_bij_cart vector_component_simps - by(auto simp add: field_simps add_divide_distrib[THEN sym]) + apply rule + unfolding vec_eq_iff interval_bij_cart vector_component_simps + apply (auto simp add: field_simps add_divide_distrib[symmetric]) + done + subsection "Derivative" -lemma has_derivative_vmul_component_cart: fixes c::"real^'a \ real^'b" and v::"real^'c" +lemma has_derivative_vmul_component_cart: + fixes c :: "real^'a \ real^'b" and v :: "real^'c" assumes "(c has_derivative c') net" shows "((\x. c(x)$k *\<^sub>R v) has_derivative (\x. (c' x)$k *\<^sub>R v)) net" - unfolding nth_conv_component - by (intro has_derivative_intros assms) + unfolding nth_conv_component by (intro has_derivative_intros assms) -lemma differentiable_at_imp_differentiable_on: "(\x\(s::(real^'n) set). f differentiable at x) \ f differentiable_on s" +lemma differentiable_at_imp_differentiable_on: + "(\x\(s::(real^'n) set). f differentiable at x) \ f differentiable_on s" unfolding differentiable_on_def by(auto intro!: differentiable_at_withinI) definition "jacobian f net = matrix(frechet_derivative f net)" -lemma jacobian_works: "(f::(real^'a) \ (real^'b)) differentiable net \ (f has_derivative (\h. (jacobian f net) *v h)) net" - apply rule unfolding jacobian_def apply(simp only: matrix_works[OF linear_frechet_derivative]) defer - apply(rule differentiableI) apply assumption unfolding frechet_derivative_works by assumption +lemma jacobian_works: + "(f::(real^'a) \ (real^'b)) differentiable net \ + (f has_derivative (\h. (jacobian f net) *v h)) net" + apply rule + unfolding jacobian_def + apply (simp only: matrix_works[OF linear_frechet_derivative]) defer + apply (rule differentiableI) + apply assumption + unfolding frechet_derivative_works + apply assumption + done -subsection {* Component of the differential must be zero if it exists at a local *) -(* maximum or minimum for that corresponding component. *} -lemma differential_zero_maxmin_component: fixes f::"real^'a \ real^'b" - assumes "0 < e" "((\y \ ball x e. (f y)$k \ (f x)$k) \ (\y\ball x e. (f x)$k \ (f y)$k))" - "f differentiable (at x)" shows "jacobian f (at x) $ k = 0" -(* FIXME: reuse proof of generic differential_zero_maxmin_component*) +subsection {* Component of the differential must be zero if it exists at a local + maximum or minimum for that corresponding component. *} -proof(rule ccontr) - def D \ "jacobian f (at x)" assume "jacobian f (at x) $ k \ 0" +lemma differential_zero_maxmin_component: + fixes f::"real^'a \ real^'b" + assumes "0 < e" "((\y \ ball x e. (f y)$k \ (f x)$k) \ (\y\ball x e. (f x)$k \ (f y)$k))" + "f differentiable (at x)" shows "jacobian f (at x) $ k = 0" +(* FIXME: reuse proof of generic differential_zero_maxmin_component*) +proof (rule ccontr) + def D \ "jacobian f (at x)" + assume "jacobian f (at x) $ k \ 0" then obtain j where j:"D$k$j \ 0" unfolding vec_eq_iff D_def by auto - hence *:"abs (jacobian f (at x) $ k $ j) / 2 > 0" unfolding D_def by auto + hence *: "abs (jacobian f (at x) $ k $ j) / 2 > 0" + unfolding D_def by auto note as = assms(3)[unfolded jacobian_works has_derivative_at_alt] guess e' using as[THEN conjunct2,rule_format,OF *] .. note e' = this guess d using real_lbound_gt_zero[OF assms(1) e'[THEN conjunct1]] .. note d = this - { fix c assume "abs c \ d" - hence *:"norm (x + c *\<^sub>R cart_basis j - x) < e'" using norm_basis[of j] d by auto - have "\(f (x + c *\<^sub>R cart_basis j) - f x - D *v (c *\<^sub>R cart_basis j)) $ k\ \ norm (f (x + c *\<^sub>R cart_basis j) - f x - D *v (c *\<^sub>R cart_basis j))" - by(rule component_le_norm_cart) - also have "\ \ \D $ k $ j\ / 2 * \c\" using e'[THEN conjunct2,rule_format,OF *] and norm_basis[of j] unfolding D_def[symmetric] by auto - finally have "\(f (x + c *\<^sub>R cart_basis j) - f x - D *v (c *\<^sub>R cart_basis j)) $ k\ \ \D $ k $ j\ / 2 * \c\" by simp - hence "\f (x + c *\<^sub>R cart_basis j) $ k - f x $ k - c * D $ k $ j\ \ \D $ k $ j\ / 2 * \c\" - unfolding vector_component_simps matrix_vector_mul_component unfolding smult_conv_scaleR[symmetric] - unfolding inner_simps dot_basis smult_conv_scaleR by simp } note * = this + { fix c + assume "abs c \ d" + hence *:"norm (x + c *\<^sub>R cart_basis j - x) < e'" + using norm_basis[of j] d by auto + have "\(f (x + c *\<^sub>R cart_basis j) - f x - D *v (c *\<^sub>R cart_basis j)) $ k\ \ + norm (f (x + c *\<^sub>R cart_basis j) - f x - D *v (c *\<^sub>R cart_basis j))" + by (rule component_le_norm_cart) + also have "\ \ \D $ k $ j\ / 2 * \c\" + using e'[THEN conjunct2,rule_format,OF *] and norm_basis[of j] + unfolding D_def[symmetric] by auto + finally + have "\(f (x + c *\<^sub>R cart_basis j) - f x - D *v (c *\<^sub>R cart_basis j)) $ k\ \ + \D $ k $ j\ / 2 * \c\" by simp + hence "\f (x + c *\<^sub>R cart_basis j) $ k - f x $ k - c * D $ k $ j\ \ + \D $ k $ j\ / 2 * \c\" + unfolding vector_component_simps matrix_vector_mul_component + unfolding smult_conv_scaleR[symmetric] + unfolding inner_simps dot_basis smult_conv_scaleR by simp + } note * = this have "x + d *\<^sub>R cart_basis j \ ball x e" "x - d *\<^sub>R cart_basis j \ ball x e" unfolding mem_ball dist_norm using norm_basis[of j] d by auto - hence **:"((f (x - d *\<^sub>R cart_basis j))$k \ (f x)$k \ (f (x + d *\<^sub>R cart_basis j))$k \ (f x)$k) \ - ((f (x - d *\<^sub>R cart_basis j))$k \ (f x)$k \ (f (x + d *\<^sub>R cart_basis j))$k \ (f x)$k)" using assms(2) by auto - have ***:"\y y1 y2 d dx::real. (y1\y\y2\y) \ (y\y1\y\y2) \ d < abs dx \ abs(y1 - y - - dx) \ d \ (abs (y2 - y - dx) \ d) \ False" by arith - show False apply(rule ***[OF **, where dx="d * D $ k $ j" and d="\D $ k $ j\ / 2 * \d\"]) - using *[of "-d"] and *[of d] and d[THEN conjunct1] and j unfolding mult_minus_left - unfolding abs_mult diff_minus_eq_add scaleR_minus_left unfolding algebra_simps by (auto intro: mult_pos_pos) + hence **: "((f (x - d *\<^sub>R cart_basis j))$k \ (f x)$k \ (f (x + d *\<^sub>R cart_basis j))$k \ (f x)$k) \ + ((f (x - d *\<^sub>R cart_basis j))$k \ (f x)$k \ (f (x + d *\<^sub>R cart_basis j))$k \ (f x)$k)" + using assms(2) by auto + have ***: "\y y1 y2 d dx::real. (y1\y\y2\y) \ (y\y1\y\y2) \ + d < abs dx \ abs(y1 - y - - dx) \ d \ (abs (y2 - y - dx) \ d) \ False" by arith + show False + apply (rule ***[OF **, where dx="d * D $ k $ j" and d="\D $ k $ j\ / 2 * \d\"]) + using *[of "-d"] and *[of d] and d[THEN conjunct1] and j + unfolding mult_minus_left + unfolding abs_mult diff_minus_eq_add scaleR_minus_left + unfolding algebra_simps + apply (auto intro: mult_pos_pos) + done qed + subsection {* Lemmas for working on @{typ "real^1"} *} lemma forall_1[simp]: "(\i::1. P i) \ P 1" - by (metis num1_eq_iff) + by (metis (full_types) num1_eq_iff) lemma ex_1[simp]: "(\x::1. P x) \ P 1" - by auto (metis num1_eq_iff) + by auto (metis (full_types) num1_eq_iff) lemma exhaust_2: - fixes x :: 2 shows "x = 1 \ x = 2" + fixes x :: 2 + shows "x = 1 \ x = 2" proof (induct x) case (of_int z) then have "0 <= z" and "z < 2" by simp_all @@ -1551,7 +1735,8 @@ by (metis exhaust_2) lemma exhaust_3: - fixes x :: 3 shows "x = 1 \ x = 2 \ x = 3" + fixes x :: 3 + shows "x = 1 \ x = 2 \ x = 3" proof (induct x) case (of_int z) then have "0 <= z" and "z < 3" by simp_all @@ -1580,17 +1765,21 @@ lemma setsum_3: "setsum f (UNIV::3 set) = f 1 + f 2 + f 3" unfolding UNIV_3 by (simp add: add_ac) -instantiation num1 :: cart_one begin -instance proof +instantiation num1 :: cart_one +begin + +instance +proof show "CARD(1) = Suc 0" by auto -qed end +qed + +end (* "lift" from 'a to 'a^1 and "drop" from 'a^1 to 'a -- FIXME: potential use of transfer *) abbreviation vec1:: "'a \ 'a ^ 1" where "vec1 x \ vec x" -abbreviation dest_vec1:: "'a ^1 \ 'a" - where "dest_vec1 x \ (x$1)" +abbreviation dest_vec1:: "'a ^1 \ 'a" where "dest_vec1 x \ (x$1)" lemma vec1_dest_vec1[simp]: "vec1(dest_vec1 x) = x" by (simp add: vec_eq_iff) @@ -1604,6 +1793,7 @@ lemma dest_vec1_eq[simp]: "dest_vec1 x = dest_vec1 y \ x = y" by (metis vec1_dest_vec1(1)) + subsection{* The collapse of the general concepts to dimension one. *} lemma vector_one: "(x::'a ^1) = (\ i. (x$1))" @@ -1624,6 +1814,7 @@ lemma dist_real: "dist(x::real ^ 1) y = abs((x$1) - (y$1))" by (auto simp add: norm_real dist_norm) + subsection{* Explicit vector construction from lists. *} definition "vector l = (\ i. foldr (\x f n. fun_upd (f (n+1)) n x) l (\n x. 0) 1 i)" @@ -1672,14 +1863,17 @@ apply (simp add: forall_3) done -lemma range_vec1[simp]:"range vec1 = UNIV" apply(rule set_eqI,rule) unfolding image_iff defer - apply(rule_tac x="dest_vec1 x" in bexI) by auto +lemma range_vec1[simp]:"range vec1 = UNIV" + apply (rule set_eqI,rule) unfolding image_iff defer + apply (rule_tac x="dest_vec1 x" in bexI) + apply auto + done lemma dest_vec1_lambda: "dest_vec1(\ i. x i) = x 1" - by (simp) + by simp lemma dest_vec1_vec: "dest_vec1(vec x) = x" - by (simp) + by simp lemma dest_vec1_sum: assumes fS: "finite S" shows "dest_vec1(setsum f S) = setsum (dest_vec1 o f) S" @@ -1696,13 +1890,19 @@ lemma abs_dest_vec1: "norm x = \dest_vec1 x\" by (metis vec1_dest_vec1(1) norm_vec1) -lemmas vec1_dest_vec1_simps = forall_vec1 vec_add[THEN sym] dist_vec1 vec_sub[THEN sym] vec1_dest_vec1 norm_vec1 vector_smult_component - vec_inj[where 'b=1] vec_cmul[THEN sym] smult_conv_scaleR[THEN sym] o_def dist_real_def real_norm_def +lemmas vec1_dest_vec1_simps = + forall_vec1 vec_add[symmetric] dist_vec1 vec_sub[symmetric] vec1_dest_vec1 norm_vec1 vector_smult_component + vec_inj[where 'b=1] vec_cmul[symmetric] smult_conv_scaleR[symmetric] o_def dist_real_def real_norm_def -lemma bounded_linear_vec1:"bounded_linear (vec1::real\real^1)" +lemma bounded_linear_vec1: "bounded_linear (vec1::real\real^1)" unfolding bounded_linear_def additive_def bounded_linear_axioms_def - unfolding smult_conv_scaleR[THEN sym] unfolding vec1_dest_vec1_simps - apply(rule conjI) defer apply(rule conjI) defer apply(rule_tac x=1 in exI) by auto + unfolding smult_conv_scaleR[symmetric] + unfolding vec1_dest_vec1_simps + apply (rule conjI) defer + apply (rule conjI) defer + apply (rule_tac x=1 in exI) + apply auto + done lemma linear_vmul_dest_vec1: fixes f:: "real^_ \ real^1" @@ -1719,7 +1919,8 @@ apply (auto simp add: vec_eq_iff matrix_vector_mult_def column_def mult_commute) done -lemma linear_to_scalars: assumes lf: "linear (f::real ^'n \ real^1)" +lemma linear_to_scalars: + assumes lf: "linear (f::real ^'n \ real^1)" shows "f = (\x. vec1(row 1 (matrix f) \ x))" apply (rule ext) apply (subst matrix_works[OF lf, symmetric]) @@ -1729,127 +1930,190 @@ lemma dest_vec1_eq_0: "dest_vec1 x = 0 \ x = 0" by (simp add: dest_vec1_eq[symmetric]) -lemma setsum_scalars: assumes fS: "finite S" +lemma setsum_scalars: + assumes fS: "finite S" shows "setsum f S = vec1 (setsum (dest_vec1 o f) S)" unfolding vec_setsum[OF fS] by simp -lemma dest_vec1_wlog_le: "(\(x::'a::linorder ^ 1) y. P x y \ P y x) \ (\x y. dest_vec1 x <= dest_vec1 y ==> P x y) \ P x y" +lemma dest_vec1_wlog_le: + "(\(x::'a::linorder ^ 1) y. P x y \ P y x) + \ (\x y. dest_vec1 x <= dest_vec1 y ==> P x y) \ P x y" apply (cases "dest_vec1 x \ dest_vec1 y") apply simp apply (subgoal_tac "dest_vec1 y \ dest_vec1 x") - apply (auto) + apply auto done text{* Lifting and dropping *} -lemma continuous_on_o_dest_vec1: fixes f::"real \ 'a::real_normed_vector" - assumes "continuous_on {a..b::real} f" shows "continuous_on {vec1 a..vec1 b} (f o dest_vec1)" +lemma continuous_on_o_dest_vec1: + fixes f::"real \ 'a::real_normed_vector" + assumes "continuous_on {a..b::real} f" + shows "continuous_on {vec1 a..vec1 b} (f o dest_vec1)" using assms unfolding continuous_on_iff apply safe - apply(erule_tac x="x$1" in ballE,erule_tac x=e in allE) apply safe - apply(rule_tac x=d in exI) apply safe unfolding o_def dist_real_def dist_real - apply(erule_tac x="dest_vec1 x'" in ballE) by(auto simp add:less_eq_vec_def) + apply (erule_tac x="x$1" in ballE,erule_tac x=e in allE) + apply safe + apply (rule_tac x=d in exI) + apply safe + unfolding o_def dist_real_def dist_real + apply (erule_tac x="dest_vec1 x'" in ballE) + apply (auto simp add:less_eq_vec_def) + done -lemma continuous_on_o_vec1: fixes f::"real^1 \ 'a::real_normed_vector" - assumes "continuous_on {a..b} f" shows "continuous_on {dest_vec1 a..dest_vec1 b} (f o vec1)" - using assms unfolding continuous_on_iff apply safe - apply(erule_tac x="vec x" in ballE,erule_tac x=e in allE) apply safe - apply(rule_tac x=d in exI) apply safe unfolding o_def dist_real_def dist_real - apply(erule_tac x="vec1 x'" in ballE) by(auto simp add:less_eq_vec_def) +lemma continuous_on_o_vec1: + fixes f::"real^1 \ 'a::real_normed_vector" + assumes "continuous_on {a..b} f" + shows "continuous_on {dest_vec1 a..dest_vec1 b} (f o vec1)" + using assms unfolding continuous_on_iff + apply safe + apply (erule_tac x="vec x" in ballE,erule_tac x=e in allE) + apply safe + apply (rule_tac x=d in exI) + apply safe + unfolding o_def dist_real_def dist_real + apply (erule_tac x="vec1 x'" in ballE) + apply (auto simp add:less_eq_vec_def) + done lemma continuous_on_vec1:"continuous_on A (vec1::real\real^1)" - by(rule linear_continuous_on[OF bounded_linear_vec1]) + by (rule linear_continuous_on[OF bounded_linear_vec1]) -lemma mem_interval_1: fixes x :: "real^1" shows - "(x \ {a .. b} \ dest_vec1 a \ dest_vec1 x \ dest_vec1 x \ dest_vec1 b)" - "(x \ {a<.. dest_vec1 a < dest_vec1 x \ dest_vec1 x < dest_vec1 b)" -by(simp_all add: vec_eq_iff less_vec_def less_eq_vec_def) +lemma mem_interval_1: + fixes x :: "real^1" + shows "(x \ {a .. b} \ dest_vec1 a \ dest_vec1 x \ dest_vec1 x \ dest_vec1 b)" + and "(x \ {a<.. dest_vec1 a < dest_vec1 x \ dest_vec1 x < dest_vec1 b)" + by (simp_all add: vec_eq_iff less_vec_def less_eq_vec_def) -lemma vec1_interval:fixes a::"real" shows - "vec1 ` {a .. b} = {vec1 a .. vec1 b}" - "vec1 ` {a<.. {a .. b} ==> x \ {a<.. (x = a) \ (x = b)" - unfolding vec_eq_iff less_vec_def less_eq_vec_def mem_interval_cart by(auto simp del:dest_vec1_eq) +lemma interval_cases_1: + fixes x :: "real^1" + shows "x \ {a .. b} ==> x \ {a<.. (x = a) \ (x = b)" + unfolding vec_eq_iff less_vec_def less_eq_vec_def mem_interval_cart + by (auto simp del:dest_vec1_eq) -lemma in_interval_1: fixes x :: "real^1" shows - "(x \ {a .. b} \ dest_vec1 a \ dest_vec1 x \ dest_vec1 x \ dest_vec1 b) \ - (x \ {a<.. dest_vec1 a < dest_vec1 x \ dest_vec1 x < dest_vec1 b)" - unfolding vec_eq_iff less_vec_def less_eq_vec_def mem_interval_cart by(auto simp del:dest_vec1_eq) +lemma in_interval_1: + fixes x :: "real^1" + shows "(x \ {a .. b} \ dest_vec1 a \ dest_vec1 x \ dest_vec1 x \ dest_vec1 b) \ + (x \ {a<.. dest_vec1 a < dest_vec1 x \ dest_vec1 x < dest_vec1 b)" + unfolding vec_eq_iff less_vec_def less_eq_vec_def mem_interval_cart + by (auto simp del:dest_vec1_eq) -lemma interval_eq_empty_1: fixes a :: "real^1" shows - "{a .. b} = {} \ dest_vec1 b < dest_vec1 a" - "{a<.. dest_vec1 b \ dest_vec1 a" +lemma interval_eq_empty_1: + fixes a :: "real^1" + shows "{a .. b} = {} \ dest_vec1 b < dest_vec1 a" + and "{a<.. dest_vec1 b \ dest_vec1 a" unfolding interval_eq_empty_cart and ex_1 by auto -lemma subset_interval_1: fixes a :: "real^1" shows - "({a .. b} \ {c .. d} \ dest_vec1 b < dest_vec1 a \ - dest_vec1 c \ dest_vec1 a \ dest_vec1 a \ dest_vec1 b \ dest_vec1 b \ dest_vec1 d)" - "({a .. b} \ {c<.. dest_vec1 b < dest_vec1 a \ - dest_vec1 c < dest_vec1 a \ dest_vec1 a \ dest_vec1 b \ dest_vec1 b < dest_vec1 d)" - "({a<.. {c .. d} \ dest_vec1 b \ dest_vec1 a \ - dest_vec1 c \ dest_vec1 a \ dest_vec1 a < dest_vec1 b \ dest_vec1 b \ dest_vec1 d)" - "({a<.. {c<.. dest_vec1 b \ dest_vec1 a \ - dest_vec1 c \ dest_vec1 a \ dest_vec1 a < dest_vec1 b \ dest_vec1 b \ dest_vec1 d)" +lemma subset_interval_1: + fixes a :: "real^1" + shows "({a .. b} \ {c .. d} \ dest_vec1 b < dest_vec1 a \ + dest_vec1 c \ dest_vec1 a \ dest_vec1 a \ dest_vec1 b \ dest_vec1 b \ dest_vec1 d)" + "({a .. b} \ {c<.. dest_vec1 b < dest_vec1 a \ + dest_vec1 c < dest_vec1 a \ dest_vec1 a \ dest_vec1 b \ dest_vec1 b < dest_vec1 d)" + "({a<.. {c .. d} \ dest_vec1 b \ dest_vec1 a \ + dest_vec1 c \ dest_vec1 a \ dest_vec1 a < dest_vec1 b \ dest_vec1 b \ dest_vec1 d)" + "({a<.. {c<.. dest_vec1 b \ dest_vec1 a \ + dest_vec1 c \ dest_vec1 a \ dest_vec1 a < dest_vec1 b \ dest_vec1 b \ dest_vec1 d)" unfolding subset_interval_cart[of a b c d] unfolding forall_1 by auto -lemma eq_interval_1: fixes a :: "real^1" shows - "{a .. b} = {c .. d} \ +lemma eq_interval_1: + fixes a :: "real^1" + shows "{a .. b} = {c .. d} \ dest_vec1 b < dest_vec1 a \ dest_vec1 d < dest_vec1 c \ dest_vec1 a = dest_vec1 c \ dest_vec1 b = dest_vec1 d" -unfolding set_eq_subset[of "{a .. b}" "{c .. d}"] -unfolding subset_interval_1(1)[of a b c d] -unfolding subset_interval_1(1)[of c d a b] -by auto + unfolding set_eq_subset[of "{a .. b}" "{c .. d}"] + unfolding subset_interval_1(1)[of a b c d] + unfolding subset_interval_1(1)[of c d a b] + by auto -lemma disjoint_interval_1: fixes a :: "real^1" shows - "{a .. b} \ {c .. d} = {} \ dest_vec1 b < dest_vec1 a \ dest_vec1 d < dest_vec1 c \ dest_vec1 b < dest_vec1 c \ dest_vec1 d < dest_vec1 a" - "{a .. b} \ {c<.. dest_vec1 b < dest_vec1 a \ dest_vec1 d \ dest_vec1 c \ dest_vec1 b \ dest_vec1 c \ dest_vec1 d \ dest_vec1 a" - "{a<.. {c .. d} = {} \ dest_vec1 b \ dest_vec1 a \ dest_vec1 d < dest_vec1 c \ dest_vec1 b \ dest_vec1 c \ dest_vec1 d \ dest_vec1 a" - "{a<.. {c<.. dest_vec1 b \ dest_vec1 a \ dest_vec1 d \ dest_vec1 c \ dest_vec1 b \ dest_vec1 c \ dest_vec1 d \ dest_vec1 a" +lemma disjoint_interval_1: + fixes a :: "real^1" + shows + "{a .. b} \ {c .. d} = {} \ + dest_vec1 b < dest_vec1 a \ dest_vec1 d < dest_vec1 c \ dest_vec1 b < dest_vec1 c \ dest_vec1 d < dest_vec1 a" + "{a .. b} \ {c<.. + dest_vec1 b < dest_vec1 a \ dest_vec1 d \ dest_vec1 c \ dest_vec1 b \ dest_vec1 c \ dest_vec1 d \ dest_vec1 a" + "{a<.. {c .. d} = {} \ + dest_vec1 b \ dest_vec1 a \ dest_vec1 d < dest_vec1 c \ dest_vec1 b \ dest_vec1 c \ dest_vec1 d \ dest_vec1 a" + "{a<.. {c<.. + dest_vec1 b \ dest_vec1 a \ dest_vec1 d \ dest_vec1 c \ dest_vec1 b \ dest_vec1 c \ dest_vec1 d \ dest_vec1 a" unfolding disjoint_interval_cart and ex_1 by auto -lemma open_closed_interval_1: fixes a :: "real^1" shows - "{a<.. dest_vec1 b ==> {a .. b} = {a<.. {a,b}" - unfolding set_eq_iff apply simp unfolding less_vec_def and less_eq_vec_def and forall_1 and dest_vec1_eq[THEN sym] by(auto simp del:dest_vec1_eq) +lemma closed_open_interval_1: + "dest_vec1 (a::real^1) \ dest_vec1 b ==> {a .. b} = {a<.. {a,b}" + unfolding set_eq_iff + apply simp + unfolding less_vec_def and less_eq_vec_def and forall_1 and dest_vec1_eq[symmetric] + apply (auto simp del:dest_vec1_eq) + done -lemma Lim_drop_le: fixes f :: "'a \ real^1" shows - "(f ---> l) net \ ~(trivial_limit net) \ eventually (\x. dest_vec1 (f x) \ b) net ==> dest_vec1 l \ b" +lemma Lim_drop_le: + fixes f :: "'a \ real^1" + shows "(f ---> l) net \ \ trivial_limit net \ + eventually (\x. dest_vec1 (f x) \ b) net ==> dest_vec1 l \ b" using Lim_component_le_cart[of f l net 1 b] by auto -lemma Lim_drop_ge: fixes f :: "'a \ real^1" shows - "(f ---> l) net \ ~(trivial_limit net) \ eventually (\x. b \ dest_vec1 (f x)) net ==> b \ dest_vec1 l" +lemma Lim_drop_ge: + fixes f :: "'a \ real^1" + shows "(f ---> l) net \ \ trivial_limit net \ + eventually (\x. b \ dest_vec1 (f x)) net ==> b \ dest_vec1 l" using Lim_component_ge_cart[of f l net b 1] by auto + text{* Also more convenient formulations of monotone convergence. *} -lemma bounded_increasing_convergent: fixes s::"nat \ real^1" +lemma bounded_increasing_convergent: + fixes s :: "nat \ real^1" assumes "bounded {s n| n::nat. True}" "\n. dest_vec1(s n) \ dest_vec1(s(Suc n))" shows "\l. (s ---> l) sequentially" -proof- - obtain a where a:"\n. \dest_vec1 (s n)\ \ a" using assms(1)[unfolded bounded_iff abs_dest_vec1] by auto +proof - + obtain a where a:"\n. \dest_vec1 (s n)\ \ a" + using assms(1)[unfolded bounded_iff abs_dest_vec1] by auto { fix m::nat have "\ n. n\m \ dest_vec1 (s m) \ dest_vec1 (s n)" - apply(induct_tac n) apply simp using assms(2) apply(erule_tac x="na" in allE) by(auto simp add: not_less_eq_eq) } - hence "\m n. m \ n \ dest_vec1 (s m) \ dest_vec1 (s n)" by auto - then obtain l where "\e>0. \N. \n\N. \dest_vec1 (s n) - l\ < e" using convergent_bounded_monotone[OF a] unfolding monoseq_def by auto + apply (induct_tac n) + apply simp + using assms(2) apply (erule_tac x="na" in allE) + apply (auto simp add: not_less_eq_eq) + done + } + hence "\m n. m \ n \ dest_vec1 (s m) \ dest_vec1 (s n)" + by auto + then obtain l where "\e>0. \N. \n\N. \dest_vec1 (s n) - l\ < e" + using convergent_bounded_monotone[OF a] unfolding monoseq_def by auto thus ?thesis unfolding LIMSEQ_def apply(rule_tac x="vec1 l" in exI) - unfolding dist_norm unfolding abs_dest_vec1 by auto + unfolding dist_norm unfolding abs_dest_vec1 by auto qed -lemma dest_vec1_simps[simp]: fixes a::"real^1" +lemma dest_vec1_simps[simp]: + fixes a :: "real^1" shows "a$1 = 0 \ a = 0" (*"a \ 1 \ dest_vec1 a \ 1" "0 \ a \ 0 \ dest_vec1 a"*) - "a \ b \ dest_vec1 a \ dest_vec1 b" "dest_vec1 (1::real^1) = 1" - by(auto simp add: less_eq_vec_def vec_eq_iff) + "a \ b \ dest_vec1 a \ dest_vec1 b" "dest_vec1 (1::real^1) = 1" + by (auto simp add: less_eq_vec_def vec_eq_iff) lemma dest_vec1_inverval: "dest_vec1 ` {a .. b} = {dest_vec1 a .. dest_vec1 b}" @@ -1861,106 +2125,162 @@ apply(rule_tac [!] allI)apply(rule_tac [!] impI) apply(rule_tac[2] x="vec1 x" in exI)apply(rule_tac[4] x="vec1 x" in exI) apply(rule_tac[6] x="vec1 x" in exI)apply(rule_tac[8] x="vec1 x" in exI) - by (auto simp add: less_vec_def less_eq_vec_def) + apply (auto simp add: less_vec_def less_eq_vec_def) + done -lemma dest_vec1_setsum: assumes "finite S" +lemma dest_vec1_setsum: + assumes "finite S" shows " dest_vec1 (setsum f S) = setsum (\x. dest_vec1 (f x)) S" using dest_vec1_sum[OF assms] by auto lemma open_dest_vec1_vimage: "open S \ open (dest_vec1 -` S)" -unfolding open_vec_def forall_1 by auto + unfolding open_vec_def forall_1 by auto lemma tendsto_dest_vec1 [tendsto_intros]: "(f ---> l) net \ ((\x. dest_vec1 (f x)) ---> dest_vec1 l) net" -by(rule tendsto_vec_nth) + by (rule tendsto_vec_nth) -lemma continuous_dest_vec1: "continuous net f \ continuous net (\x. dest_vec1 (f x))" +lemma continuous_dest_vec1: + "continuous net f \ continuous net (\x. dest_vec1 (f x))" unfolding continuous_def by (rule tendsto_dest_vec1) lemma forall_dest_vec1: "(\x. P x) \ (\x. P(dest_vec1 x))" - apply safe defer apply(erule_tac x="vec1 x" in allE) by auto + apply safe defer + apply (erule_tac x="vec1 x" in allE) + apply auto + done lemma forall_of_dest_vec1: "(\v. P (\x. dest_vec1 (v x))) \ (\x. P x)" - apply rule apply rule apply(erule_tac x="(vec1 \ x)" in allE) unfolding o_def vec1_dest_vec1 by auto + apply rule + apply rule + apply (erule_tac x="vec1 \ x" in allE) + unfolding o_def vec1_dest_vec1 + apply auto + done lemma forall_of_dest_vec1': "(\v. P (dest_vec1 v)) \ (\x. P x)" - apply rule apply rule apply(erule_tac x="(vec1 x)" in allE) defer apply rule - apply(erule_tac x="dest_vec1 v" in allE) unfolding o_def vec1_dest_vec1 by auto + apply rule + apply rule + apply (erule_tac x="(vec1 x)" in allE) defer + apply rule + apply (erule_tac x="dest_vec1 v" in allE) + unfolding o_def vec1_dest_vec1 + apply auto + done -lemma dist_vec1_0[simp]: "dist(vec1 (x::real)) 0 = norm x" unfolding dist_norm by auto +lemma dist_vec1_0[simp]: "dist(vec1 (x::real)) 0 = norm x" + unfolding dist_norm by auto -lemma bounded_linear_vec1_dest_vec1: fixes f::"real \ real" - shows "linear (vec1 \ f \ dest_vec1) = bounded_linear f" (is "?l = ?r") proof- - { assume ?l guess K using linear_bounded[OF `?l`] .. - hence "\K. \x. \f x\ \ \x\ * K" apply(rule_tac x=K in exI) - unfolding vec1_dest_vec1_simps by (auto simp add:field_simps) } - thus ?thesis unfolding linear_def bounded_linear_def additive_def bounded_linear_axioms_def o_def - unfolding vec1_dest_vec1_simps by auto qed +lemma bounded_linear_vec1_dest_vec1: + fixes f :: "real \ real" + shows "linear (vec1 \ f \ dest_vec1) = bounded_linear f" (is "?l = ?r") +proof - + { assume ?l + then have "\K. \x. norm ((vec1 \ f \ dest_vec1) x) \ K * norm x" by (rule linear_bounded) + then guess K .. + hence "\K. \x. \f x\ \ \x\ * K" + apply(rule_tac x=K in exI) + unfolding vec1_dest_vec1_simps by (auto simp add:field_simps) + } + thus ?thesis + unfolding linear_def bounded_linear_def additive_def bounded_linear_axioms_def o_def + unfolding vec1_dest_vec1_simps by auto +qed -lemma vec1_le[simp]:fixes a::real shows "vec1 a \ vec1 b \ a \ b" +lemma vec1_le[simp]: fixes a :: real shows "vec1 a \ vec1 b \ a \ b" unfolding less_eq_vec_def by auto -lemma vec1_less[simp]:fixes a::real shows "vec1 a < vec1 b \ a < b" +lemma vec1_less[simp]: fixes a :: real shows "vec1 a < vec1 b \ a < b" unfolding less_vec_def by auto subsection {* Derivatives on real = Derivatives on @{typ "real^1"} *} -lemma has_derivative_within_vec1_dest_vec1: fixes f::"real\real" shows - "((vec1 \ f \ dest_vec1) has_derivative (vec1 \ f' \ dest_vec1)) (at (vec1 x) within vec1 ` s) - = (f has_derivative f') (at x within s)" - unfolding has_derivative_within unfolding bounded_linear_vec1_dest_vec1[unfolded linear_conv_bounded_linear] +lemma has_derivative_within_vec1_dest_vec1: + fixes f :: "real \ real" + shows "((vec1 \ f \ dest_vec1) has_derivative (vec1 \ f' \ dest_vec1)) (at (vec1 x) within vec1 ` s) + = (f has_derivative f') (at x within s)" + unfolding has_derivative_within + unfolding bounded_linear_vec1_dest_vec1[unfolded linear_conv_bounded_linear] unfolding o_def Lim_within Ball_def unfolding forall_vec1 - unfolding vec1_dest_vec1_simps dist_vec1_0 image_iff by auto + unfolding vec1_dest_vec1_simps dist_vec1_0 image_iff + by auto -lemma has_derivative_at_vec1_dest_vec1: fixes f::"real\real" shows - "((vec1 \ f \ dest_vec1) has_derivative (vec1 \ f' \ dest_vec1)) (at (vec1 x)) = (f has_derivative f') (at x)" - using has_derivative_within_vec1_dest_vec1[where s=UNIV, unfolded range_vec1 within_UNIV] by auto +lemma has_derivative_at_vec1_dest_vec1: + fixes f :: "real \ real" + shows "((vec1 \ f \ dest_vec1) has_derivative (vec1 \ f' \ dest_vec1)) (at (vec1 x)) = (f has_derivative f') (at x)" + using has_derivative_within_vec1_dest_vec1[where s=UNIV, unfolded range_vec1 within_UNIV] + by auto -lemma bounded_linear_vec1': fixes f::"'a::real_normed_vector\real" +lemma bounded_linear_vec1': + fixes f :: "'a::real_normed_vector\real" shows "bounded_linear f = bounded_linear (vec1 \ f)" unfolding bounded_linear_def additive_def bounded_linear_axioms_def o_def unfolding vec1_dest_vec1_simps by auto -lemma bounded_linear_dest_vec1: fixes f::"real\'a::real_normed_vector" +lemma bounded_linear_dest_vec1: + fixes f :: "real\'a::real_normed_vector" shows "bounded_linear f = bounded_linear (f \ dest_vec1)" unfolding bounded_linear_def additive_def bounded_linear_axioms_def o_def - unfolding vec1_dest_vec1_simps by auto + unfolding vec1_dest_vec1_simps + by auto -lemma has_derivative_at_vec1: fixes f::"'a::real_normed_vector\real" shows - "(f has_derivative f') (at x) = ((vec1 \ f) has_derivative (vec1 \ f')) (at x)" - unfolding has_derivative_at unfolding bounded_linear_vec1'[unfolded linear_conv_bounded_linear] - unfolding o_def Lim_at unfolding vec1_dest_vec1_simps dist_vec1_0 by auto +lemma has_derivative_at_vec1: + fixes f :: "'a::real_normed_vector\real" + shows "(f has_derivative f') (at x) = ((vec1 \ f) has_derivative (vec1 \ f')) (at x)" + unfolding has_derivative_at + unfolding bounded_linear_vec1'[unfolded linear_conv_bounded_linear] + unfolding o_def Lim_at + unfolding vec1_dest_vec1_simps dist_vec1_0 + by auto -lemma has_derivative_within_dest_vec1:fixes f::"real\'a::real_normed_vector" shows - "((f \ dest_vec1) has_derivative (f' \ dest_vec1)) (at (vec1 x) within vec1 ` s) = (f has_derivative f') (at x within s)" - unfolding has_derivative_within bounded_linear_dest_vec1 unfolding o_def Lim_within Ball_def - unfolding vec1_dest_vec1_simps dist_vec1_0 image_iff by auto +lemma has_derivative_within_dest_vec1: + fixes f :: "real\'a::real_normed_vector" + shows "((f \ dest_vec1) has_derivative (f' \ dest_vec1)) (at (vec1 x) within vec1 ` s) = + (f has_derivative f') (at x within s)" + unfolding has_derivative_within bounded_linear_dest_vec1 + unfolding o_def Lim_within Ball_def + unfolding vec1_dest_vec1_simps dist_vec1_0 image_iff + by auto -lemma has_derivative_at_dest_vec1:fixes f::"real\'a::real_normed_vector" shows - "((f \ dest_vec1) has_derivative (f' \ dest_vec1)) (at (vec1 x)) = (f has_derivative f') (at x)" +lemma has_derivative_at_dest_vec1: + fixes f :: "real\'a::real_normed_vector" + shows "((f \ dest_vec1) has_derivative (f' \ dest_vec1)) (at (vec1 x)) = + (f has_derivative f') (at x)" using has_derivative_within_dest_vec1[where s=UNIV] by simp + subsection {* In particular if we have a mapping into @{typ "real^1"}. *} -lemma onorm_vec1: fixes f::"real \ real" - shows "onorm (\x. vec1 (f (dest_vec1 x))) = onorm f" proof- - have "\x::real^1. norm x = 1 \ x\{vec1 -1, vec1 (1::real)}" unfolding forall_vec1 by(auto simp add:vec_eq_iff) - hence 1:"{x. norm x = 1} = {vec1 -1, vec1 (1::real)}" by auto - have 2:"{norm (vec1 (f (dest_vec1 x))) |x. norm x = 1} = (\x. norm (vec1 (f (dest_vec1 x)))) ` {x. norm x=1}" by auto - have "\x::real. norm x = 1 \ x\{-1, 1}" by auto hence 3:"{x. norm x = 1} = {-1, (1::real)}" by auto +lemma onorm_vec1: + fixes f::"real \ real" + shows "onorm (\x. vec1 (f (dest_vec1 x))) = onorm f" +proof - + have "\x::real^1. norm x = 1 \ x\{vec1 -1, vec1 (1::real)}" + unfolding forall_vec1 by (auto simp add: vec_eq_iff) + hence 1: "{x. norm x = 1} = {vec1 -1, vec1 (1::real)}" by auto + have 2: "{norm (vec1 (f (dest_vec1 x))) |x. norm x = 1} = + (\x. norm (vec1 (f (dest_vec1 x)))) ` {x. norm x=1}" + by auto + have "\x::real. norm x = 1 \ x\{-1, 1}" by auto + hence 3:"{x. norm x = 1} = {-1, (1::real)}" by auto have 4:"{norm (f x) |x. norm x = 1} = (\x. norm (f x)) ` {x. norm x=1}" by auto - show ?thesis unfolding onorm_def 1 2 3 4 by(simp add:Sup_finite_Max) qed + show ?thesis + unfolding onorm_def 1 2 3 4 by (simp add:Sup_finite_Max) +qed lemma convex_vec1:"convex (vec1 ` s) = convex (s::real set)" - unfolding convex_def Ball_def forall_vec1 unfolding vec1_dest_vec1_simps image_iff by auto + unfolding convex_def Ball_def forall_vec1 + unfolding vec1_dest_vec1_simps image_iff + by auto lemma bounded_linear_component_cart[intro]: "bounded_linear (\x::real^'n. x $ k)" - apply(rule bounded_linearI[where K=1]) + apply (rule bounded_linearI[where K=1]) using component_le_norm_cart[of _ k] unfolding real_norm_def by auto -lemma bounded_vec1[intro]: "bounded s \ bounded (vec1 ` (s::real set))" +lemma bounded_vec1[intro]: "bounded s \ bounded (vec1 ` (s::real set))" unfolding bounded_def apply safe apply(rule_tac x="vec1 x" in exI,rule_tac x=e in exI) - by(auto simp add: dist_real dist_real_def) + apply (auto simp add: dist_real dist_real_def) + done (*lemma content_closed_interval_cases_cart: "content {a..b::real^'n} = @@ -1975,15 +2295,19 @@ sorry*) -lemma integral_component_eq_cart[simp]: fixes f::"'n::ordered_euclidean_space \ real^'m" - assumes "f integrable_on s" shows "integral s (\x. f x $ k) = integral s f $ k" +lemma integral_component_eq_cart[simp]: + fixes f :: "'n::ordered_euclidean_space \ real^'m" + assumes "f integrable_on s" + shows "integral s (\x. f x $ k) = integral s f $ k" using integral_linear[OF assms(1) bounded_linear_component_cart,unfolded o_def] . lemma interval_split_cart: "{a..b::real^'n} \ {x. x$k \ c} = {a .. (\ i. if i = k then min (b$k) c else b$i)}" "{a..b} \ {x. x$k \ c} = {(\ i. if i = k then max (a$k) c else a$i) .. b}" - apply(rule_tac[!] set_eqI) unfolding Int_iff mem_interval_cart mem_Collect_eq - unfolding vec_lambda_beta by auto + apply (rule_tac[!] set_eqI) + unfolding Int_iff mem_interval_cart mem_Collect_eq + unfolding vec_lambda_beta + by auto (*lemma content_split_cart: "content {a..b::real^'n} = content({a..b} \ {x. x$k \ c}) + content({a..b} \ {x. x$k >= c})" @@ -2006,7 +2330,8 @@ assumes "(f has_integral k) {a..b}" shows "((\x. vec1 (f x)) has_integral (vec1 k)) {a..b}" proof - - have *: "\p. (\(x, k)\p. content k *\<^sub>R vec1 (f x)) - vec1 k = vec1 ((\(x, k)\p. content k *\<^sub>R f x) - k)" + have *: "\p. (\(x, k)\p. content k *\<^sub>R vec1 (f x)) - vec1 k = + vec1 ((\(x, k)\p. content k *\<^sub>R f x) - k)" unfolding vec_sub vec_eq_iff by (auto simp add: split_beta) show ?thesis using assms unfolding has_integral