hoelzl@37489: hoelzl@37489: header {*Instanciates the finite cartesian product of euclidean spaces as a euclidean space.*} hoelzl@37489: hoelzl@37489: theory Cartesian_Euclidean_Space hoelzl@37489: imports Finite_Cartesian_Product Integration hoelzl@37489: begin hoelzl@37489: hoelzl@37489: lemma delta_mult_idempotent: hoelzl@37489: "(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) hoelzl@37489: hoelzl@37489: lemma setsum_Plus: hoelzl@37489: "\finite A; finite B\ \ hoelzl@37489: (\x\A <+> B. g x) = (\x\A. g (Inl x)) + (\x\B. g (Inr x))" hoelzl@37489: unfolding Plus_def hoelzl@37489: by (subst setsum_Un_disjoint, auto simp add: setsum_reindex) hoelzl@37489: hoelzl@37489: lemma setsum_UNIV_sum: hoelzl@37489: fixes g :: "'a::finite + 'b::finite \ _" hoelzl@37489: shows "(\x\UNIV. g x) = (\x\UNIV. g (Inl x)) + (\x\UNIV. g (Inr x))" hoelzl@37489: apply (subst UNIV_Plus_UNIV [symmetric]) hoelzl@37489: apply (rule setsum_Plus [OF finite finite]) hoelzl@37489: done hoelzl@37489: hoelzl@37489: lemma setsum_mult_product: hoelzl@37489: "setsum h {..i\{..j\{..j. j + i * B) {..j. j + i * B) ` {.. {i * B.. (\j. j + i * B) ` {.. (\ x y. (\ i. (x$i) * (y$i)))" hoelzl@37489: instance .. hoelzl@37489: end hoelzl@37489: hoelzl@37489: instantiation cart :: (one,finite) one hoelzl@37489: begin hoelzl@37489: definition vector_one_def : "1 \ (\ i. 1)" hoelzl@37489: instance .. hoelzl@37489: end hoelzl@37489: hoelzl@37489: instantiation cart :: (ord,finite) ord hoelzl@37489: begin hoelzl@37489: definition vector_le_def: hoelzl@37489: "less_eq (x :: 'a ^'b) y = (ALL i. x$i <= y$i)" hoelzl@37489: definition vector_less_def: "less (x :: 'a ^'b) y = (ALL i. x$i < y$i)" hoelzl@37489: instance by (intro_classes) hoelzl@37489: end hoelzl@37489: hoelzl@37489: text{* The ordering on one-dimensional vectors is linear. *} hoelzl@37489: hoelzl@37489: class cart_one = assumes UNIV_one: "card (UNIV \ 'a set) = Suc 0" hoelzl@37489: begin hoelzl@37489: subclass finite hoelzl@37489: proof from UNIV_one show "finite (UNIV :: 'a set)" hoelzl@37489: by (auto intro!: card_ge_0_finite) qed hoelzl@37489: end hoelzl@37489: hoelzl@37489: instantiation cart :: (linorder,cart_one) linorder begin hoelzl@37489: instance proof hoelzl@37489: guess a B using UNIV_one[where 'a='b] unfolding card_Suc_eq apply- by(erule exE)+ hoelzl@37489: hence *:"UNIV = {a}" by auto hoelzl@37489: have "\P. (\i\UNIV. P i) \ P a" unfolding * by auto hence all:"\P. (\i. P i) \ P a" by auto hoelzl@37489: fix x y z::"'a^'b::cart_one" note * = vector_le_def vector_less_def all Cart_eq hoelzl@37489: show "x\x" "(x < y) = (x \ y \ \ y \ x)" "x\y \ y\x" unfolding * by(auto simp only:field_simps) hoelzl@37489: { assume "x\y" "y\z" thus "x\z" unfolding * by(auto simp only:field_simps) } hoelzl@37489: { assume "x\y" "y\x" thus "x=y" unfolding * by(auto simp only:field_simps) } hoelzl@37489: qed end hoelzl@37489: hoelzl@37489: text{* Constant Vectors *} hoelzl@37489: hoelzl@37489: definition "vec x = (\ i. x)" hoelzl@37489: hoelzl@37489: text{* Also the scalar-vector multiplication. *} hoelzl@37489: hoelzl@37489: definition vector_scalar_mult:: "'a::times \ 'a ^ 'n \ 'a ^ 'n" (infixl "*s" 70) hoelzl@37489: where "c *s x = (\ i. c * (x$i))" hoelzl@37489: hoelzl@37489: subsection {* A naive proof procedure to lift really trivial arithmetic stuff from the basis of the vector space. *} hoelzl@37489: hoelzl@37489: method_setup vector = {* hoelzl@37489: let hoelzl@37489: val ss1 = HOL_basic_ss addsimps [@{thm setsum_addf} RS sym, hoelzl@37489: @{thm setsum_subtractf} RS sym, @{thm setsum_right_distrib}, hoelzl@37489: @{thm setsum_left_distrib}, @{thm setsum_negf} RS sym] hoelzl@37489: val ss2 = @{simpset} addsimps hoelzl@37489: [@{thm vector_add_def}, @{thm vector_mult_def}, hoelzl@37489: @{thm vector_minus_def}, @{thm vector_uminus_def}, hoelzl@37489: @{thm vector_one_def}, @{thm vector_zero_def}, @{thm vec_def}, hoelzl@37489: @{thm vector_scaleR_def}, hoelzl@37489: @{thm Cart_lambda_beta}, @{thm vector_scalar_mult_def}] hoelzl@37489: fun vector_arith_tac ths = hoelzl@37489: simp_tac ss1 hoelzl@37489: THEN' (fn i => rtac @{thm setsum_cong2} i hoelzl@37489: ORELSE rtac @{thm setsum_0'} i hoelzl@37489: ORELSE simp_tac (HOL_basic_ss addsimps [@{thm "Cart_eq"}]) i) hoelzl@37489: (* THEN' TRY o clarify_tac HOL_cs THEN' (TRY o rtac @{thm iffI}) *) hoelzl@37489: THEN' asm_full_simp_tac (ss2 addsimps ths) hoelzl@37489: in hoelzl@37489: Attrib.thms >> (fn ths => K (SIMPLE_METHOD' (vector_arith_tac ths))) hoelzl@37489: end wenzelm@43685: *} "lift trivial vector statements to real arith statements" hoelzl@37489: hoelzl@37489: lemma vec_0[simp]: "vec 0 = 0" by (vector vector_zero_def) hoelzl@37489: lemma vec_1[simp]: "vec 1 = 1" by (vector vector_one_def) hoelzl@37489: hoelzl@37489: lemma vec_inj[simp]: "vec x = vec y \ x = y" by vector hoelzl@37489: hoelzl@37489: lemma vec_in_image_vec: "vec x \ (vec ` S) \ x \ S" by auto hoelzl@37489: hoelzl@37489: lemma vec_add: "vec(x + y) = vec x + vec y" by (vector vec_def) hoelzl@37489: lemma vec_sub: "vec(x - y) = vec x - vec y" by (vector vec_def) hoelzl@37489: lemma vec_cmul: "vec(c * x) = c *s vec x " by (vector vec_def) hoelzl@37489: lemma vec_neg: "vec(- x) = - vec x " by (vector vec_def) hoelzl@37489: hoelzl@37489: lemma vec_setsum: assumes fS: "finite S" hoelzl@37489: shows "vec(setsum f S) = setsum (vec o f) S" hoelzl@37489: apply (induct rule: finite_induct[OF fS]) hoelzl@37489: apply (simp) hoelzl@37489: apply (auto simp add: vec_add) hoelzl@37489: done hoelzl@37489: hoelzl@37489: text{* Obvious "component-pushing". *} hoelzl@37489: hoelzl@37489: lemma vec_component [simp]: "vec x $ i = x" hoelzl@37489: by (vector vec_def) hoelzl@37489: hoelzl@37489: lemma vector_mult_component [simp]: "(x * y)$i = x$i * y$i" hoelzl@37489: by vector hoelzl@37489: hoelzl@37489: lemma vector_smult_component [simp]: "(c *s y)$i = c * (y$i)" hoelzl@37489: by vector hoelzl@37489: hoelzl@37489: lemma cond_component: "(if b then x else y)$i = (if b then x$i else y$i)" by vector hoelzl@37489: hoelzl@37489: lemmas vector_component = hoelzl@37489: vec_component vector_add_component vector_mult_component hoelzl@37489: vector_smult_component vector_minus_component vector_uminus_component hoelzl@37489: vector_scaleR_component cond_component hoelzl@37489: hoelzl@37489: subsection {* Some frequently useful arithmetic lemmas over vectors. *} hoelzl@37489: hoelzl@37489: instance cart :: (semigroup_mult,finite) semigroup_mult hoelzl@37489: apply (intro_classes) by (vector mult_assoc) hoelzl@37489: hoelzl@37489: instance cart :: (monoid_mult,finite) monoid_mult hoelzl@37489: apply (intro_classes) by vector+ hoelzl@37489: hoelzl@37489: instance cart :: (ab_semigroup_mult,finite) ab_semigroup_mult hoelzl@37489: apply (intro_classes) by (vector mult_commute) hoelzl@37489: hoelzl@37489: instance cart :: (ab_semigroup_idem_mult,finite) ab_semigroup_idem_mult hoelzl@37489: apply (intro_classes) by (vector mult_idem) hoelzl@37489: hoelzl@37489: instance cart :: (comm_monoid_mult,finite) comm_monoid_mult hoelzl@37489: apply (intro_classes) by vector hoelzl@37489: hoelzl@37489: instance cart :: (semiring,finite) semiring hoelzl@37489: apply (intro_classes) by (vector field_simps)+ hoelzl@37489: hoelzl@37489: instance cart :: (semiring_0,finite) semiring_0 hoelzl@37489: apply (intro_classes) by (vector field_simps)+ hoelzl@37489: instance cart :: (semiring_1,finite) semiring_1 hoelzl@37489: apply (intro_classes) by vector hoelzl@37489: instance cart :: (comm_semiring,finite) comm_semiring hoelzl@37489: apply (intro_classes) by (vector field_simps)+ hoelzl@37489: hoelzl@37489: instance cart :: (comm_semiring_0,finite) comm_semiring_0 by (intro_classes) hoelzl@37489: instance cart :: (cancel_comm_monoid_add, finite) cancel_comm_monoid_add .. hoelzl@37489: instance cart :: (semiring_0_cancel,finite) semiring_0_cancel by (intro_classes) hoelzl@37489: instance cart :: (comm_semiring_0_cancel,finite) comm_semiring_0_cancel by (intro_classes) hoelzl@37489: instance cart :: (ring,finite) ring by (intro_classes) hoelzl@37489: instance cart :: (semiring_1_cancel,finite) semiring_1_cancel by (intro_classes) hoelzl@37489: instance cart :: (comm_semiring_1,finite) comm_semiring_1 by (intro_classes) hoelzl@37489: hoelzl@37489: instance cart :: (ring_1,finite) ring_1 .. hoelzl@37489: hoelzl@37489: instance cart :: (real_algebra,finite) real_algebra hoelzl@37489: apply intro_classes hoelzl@37489: apply (simp_all add: vector_scaleR_def field_simps) hoelzl@37489: apply vector hoelzl@37489: apply vector hoelzl@37489: done hoelzl@37489: hoelzl@37489: instance cart :: (real_algebra_1,finite) real_algebra_1 .. hoelzl@37489: hoelzl@37489: lemma of_nat_index: hoelzl@37489: "(of_nat n :: 'a::semiring_1 ^'n)$i = of_nat n" hoelzl@37489: apply (induct n) hoelzl@37489: apply vector hoelzl@37489: apply vector hoelzl@37489: done hoelzl@37489: hoelzl@37489: lemma one_index[simp]: hoelzl@37489: "(1 :: 'a::one ^'n)$i = 1" by vector hoelzl@37489: haftmann@38844: instance cart :: (semiring_char_0, finite) semiring_char_0 haftmann@38844: proof haftmann@38844: fix m n :: nat haftmann@38844: show "inj (of_nat :: nat \ 'a ^ 'b)" haftmann@38844: by (auto intro!: injI simp add: Cart_eq of_nat_index) hoelzl@37489: qed hoelzl@37489: haftmann@38844: instance cart :: (comm_ring_1,finite) comm_ring_1 .. haftmann@38844: instance cart :: (ring_char_0,finite) ring_char_0 .. hoelzl@37489: hoelzl@37489: instance cart :: (real_vector,finite) real_vector .. hoelzl@37489: hoelzl@37489: lemma vector_smult_assoc: "a *s (b *s x) = ((a::'a::semigroup_mult) * b) *s x" hoelzl@37489: by (vector mult_assoc) hoelzl@37489: lemma vector_sadd_rdistrib: "((a::'a::semiring) + b) *s x = a *s x + b *s x" hoelzl@37489: by (vector field_simps) hoelzl@37489: lemma vector_add_ldistrib: "(c::'a::semiring) *s (x + y) = c *s x + c *s y" hoelzl@37489: by (vector field_simps) hoelzl@37489: lemma vector_smult_lzero[simp]: "(0::'a::mult_zero) *s x = 0" by vector hoelzl@37489: lemma vector_smult_lid[simp]: "(1::'a::monoid_mult) *s x = x" by vector hoelzl@37489: lemma vector_ssub_ldistrib: "(c::'a::ring) *s (x - y) = c *s x - c *s y" hoelzl@37489: by (vector field_simps) hoelzl@37489: lemma vector_smult_rneg: "(c::'a::ring) *s -x = -(c *s x)" by vector hoelzl@37489: lemma vector_smult_lneg: "- (c::'a::ring) *s x = -(c *s x)" by vector hoelzl@37489: lemma vector_sneg_minus1: "-x = (- (1::'a::ring_1)) *s x" by vector hoelzl@37489: lemma vector_smult_rzero[simp]: "c *s 0 = (0::'a::mult_zero ^ 'n)" by vector hoelzl@37489: lemma vector_sub_rdistrib: "((a::'a::ring) - b) *s x = a *s x - b *s x" hoelzl@37489: by (vector field_simps) hoelzl@37489: hoelzl@37489: lemma vec_eq[simp]: "(vec m = vec n) \ (m = n)" hoelzl@37489: by (simp add: Cart_eq) hoelzl@37489: hoelzl@37489: lemma norm_eq_0_imp: "norm x = 0 ==> x = (0::real ^'n)" by (metis norm_eq_zero) hoelzl@37489: lemma vector_mul_eq_0[simp]: "(a *s x = 0) \ a = (0::'a::idom) \ x = 0" hoelzl@37489: by vector hoelzl@37489: lemma vector_mul_lcancel[simp]: "a *s x = a *s y \ a = (0::real) \ x = y" hoelzl@37489: by (metis eq_iff_diff_eq_0 vector_mul_eq_0 vector_ssub_ldistrib) hoelzl@37489: lemma vector_mul_rcancel[simp]: "a *s x = b *s x \ (a::real) = b \ x = 0" hoelzl@37489: by (metis eq_iff_diff_eq_0 vector_mul_eq_0 vector_sub_rdistrib) hoelzl@37489: lemma vector_mul_lcancel_imp: "a \ (0::real) ==> a *s x = a *s y ==> (x = y)" hoelzl@37489: by (metis vector_mul_lcancel) hoelzl@37489: lemma vector_mul_rcancel_imp: "x \ 0 \ (a::real) *s x = b *s x ==> a = b" hoelzl@37489: by (metis vector_mul_rcancel) hoelzl@37489: hoelzl@37489: lemma component_le_norm_cart: "\x$i\ <= norm x" hoelzl@37489: apply (simp add: norm_vector_def) hoelzl@37489: apply (rule member_le_setL2, simp_all) hoelzl@37489: done hoelzl@37489: hoelzl@37489: lemma norm_bound_component_le_cart: "norm x <= e ==> \x$i\ <= e" hoelzl@37489: by (metis component_le_norm_cart order_trans) hoelzl@37489: hoelzl@37489: lemma norm_bound_component_lt_cart: "norm x < e ==> \x$i\ < e" hoelzl@37489: by (metis component_le_norm_cart basic_trans_rules(21)) hoelzl@37489: hoelzl@37489: lemma norm_le_l1_cart: "norm x <= setsum(\i. \x$i\) UNIV" hoelzl@37489: by (simp add: norm_vector_def setL2_le_setsum) hoelzl@37489: hoelzl@37489: lemma scalar_mult_eq_scaleR: "c *s x = c *\<^sub>R x" hoelzl@37489: unfolding vector_scaleR_def vector_scalar_mult_def by simp hoelzl@37489: hoelzl@37489: lemma dist_mul[simp]: "dist (c *s x) (c *s y) = \c\ * dist x y" hoelzl@37489: unfolding dist_norm scalar_mult_eq_scaleR hoelzl@37489: unfolding scaleR_right_diff_distrib[symmetric] by simp hoelzl@37489: hoelzl@37489: lemma setsum_component [simp]: hoelzl@37489: fixes f:: " 'a \ ('b::comm_monoid_add) ^'n" hoelzl@37489: shows "(setsum f S)$i = setsum (\x. (f x)$i) S" hoelzl@37489: by (cases "finite S", induct S set: finite, simp_all) hoelzl@37489: hoelzl@37489: lemma setsum_eq: "setsum f S = (\ i. setsum (\x. (f x)$i ) S)" hoelzl@37489: by (simp add: Cart_eq) hoelzl@37489: hoelzl@37489: lemma setsum_cmul: hoelzl@37489: fixes f:: "'c \ ('a::semiring_1)^'n" hoelzl@37489: shows "setsum (\x. c *s f x) S = c *s setsum f S" hoelzl@37489: by (simp add: Cart_eq setsum_right_distrib) hoelzl@37489: hoelzl@37489: (* TODO: use setsum_norm_allsubsets_bound *) hoelzl@37489: lemma setsum_norm_allsubsets_bound_cart: hoelzl@37489: fixes f:: "'a \ real ^'n" hoelzl@37489: assumes fP: "finite P" and fPs: "\Q. Q \ P \ norm (setsum f Q) \ e" hoelzl@37489: shows "setsum (\x. norm (f x)) P \ 2 * real CARD('n) * e" hoelzl@37489: proof- hoelzl@37489: let ?d = "real CARD('n)" hoelzl@37489: let ?nf = "\x. norm (f x)" hoelzl@37489: let ?U = "UNIV :: 'n set" hoelzl@37489: have th0: "setsum (\x. setsum (\i. \f x $ i\) ?U) P = setsum (\i. setsum (\x. \f x $ i\) P) ?U" hoelzl@37489: by (rule setsum_commute) hoelzl@37489: have th1: "2 * ?d * e = of_nat (card ?U) * (2 * e)" by (simp add: real_of_nat_def) hoelzl@37489: have "setsum ?nf P \ setsum (\x. setsum (\i. \f x $ i\) ?U) P" hoelzl@37489: apply (rule setsum_mono) by (rule norm_le_l1_cart) hoelzl@37489: also have "\ \ 2 * ?d * e" hoelzl@37489: unfolding th0 th1 hoelzl@37489: proof(rule setsum_bounded) hoelzl@37489: fix i assume i: "i \ ?U" hoelzl@37489: let ?Pp = "{x. x\ P \ f x $ i \ 0}" hoelzl@37489: let ?Pn = "{x. x \ P \ f x $ i < 0}" hoelzl@37489: have thp: "P = ?Pp \ ?Pn" by auto hoelzl@37489: have thp0: "?Pp \ ?Pn ={}" by auto hoelzl@37489: have PpP: "?Pp \ P" and PnP: "?Pn \ P" by blast+ hoelzl@37489: have Ppe:"setsum (\x. \f x $ i\) ?Pp \ e" hoelzl@37489: using component_le_norm_cart[of "setsum (\x. f x) ?Pp" i] fPs[OF PpP] hoelzl@37489: by (auto intro: abs_le_D1) hoelzl@37489: have Pne: "setsum (\x. \f x $ i\) ?Pn \ e" hoelzl@37489: using component_le_norm_cart[of "setsum (\x. - f x) ?Pn" i] fPs[OF PnP] hoelzl@37489: by (auto simp add: setsum_negf intro: abs_le_D1) hoelzl@37489: have "setsum (\x. \f x $ i\) P = setsum (\x. \f x $ i\) ?Pp + setsum (\x. \f x $ i\) ?Pn" hoelzl@37489: apply (subst thp) hoelzl@37489: apply (rule setsum_Un_zero) hoelzl@37489: using fP thp0 by auto hoelzl@37489: also have "\ \ 2*e" using Pne Ppe by arith hoelzl@37489: finally show "setsum (\x. \f x $ i\) P \ 2*e" . hoelzl@37489: qed hoelzl@37489: finally show ?thesis . hoelzl@37489: qed hoelzl@37489: hoelzl@37489: subsection {* A bijection between 'n::finite and {.. ('n::finite)" where hoelzl@37489: "cart_bij_nat = (SOME p. bij_betw p {.. \ cart_bij_nat" hoelzl@37489: definition "\' = inv_into {..::nat \ ('n::finite))" hoelzl@37489: hoelzl@37489: lemma bij_betw_pi: hoelzl@37489: "bij_betw \ {..x. bij_betw x {..' (UNIV::'n set) {..'_def by auto hoelzl@37489: hoelzl@37489: lemma pi'_inj[intro]: "inj \'" hoelzl@37489: using bij_betw_pi' unfolding bij_betw_def by auto hoelzl@37489: hoelzl@37489: lemma pi'_range[intro]: "\i::'n. \' i < CARD('n::finite)" hoelzl@37489: using bij_betw_pi' unfolding bij_betw_def by auto hoelzl@37489: hoelzl@37489: lemma \\'[simp]: "\i::'n::finite. \ (\' i) = i" hoelzl@37489: using bij_betw_pi by (auto intro!: f_inv_into_f simp: \'_def bij_betw_def) hoelzl@37489: hoelzl@37489: lemma \'\[simp]: "\i. i\{.. \' (\ i::'n) = i" hoelzl@37489: using bij_betw_pi by (auto intro!: inv_into_f_eq simp: \'_def bij_betw_def) hoelzl@37489: hoelzl@37489: lemma \\'_alt[simp]: "\i. i \' (\ i::'n) = i" hoelzl@37489: by auto hoelzl@37489: hoelzl@37489: lemma \_inj_on: "inj_on (\::nat\'n::finite) {.. j::'b. if j = \(i div DIM('a)) then basis (i mod DIM('a)) else 0) hoelzl@37489: else 0)" hoelzl@37489: hoelzl@37489: lemma basis_eq: hoelzl@37489: assumes "i < CARD('b)" and "j < DIM('a)" hoelzl@37489: shows "basis (j + i * DIM('a)) = (\ k. if k = \ i then basis j else 0)" hoelzl@37489: proof - hoelzl@37489: have "j + i * DIM('a) < DIM('a) * (i + 1)" using assms by (auto simp: field_simps) hoelzl@37489: also have "\ \ DIM('a) * CARD('b)" using assms unfolding mult_le_cancel1 by auto hoelzl@37489: finally show ?thesis hoelzl@37489: unfolding basis_cart_def using assms by (auto simp: Cart_eq not_less field_simps) hoelzl@37489: qed hoelzl@37489: hoelzl@37489: lemma basis_eq_pi': hoelzl@37489: assumes "j < DIM('a)" hoelzl@37489: shows "basis (j + \' i * DIM('a)) $ k = (if k = i then basis j else 0)" hoelzl@37489: apply (subst basis_eq) hoelzl@37489: using pi'_range assms by simp_all hoelzl@37489: hoelzl@37489: lemma split_times_into_modulo[consumes 1]: hoelzl@37489: fixes k :: nat hoelzl@37489: assumes "k < A * B" hoelzl@37489: obtains i j where "i < A" and "j < B" and "k = j + i * B" hoelzl@37489: proof hoelzl@37489: have "A * B \ 0" hoelzl@37489: proof assume "A * B = 0" with assms show False by simp qed hoelzl@37489: hence "0 < B" by auto hoelzl@37489: thus "k mod B < B" using `0 < B` by auto hoelzl@37489: next hoelzl@37489: have "k div B * B \ k div B * B + k mod B" by (rule le_add1) hoelzl@37489: also have "... < A * B" using assms by simp hoelzl@37489: finally show "k div B < A" by auto hoelzl@37489: qed simp hoelzl@37489: hoelzl@37489: lemma split_CARD_DIM[consumes 1]: hoelzl@37489: fixes k :: nat hoelzl@37489: assumes k: "k < CARD('b) * DIM('a)" hoelzl@37489: obtains i and j::'b where "i < DIM('a)" "k = i + \' j * DIM('a)" hoelzl@37489: proof - hoelzl@37489: from split_times_into_modulo[OF k] guess i j . note ij = this hoelzl@37489: show thesis hoelzl@37489: proof hoelzl@37489: show "j < DIM('a)" using ij by simp hoelzl@37489: show "k = j + \' (\ i :: 'b) * DIM('a)" hoelzl@37489: using ij by simp hoelzl@37489: qed hoelzl@37489: qed hoelzl@37489: hoelzl@37489: lemma linear_less_than_times: hoelzl@37489: fixes i j A B :: nat assumes "i < B" "j < A" hoelzl@37489: shows "j + i * A < B * A" hoelzl@37489: proof - hoelzl@37489: have "i * A + j < (Suc i)*A" using `j < A` by simp hoelzl@37489: also have "\ \ B * A" using `i < B` unfolding mult_le_cancel2 by simp hoelzl@37489: finally show ?thesis by simp hoelzl@37489: qed hoelzl@37489: hoelzl@37489: instance hoelzl@37489: proof hoelzl@37489: let ?b = "basis :: nat \ 'a^'b" hoelzl@37489: let ?b' = "basis :: nat \ 'a" hoelzl@37489: hoelzl@37489: have setsum_basis: hoelzl@37489: "\f. (\x\range basis. f (x::'a)) = f 0 + (\i real" assume "j < DIM('a)" hoelzl@37489: let ?x = "j + \' i * DIM('a)" hoelzl@37489: show "(\k\{..R ?b k) \ ?b ?x" hoelzl@37489: unfolding Cart_eq not_all hoelzl@37489: proof hoelzl@37489: have "(\j. j + \' i*DIM('a))`({..' i*DIM('a)..' i) * DIM('a)} - {?x}"(is "?S = ?I - _") hoelzl@37489: proof safe hoelzl@37489: fix y assume "y \ ?I" hoelzl@37489: moreover def k \ "y - \' i*DIM('a)" hoelzl@37489: ultimately have "k < DIM('a)" and "y = k + \' i * DIM('a)" by auto hoelzl@37489: moreover assume "y \ ?S" hoelzl@37489: ultimately show "y = j + \' i * DIM('a)" by auto hoelzl@37489: qed auto hoelzl@37489: hoelzl@37489: have "(\k\{..R ?b k) $ i = hoelzl@37489: (\k\{..R ?b k $ i)" by simp hoelzl@37489: also have "\ = (\k\?S. u(?b k) *\<^sub>R ?b k $ i)" hoelzl@37489: unfolding `?S = ?I - {?x}` hoelzl@37489: proof (safe intro!: setsum_mono_zero_cong_right) hoelzl@37489: fix y assume "y \ {\' i*DIM('a)..' i) * DIM('a)}" hoelzl@37489: moreover have "Suc (\' i) * DIM('a) \ CARD('b) * DIM('a)" hoelzl@37489: unfolding mult_le_cancel2 using pi'_range[of i] by simp hoelzl@37489: ultimately show "y < CARD('b) * DIM('a)" by simp hoelzl@37489: next hoelzl@37489: fix y assume "y < CARD('b) * DIM('a)" hoelzl@37489: with split_CARD_DIM guess l k . note y = this hoelzl@37489: moreover assume "u (?b y) *\<^sub>R ?b y $ i \ 0" hoelzl@37489: ultimately show "y \ {\' i*DIM('a)..' i) * DIM('a)}" hoelzl@37489: by (auto simp: basis_eq_pi' split: split_if_asm) hoelzl@37489: qed simp hoelzl@37489: also have "\ = (\k\{..' i*DIM('a))) *\<^sub>R (?b' k))" hoelzl@37489: by (subst setsum_reindex) (auto simp: basis_eq_pi' intro!: inj_onI) hoelzl@37489: also have "\ \ ?b ?x $ i" hoelzl@37489: proof - hoelzl@37664: note independent_eq_inj_on[THEN iffD1, OF basis_inj independent_basis, rule_format] hoelzl@37489: note this[of j "\v. u (\ ka::'b. if ka = i then v else (0\'a))"] hoelzl@37489: thus ?thesis by (simp add: `j < DIM('a)` basis_eq pi'_range) hoelzl@37489: qed hoelzl@37489: finally show "(\k\{..R ?b k) $ i \ ?b ?x $ i" . hoelzl@37489: qed hoelzl@37489: qed hoelzl@37489: ultimately hoelzl@37489: show "\d>0. ?b ` {d..} = {0} \ independent (?b ` {.. inj_on ?b {.. ?b ` {..i j. j < DIM('a) \ (THE k. (?if i j) $ k \ 0) = i" hoelzl@37489: by (rule the_equality) (simp_all split: split_if_asm add: basis_neq_0) hoelzl@37489: { fix x :: 'a hoelzl@37489: have "x \ span (range basis)" hoelzl@37664: using span_basis by (auto simp: range_basis) hoelzl@37489: hence "\u. (\xR ?b' x) = x" hoelzl@37489: by (subst (asm) span_finite) (auto simp: setsum_basis) } hoelzl@37489: hence "\i. \u. (\xR ?b' x) = i" by auto hoelzl@37489: then obtain u where u: "\i. (\xR ?b' x) = i" hoelzl@37489: by (auto dest: choice) hoelzl@37489: have "\u. \i. (\jR ?b' j) = x $ i" hoelzl@37489: apply (rule exI[of _ "\v. let i = (THE i. v$i \ 0) in u (x$i) (v$i)"]) hoelzl@37489: using The_if u by simp } hoelzl@37489: moreover hoelzl@37489: have "\i::'b. {.. {x. i = \ x} = {\' i}" hoelzl@37489: using pi'_range[where 'n='b] by auto hoelzl@37489: moreover hoelzl@37489: have "range ?b = {0} \ ?b ` {.. (0::'a^'b)" unfolding basis_cart_def hoelzl@37489: using * basis_finite[where 'a='a] hoelzl@37489: linear_less_than_times[of i "CARD('b)" j "DIM('a)"] hoelzl@37489: by (auto simp: A B field_simps Cart_eq basis_eq_0_iff) hoelzl@37489: qed (auto simp: basis_cart_def) hoelzl@37489: hoelzl@37489: lemma if_distr: "(if P then f else g) $ i = (if P then f $ i else g $ i)" by simp hoelzl@37489: hoelzl@37489: lemma split_dimensions'[consumes 1]: hoelzl@37489: assumes "k < DIM('a::real_basis^'b)" hoelzl@37489: obtains i j where "i < CARD('b::finite)" and "j < DIM('a::real_basis)" and "k = j + i * DIM('a::real_basis)" hoelzl@37489: using split_times_into_modulo[OF assms[simplified]] . hoelzl@37489: hoelzl@37489: lemma cart_euclidean_bound[intro]: hoelzl@37489: assumes j:"j < DIM('a::{real_basis})" hoelzl@37489: shows "j + \' (i::'b::finite) * DIM('a) < CARD('b) * DIM('a::real_basis)" hoelzl@37489: using linear_less_than_times[OF pi'_range j, of i] . hoelzl@37489: hoelzl@37489: instance cart :: (real_basis_with_inner,finite) real_basis_with_inner .. hoelzl@37489: hoelzl@37489: lemma (in real_basis) forall_CARD_DIM: hoelzl@37489: "(\i (\(i::'b::finite) j. j P (j + \' i * DIM('a)))" hoelzl@37489: (is "?l \ ?r") hoelzl@37489: proof (safe elim!: split_times_into_modulo) hoelzl@37489: fix i :: 'b and j assume "j < DIM('a)" hoelzl@37489: note linear_less_than_times[OF pi'_range[of i] this] hoelzl@37489: moreover assume "?l" hoelzl@37489: ultimately show "P (j + \' i * DIM('a))" by auto hoelzl@37489: next hoelzl@37489: fix i j assume "i < CARD('b)" "j < DIM('a)" and "?r" hoelzl@37489: from `?r`[rule_format, OF `j < DIM('a)`, of "\ i"] `i < CARD('b)` hoelzl@37489: show "P (j + i * DIM('a))" by simp hoelzl@37489: qed hoelzl@37489: hoelzl@37489: lemma (in real_basis) exists_CARD_DIM: hoelzl@37489: "(\i (\i::'b::finite. \j' i * DIM('a)))" hoelzl@37489: using forall_CARD_DIM[where 'b='b, of "\x. \ P x"] by blast hoelzl@37489: hoelzl@37489: lemma forall_CARD: hoelzl@37489: "(\i (\i::'b::finite. P (\' i))" hoelzl@37489: using forall_CARD_DIM[where 'a=real, of P] by simp hoelzl@37489: hoelzl@37489: lemma exists_CARD: hoelzl@37489: "(\i (\i::'b::finite. P (\' i))" hoelzl@37489: using exists_CARD_DIM[where 'a=real, of P] by simp hoelzl@37489: hoelzl@37489: lemmas cart_simps = forall_CARD_DIM exists_CARD_DIM forall_CARD exists_CARD hoelzl@37489: hoelzl@37489: lemma cart_euclidean_nth[simp]: hoelzl@37489: fixes x :: "('a::real_basis_with_inner, 'b::finite) cart" hoelzl@37489: assumes j:"j < DIM('a)" hoelzl@37489: shows "x $$ (j + \' i * DIM('a)) = x $ i $$ j" hoelzl@37489: unfolding euclidean_component_def inner_vector_def basis_eq_pi'[OF j] if_distrib cond_application_beta hoelzl@37489: by (simp add: setsum_cases) hoelzl@37489: hoelzl@37489: lemma real_euclidean_nth: hoelzl@37489: fixes x :: "real^'n" hoelzl@37489: shows "x $$ \' i = (x $ i :: real)" hoelzl@37489: using cart_euclidean_nth[where 'a=real, of 0 x i] by simp hoelzl@37489: hoelzl@37489: lemmas nth_conv_component = real_euclidean_nth[symmetric] hoelzl@37489: hoelzl@37489: lemma mult_split_eq: hoelzl@37489: fixes A :: nat assumes "x < A" "y < A" hoelzl@37489: shows "x + i * A = y + j * A \ x = y \ i = j" hoelzl@37489: proof hoelzl@37489: assume *: "x + i * A = y + j * A" hoelzl@37489: { fix x y i j assume "i < j" "x < A" and *: "x + i * A = y + j * A" hoelzl@37489: hence "x + i * A < Suc i * A" using `x < A` by simp hoelzl@37489: also have "\ \ j * A" using `i < j` unfolding mult_le_cancel2 by simp hoelzl@37489: also have "\ \ y + j * A" by simp hoelzl@37489: finally have "i = j" using * by simp } hoelzl@37489: note eq = this hoelzl@37489: hoelzl@37489: have "i = j" hoelzl@37489: proof (cases rule: linorder_cases) hoelzl@37489: assume "i < j" from eq[OF this `x < A` *] show "i = j" by simp hoelzl@37489: next hoelzl@37489: assume "j < i" from eq[OF this `y < A` *[symmetric]] show "i = j" by simp hoelzl@37489: qed simp hoelzl@37489: thus "x = y \ i = j" using * by simp hoelzl@37489: qed simp hoelzl@37489: hoelzl@37489: instance cart :: (euclidean_space,finite) euclidean_space hoelzl@37489: proof (default, safe elim!: split_dimensions') hoelzl@37489: let ?b = "basis :: nat \ 'a^'b" hoelzl@37489: have if_distrib_op: "\f P Q a b c d. hoelzl@37489: f (if P then a else b) (if Q then c else d) = hoelzl@37489: (if P then if Q then f a c else f a d else if Q then f b c else f b d)" hoelzl@37489: by simp hoelzl@37489: hoelzl@37489: fix i j k l hoelzl@37489: assume "i < CARD('b)" "k < CARD('b)" "j < DIM('a)" "l < DIM('a)" hoelzl@37489: thus "?b (j + i * DIM('a)) \ ?b (l + k * DIM('a)) = hoelzl@37489: (if j + i * DIM('a) = l + k * DIM('a) then 1 else 0)" hoelzl@37489: using inj_on_iff[OF \_inj_on[where 'n='b], of k i] hoelzl@37489: by (auto simp add: basis_eq inner_vector_def if_distrib_op[of inner] setsum_cases basis_orthonormal mult_split_eq) hoelzl@37489: qed hoelzl@37489: hoelzl@37489: instance cart :: (ordered_euclidean_space,finite) ordered_euclidean_space hoelzl@37489: proof hoelzl@37489: fix x y::"'a^'b" hoelzl@37489: show "(x \ y) = (\i y $$ i)" hoelzl@37489: unfolding vector_le_def apply(subst eucl_le) by (simp add: cart_simps) hoelzl@37489: show"(x < y) = (\i i. if i = k then 1 else 0)" hoelzl@37489: hoelzl@37489: lemma basis_component [simp]: "cart_basis k $ i = (if k=i then 1 else 0)" hoelzl@37489: unfolding cart_basis_def by simp hoelzl@37489: hoelzl@37489: lemma norm_basis[simp]: hoelzl@37489: shows "norm (cart_basis k :: real ^'n) = 1" hoelzl@37489: apply (simp add: cart_basis_def norm_eq_sqrt_inner) unfolding inner_vector_def hoelzl@37489: apply (vector delta_mult_idempotent) hoelzl@37489: using setsum_delta[of "UNIV :: 'n set" "k" "\k. 1::real"] by auto hoelzl@37489: hoelzl@37489: lemma norm_basis_1: "norm(cart_basis 1 :: real ^'n::{finite,one}) = 1" hoelzl@37489: by (rule norm_basis) hoelzl@37489: hoelzl@37489: lemma vector_choose_size: "0 <= c ==> \(x::real^'n). norm x = c" hoelzl@37489: by (rule exI[where x="c *\<^sub>R cart_basis arbitrary"]) simp hoelzl@37489: hoelzl@37489: lemma vector_choose_dist: assumes e: "0 <= e" hoelzl@37489: shows "\(y::real^'n). dist x y = e" hoelzl@37489: proof- hoelzl@37489: from vector_choose_size[OF e] obtain c:: "real ^'n" where "norm c = e" hoelzl@37489: by blast hoelzl@37489: then have "dist x (x - c) = e" by (simp add: dist_norm) hoelzl@37489: then show ?thesis by blast hoelzl@37489: qed hoelzl@37489: hoelzl@37489: lemma basis_inj[intro]: "inj (cart_basis :: 'n \ real ^'n)" hoelzl@37489: by (simp add: inj_on_def Cart_eq) hoelzl@37489: hoelzl@37489: lemma basis_expansion: hoelzl@37489: "setsum (\i. (x$i) *s cart_basis i) UNIV = (x::('a::ring_1) ^'n)" (is "?lhs = ?rhs" is "setsum ?f ?S = _") hoelzl@37489: by (auto simp add: Cart_eq if_distrib setsum_delta[of "?S", where ?'b = "'a", simplified] cong del: if_weak_cong) hoelzl@37489: hoelzl@37489: lemma smult_conv_scaleR: "c *s x = scaleR c x" hoelzl@37489: unfolding vector_scalar_mult_def vector_scaleR_def by simp hoelzl@37489: hoelzl@37489: lemma basis_expansion': hoelzl@37489: "setsum (\i. (x$i) *\<^sub>R cart_basis i) UNIV = x" hoelzl@37489: by (rule basis_expansion [where 'a=real, unfolded smult_conv_scaleR]) hoelzl@37489: hoelzl@37489: lemma basis_expansion_unique: hoelzl@37489: "setsum (\i. f i *s cart_basis i) UNIV = (x::('a::comm_ring_1) ^'n) \ (\i. f i = x$i)" hoelzl@37489: by (simp add: Cart_eq setsum_delta if_distrib cong del: if_weak_cong) hoelzl@37489: hoelzl@37489: lemma dot_basis: hoelzl@37489: shows "cart_basis i \ x = x$i" "x \ (cart_basis i) = (x$i)" hoelzl@37489: by (auto simp add: inner_vector_def cart_basis_def cond_application_beta if_distrib setsum_delta hoelzl@37489: cong del: if_weak_cong) hoelzl@37489: hoelzl@37489: lemma inner_basis: hoelzl@37489: fixes x :: "'a::{real_inner, real_algebra_1} ^ 'n" hoelzl@37489: shows "inner (cart_basis i) x = inner 1 (x $ i)" hoelzl@37489: and "inner x (cart_basis i) = inner (x $ i) 1" hoelzl@37489: unfolding inner_vector_def cart_basis_def hoelzl@37489: by (auto simp add: cond_application_beta if_distrib setsum_delta cong del: if_weak_cong) hoelzl@37489: hoelzl@37489: lemma basis_eq_0: "cart_basis i = (0::'a::semiring_1^'n) \ False" hoelzl@37489: by (auto simp add: Cart_eq) hoelzl@37489: hoelzl@37489: lemma basis_nonzero: hoelzl@37489: shows "cart_basis k \ (0:: 'a::semiring_1 ^'n)" hoelzl@37489: by (simp add: basis_eq_0) hoelzl@37489: hoelzl@37489: text {* some lemmas to map between Eucl and Cart *} hoelzl@37489: lemma basis_real_n[simp]:"(basis (\' i)::real^'a) = cart_basis i" hoelzl@37489: unfolding basis_cart_def using pi'_range[where 'n='a] hoelzl@37489: by (auto simp: Cart_eq Cart_lambda_beta) hoelzl@37489: hoelzl@37489: subsection {* Orthogonality on cartesian products *} hoelzl@37489: hoelzl@37489: lemma orthogonal_basis: hoelzl@37489: shows "orthogonal (cart_basis i) x \ x$i = (0::real)" hoelzl@37489: by (auto simp add: orthogonal_def inner_vector_def cart_basis_def if_distrib hoelzl@37489: cond_application_beta setsum_delta cong del: if_weak_cong) hoelzl@37489: hoelzl@37489: lemma orthogonal_basis_basis: hoelzl@37489: shows "orthogonal (cart_basis i :: real^'n) (cart_basis j) \ i \ j" hoelzl@37489: unfolding orthogonal_basis[of i] basis_component[of j] by simp hoelzl@37489: hoelzl@37489: subsection {* Linearity on cartesian products *} hoelzl@37489: hoelzl@37489: lemma linear_vmul_component: hoelzl@37489: assumes lf: "linear f" hoelzl@37489: shows "linear (\x. f x $ k *\<^sub>R v)" hoelzl@37489: using lf hoelzl@37489: by (auto simp add: linear_def algebra_simps) hoelzl@37489: hoelzl@37489: hoelzl@37489: subsection{* Adjoints on cartesian products *} hoelzl@37489: hoelzl@37489: text {* TODO: The following lemmas about adjoints should hold for any hoelzl@37489: Hilbert space (i.e. complete inner product space). hoelzl@37489: (see \url{http://en.wikipedia.org/wiki/Hermitian_adjoint}) hoelzl@37489: *} hoelzl@37489: hoelzl@37489: lemma adjoint_works_lemma: hoelzl@37489: fixes f:: "real ^'n \ real ^'m" hoelzl@37489: assumes lf: "linear f" hoelzl@37489: shows "\x y. f x \ y = x \ adjoint f y" hoelzl@37489: proof- hoelzl@37489: let ?N = "UNIV :: 'n set" hoelzl@37489: let ?M = "UNIV :: 'm set" hoelzl@37489: have fN: "finite ?N" by simp hoelzl@37489: have fM: "finite ?M" by simp hoelzl@37489: {fix y:: "real ^ 'm" hoelzl@37489: let ?w = "(\ i. (f (cart_basis i) \ y)) :: real ^ 'n" hoelzl@37489: {fix x hoelzl@37489: have "f x \ y = f (setsum (\i. (x$i) *\<^sub>R cart_basis i) ?N) \ y" hoelzl@37489: by (simp only: basis_expansion') hoelzl@37489: also have "\ = (setsum (\i. (x$i) *\<^sub>R f (cart_basis i)) ?N) \ y" hoelzl@37489: unfolding linear_setsum[OF lf fN] hoelzl@37489: by (simp add: linear_cmul[OF lf]) hoelzl@37489: finally have "f x \ y = x \ ?w" hoelzl@37489: apply (simp only: ) hoelzl@37489: apply (simp add: inner_vector_def setsum_left_distrib setsum_right_distrib setsum_commute[of _ ?M ?N] field_simps) hoelzl@37489: done} hoelzl@37489: } hoelzl@37489: then show ?thesis unfolding adjoint_def hoelzl@37489: some_eq_ex[of "\f'. \x y. f x \ y = x \ f' y"] hoelzl@37489: using choice_iff[of "\a b. \x. f x \ a = x \ b "] hoelzl@37489: by metis hoelzl@37489: qed hoelzl@37489: hoelzl@37489: lemma adjoint_works: hoelzl@37489: fixes f:: "real ^'n \ real ^'m" hoelzl@37489: assumes lf: "linear f" hoelzl@37489: shows "x \ adjoint f y = f x \ y" hoelzl@37489: using adjoint_works_lemma[OF lf] by metis hoelzl@37489: hoelzl@37489: lemma adjoint_linear: hoelzl@37489: fixes f:: "real ^'n \ real ^'m" hoelzl@37489: assumes lf: "linear f" hoelzl@37489: shows "linear (adjoint f)" hoelzl@37489: unfolding linear_def vector_eq_ldot[where 'a="real^'n", symmetric] apply safe hoelzl@37489: unfolding inner_simps smult_conv_scaleR adjoint_works[OF lf] by auto hoelzl@37489: hoelzl@37489: lemma adjoint_clauses: hoelzl@37489: fixes f:: "real ^'n \ real ^'m" hoelzl@37489: assumes lf: "linear f" hoelzl@37489: shows "x \ adjoint f y = f x \ y" hoelzl@37489: and "adjoint f y \ x = y \ f x" hoelzl@37489: by (simp_all add: adjoint_works[OF lf] inner_commute) hoelzl@37489: hoelzl@37489: lemma adjoint_adjoint: hoelzl@37489: fixes f:: "real ^'n \ real ^'m" hoelzl@37489: assumes lf: "linear f" hoelzl@37489: shows "adjoint (adjoint f) = f" hoelzl@37489: by (rule adjoint_unique, simp add: adjoint_clauses [OF lf]) hoelzl@37489: hoelzl@37489: hoelzl@37489: subsection {* Matrix operations *} hoelzl@37489: hoelzl@37489: text{* Matrix notation. NB: an MxN matrix is of type @{typ "'a^'n^'m"}, not @{typ "'a^'m^'n"} *} hoelzl@37489: hoelzl@37489: definition matrix_matrix_mult :: "('a::semiring_1) ^'n^'m \ 'a ^'p^'n \ 'a ^ 'p ^'m" (infixl "**" 70) hoelzl@37489: where "m ** m' == (\ i j. setsum (\k. ((m$i)$k) * ((m'$k)$j)) (UNIV :: 'n set)) ::'a ^ 'p ^'m" hoelzl@37489: hoelzl@37489: definition matrix_vector_mult :: "('a::semiring_1) ^'n^'m \ 'a ^'n \ 'a ^ 'm" (infixl "*v" 70) hoelzl@37489: where "m *v x \ (\ i. setsum (\j. ((m$i)$j) * (x$j)) (UNIV ::'n set)) :: 'a^'m" hoelzl@37489: hoelzl@37489: definition vector_matrix_mult :: "'a ^ 'm \ ('a::semiring_1) ^'n^'m \ 'a ^'n " (infixl "v*" 70) hoelzl@37489: where "v v* m == (\ j. setsum (\i. ((m$i)$j) * (v$i)) (UNIV :: 'm set)) :: 'a^'n" hoelzl@37489: hoelzl@37489: definition "(mat::'a::zero => 'a ^'n^'n) k = (\ i j. if i = j then k else 0)" hoelzl@37489: definition transpose where hoelzl@37489: "(transpose::'a^'n^'m \ 'a^'m^'n) A = (\ i j. ((A$j)$i))" hoelzl@37489: definition "(row::'m => 'a ^'n^'m \ 'a ^'n) i A = (\ j. ((A$i)$j))" hoelzl@37489: definition "(column::'n =>'a^'n^'m =>'a^'m) j A = (\ i. ((A$i)$j))" hoelzl@37489: definition "rows(A::'a^'n^'m) = { row i A | i. i \ (UNIV :: 'm set)}" hoelzl@37489: definition "columns(A::'a^'n^'m) = { column i A | i. i \ (UNIV :: 'n set)}" hoelzl@37489: hoelzl@37489: lemma mat_0[simp]: "mat 0 = 0" by (vector mat_def) hoelzl@37489: lemma matrix_add_ldistrib: "(A ** (B + C)) = (A ** B) + (A ** C)" hoelzl@37489: by (vector matrix_matrix_mult_def setsum_addf[symmetric] field_simps) hoelzl@37489: hoelzl@37489: lemma matrix_mul_lid: hoelzl@37489: fixes A :: "'a::semiring_1 ^ 'm ^ 'n" hoelzl@37489: shows "mat 1 ** A = A" hoelzl@37489: apply (simp add: matrix_matrix_mult_def mat_def) hoelzl@37489: apply vector hoelzl@37489: by (auto simp only: if_distrib cond_application_beta setsum_delta'[OF finite] mult_1_left mult_zero_left if_True UNIV_I) hoelzl@37489: hoelzl@37489: hoelzl@37489: lemma matrix_mul_rid: hoelzl@37489: fixes A :: "'a::semiring_1 ^ 'm ^ 'n" hoelzl@37489: shows "A ** mat 1 = A" hoelzl@37489: apply (simp add: matrix_matrix_mult_def mat_def) hoelzl@37489: apply vector hoelzl@37489: 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) hoelzl@37489: hoelzl@37489: lemma matrix_mul_assoc: "A ** (B ** C) = (A ** B) ** C" hoelzl@37489: apply (vector matrix_matrix_mult_def setsum_right_distrib setsum_left_distrib mult_assoc) hoelzl@37489: apply (subst setsum_commute) hoelzl@37489: apply simp hoelzl@37489: done hoelzl@37489: hoelzl@37489: lemma matrix_vector_mul_assoc: "A *v (B *v x) = (A ** B) *v x" hoelzl@37489: apply (vector matrix_matrix_mult_def matrix_vector_mult_def setsum_right_distrib setsum_left_distrib mult_assoc) hoelzl@37489: apply (subst setsum_commute) hoelzl@37489: apply simp hoelzl@37489: done hoelzl@37489: hoelzl@37489: lemma matrix_vector_mul_lid: "mat 1 *v x = (x::'a::semiring_1 ^ 'n)" hoelzl@37489: apply (vector matrix_vector_mult_def mat_def) hoelzl@37489: by (simp add: if_distrib cond_application_beta hoelzl@37489: setsum_delta' cong del: if_weak_cong) hoelzl@37489: hoelzl@37489: lemma matrix_transpose_mul: "transpose(A ** B) = transpose B ** transpose (A::'a::comm_semiring_1^_^_)" hoelzl@37489: by (simp add: matrix_matrix_mult_def transpose_def Cart_eq mult_commute) hoelzl@37489: hoelzl@37489: lemma matrix_eq: hoelzl@37489: fixes A B :: "'a::semiring_1 ^ 'n ^ 'm" hoelzl@37489: shows "A = B \ (\x. A *v x = B *v x)" (is "?lhs \ ?rhs") hoelzl@37489: apply auto hoelzl@37489: apply (subst Cart_eq) hoelzl@37489: apply clarify hoelzl@37489: apply (clarsimp simp add: matrix_vector_mult_def cart_basis_def if_distrib cond_application_beta Cart_eq cong del: if_weak_cong) hoelzl@37489: apply (erule_tac x="cart_basis ia" in allE) hoelzl@37489: apply (erule_tac x="i" in allE) hoelzl@37489: by (auto simp add: cart_basis_def if_distrib cond_application_beta setsum_delta[OF finite] cong del: if_weak_cong) hoelzl@37489: hoelzl@37489: lemma matrix_vector_mul_component: hoelzl@37489: shows "((A::real^_^_) *v x)$k = (A$k) \ x" hoelzl@37489: by (simp add: matrix_vector_mult_def inner_vector_def) hoelzl@37489: hoelzl@37489: lemma dot_lmul_matrix: "((x::real ^_) v* A) \ y = x \ (A *v y)" hoelzl@37489: apply (simp add: inner_vector_def matrix_vector_mult_def vector_matrix_mult_def setsum_left_distrib setsum_right_distrib mult_ac) hoelzl@37489: apply (subst setsum_commute) hoelzl@37489: by simp hoelzl@37489: hoelzl@37489: lemma transpose_mat: "transpose (mat n) = mat n" hoelzl@37489: by (vector transpose_def mat_def) hoelzl@37489: hoelzl@37489: lemma transpose_transpose: "transpose(transpose A) = A" hoelzl@37489: by (vector transpose_def) hoelzl@37489: hoelzl@37489: lemma row_transpose: hoelzl@37489: fixes A:: "'a::semiring_1^_^_" hoelzl@37489: shows "row i (transpose A) = column i A" hoelzl@37489: by (simp add: row_def column_def transpose_def Cart_eq) hoelzl@37489: hoelzl@37489: lemma column_transpose: hoelzl@37489: fixes A:: "'a::semiring_1^_^_" hoelzl@37489: shows "column i (transpose A) = row i A" hoelzl@37489: by (simp add: row_def column_def transpose_def Cart_eq) hoelzl@37489: hoelzl@37489: lemma rows_transpose: "rows(transpose (A::'a::semiring_1^_^_)) = columns A" nipkow@39535: by (auto simp add: rows_def columns_def row_transpose intro: set_eqI) hoelzl@37489: hoelzl@37489: lemma columns_transpose: "columns(transpose (A::'a::semiring_1^_^_)) = rows A" by (metis transpose_transpose rows_transpose) hoelzl@37489: hoelzl@37489: text{* Two sometimes fruitful ways of looking at matrix-vector multiplication. *} hoelzl@37489: hoelzl@37489: lemma matrix_mult_dot: "A *v x = (\ i. A$i \ x)" hoelzl@37489: by (simp add: matrix_vector_mult_def inner_vector_def) hoelzl@37489: hoelzl@37489: lemma matrix_mult_vsum: "(A::'a::comm_semiring_1^'n^'m) *v x = setsum (\i. (x$i) *s column i A) (UNIV:: 'n set)" hoelzl@37489: by (simp add: matrix_vector_mult_def Cart_eq column_def mult_commute) hoelzl@37489: hoelzl@37489: lemma vector_componentwise: hoelzl@37489: "(x::'a::ring_1^'n) = (\ j. setsum (\i. (x$i) * (cart_basis i :: 'a^'n)$j) (UNIV :: 'n set))" hoelzl@37489: apply (subst basis_expansion[symmetric]) hoelzl@37489: by (vector Cart_eq setsum_component) hoelzl@37489: hoelzl@37489: lemma linear_componentwise: hoelzl@37489: fixes f:: "real ^'m \ real ^ _" hoelzl@37489: assumes lf: "linear f" hoelzl@37489: shows "(f x)$j = setsum (\i. (x$i) * (f (cart_basis i)$j)) (UNIV :: 'm set)" (is "?lhs = ?rhs") hoelzl@37489: proof- hoelzl@37489: let ?M = "(UNIV :: 'm set)" hoelzl@37489: let ?N = "(UNIV :: 'n set)" hoelzl@37489: have fM: "finite ?M" by simp hoelzl@37489: have "?rhs = (setsum (\i.(x$i) *\<^sub>R f (cart_basis i) ) ?M)$j" hoelzl@37489: unfolding vector_smult_component[symmetric] smult_conv_scaleR hoelzl@37489: unfolding setsum_component[of "(\i.(x$i) *\<^sub>R f (cart_basis i :: real^'m))" ?M] hoelzl@37489: .. hoelzl@37489: then show ?thesis unfolding linear_setsum_mul[OF lf fM, symmetric] basis_expansion' .. hoelzl@37489: qed hoelzl@37489: hoelzl@37489: text{* Inverse matrices (not necessarily square) *} hoelzl@37489: hoelzl@37489: definition "invertible(A::'a::semiring_1^'n^'m) \ (\A'::'a^'m^'n. A ** A' = mat 1 \ A' ** A = mat 1)" hoelzl@37489: hoelzl@37489: definition "matrix_inv(A:: 'a::semiring_1^'n^'m) = hoelzl@37489: (SOME A'::'a^'m^'n. A ** A' = mat 1 \ A' ** A = mat 1)" hoelzl@37489: hoelzl@37489: text{* Correspondence between matrices and linear operators. *} hoelzl@37489: hoelzl@37489: definition matrix:: "('a::{plus,times, one, zero}^'m \ 'a ^ 'n) \ 'a^'m^'n" hoelzl@37489: where "matrix f = (\ i j. (f(cart_basis j))$i)" hoelzl@37489: hoelzl@37489: lemma matrix_vector_mul_linear: "linear(\x. A *v (x::real ^ _))" hoelzl@37489: by (simp add: linear_def matrix_vector_mult_def Cart_eq field_simps setsum_right_distrib setsum_addf) hoelzl@37489: hoelzl@37489: lemma matrix_works: assumes lf: "linear f" shows "matrix f *v x = f (x::real ^ 'n)" hoelzl@37489: apply (simp add: matrix_def matrix_vector_mult_def Cart_eq mult_commute) hoelzl@37489: apply clarify hoelzl@37489: apply (rule linear_componentwise[OF lf, symmetric]) hoelzl@37489: done hoelzl@37489: hoelzl@37489: lemma matrix_vector_mul: "linear f ==> f = (\x. matrix f *v (x::real ^ 'n))" by (simp add: ext matrix_works) hoelzl@37489: hoelzl@37489: lemma matrix_of_matrix_vector_mul: "matrix(\x. A *v (x :: real ^ 'n)) = A" hoelzl@37489: by (simp add: matrix_eq matrix_vector_mul_linear matrix_works) hoelzl@37489: hoelzl@37489: lemma matrix_compose: hoelzl@37489: assumes lf: "linear (f::real^'n \ real^'m)" hoelzl@37489: and lg: "linear (g::real^'m \ real^_)" hoelzl@37489: shows "matrix (g o f) = matrix g ** matrix f" hoelzl@37489: using lf lg linear_compose[OF lf lg] matrix_works[OF linear_compose[OF lf lg]] hoelzl@37489: by (simp add: matrix_eq matrix_works matrix_vector_mul_assoc[symmetric] o_def) hoelzl@37489: hoelzl@37489: lemma matrix_vector_column:"(A::'a::comm_semiring_1^'n^_) *v x = setsum (\i. (x$i) *s ((transpose A)$i)) (UNIV:: 'n set)" hoelzl@37489: by (simp add: matrix_vector_mult_def transpose_def Cart_eq mult_commute) hoelzl@37489: hoelzl@37489: lemma adjoint_matrix: "adjoint(\x. (A::real^'n^'m) *v x) = (\x. transpose A *v x)" hoelzl@37489: apply (rule adjoint_unique) hoelzl@37489: apply (simp add: transpose_def inner_vector_def matrix_vector_mult_def setsum_left_distrib setsum_right_distrib) hoelzl@37489: apply (subst setsum_commute) hoelzl@37489: apply (auto simp add: mult_ac) hoelzl@37489: done hoelzl@37489: hoelzl@37489: lemma matrix_adjoint: assumes lf: "linear (f :: real^'n \ real ^'m)" hoelzl@37489: shows "matrix(adjoint f) = transpose(matrix f)" hoelzl@37489: apply (subst matrix_vector_mul[OF lf]) hoelzl@37489: unfolding adjoint_matrix matrix_of_matrix_vector_mul .. hoelzl@37489: hoelzl@37494: section {* lambda skolemization on cartesian products *} hoelzl@37489: hoelzl@37489: (* FIXME: rename do choice_cart *) hoelzl@37489: hoelzl@37489: lemma lambda_skolem: "(\i. \x. P i x) \ hoelzl@37494: (\x::'a ^ 'n. \i. P i (x $ i))" (is "?lhs \ ?rhs") hoelzl@37489: proof- hoelzl@37489: let ?S = "(UNIV :: 'n set)" hoelzl@37489: {assume H: "?rhs" hoelzl@37489: then have ?lhs by auto} hoelzl@37489: moreover hoelzl@37489: {assume H: "?lhs" hoelzl@37489: then obtain f where f:"\i. P i (f i)" unfolding choice_iff by metis hoelzl@37489: let ?x = "(\ i. (f i)) :: 'a ^ 'n" hoelzl@37489: {fix i hoelzl@37489: from f have "P i (f i)" by metis hoelzl@37494: then have "P i (?x $ i)" by auto hoelzl@37489: } hoelzl@37489: hence "\i. P i (?x$i)" by metis hoelzl@37489: hence ?rhs by metis } hoelzl@37489: ultimately show ?thesis by metis hoelzl@37489: qed hoelzl@37489: hoelzl@37489: subsection {* Standard bases are a spanning set, and obviously finite. *} hoelzl@37489: hoelzl@37489: lemma span_stdbasis:"span {cart_basis i :: real^'n | i. i \ (UNIV :: 'n set)} = UNIV" nipkow@39535: apply (rule set_eqI) hoelzl@37489: apply auto hoelzl@37489: apply (subst basis_expansion'[symmetric]) hoelzl@37489: apply (rule span_setsum) hoelzl@37489: apply simp hoelzl@37489: apply auto hoelzl@37489: apply (rule span_mul) hoelzl@37489: apply (rule span_superset) hoelzl@37489: apply (auto simp add: Collect_def mem_def) hoelzl@37489: done hoelzl@37489: hoelzl@37489: lemma finite_stdbasis: "finite {cart_basis i ::real^'n |i. i\ (UNIV:: 'n set)}" (is "finite ?S") hoelzl@37489: proof- hoelzl@37489: have eq: "?S = cart_basis ` UNIV" by blast hoelzl@37489: show ?thesis unfolding eq by auto hoelzl@37489: qed hoelzl@37489: hoelzl@37489: lemma card_stdbasis: "card {cart_basis i ::real^'n |i. i\ (UNIV :: 'n set)} = CARD('n)" (is "card ?S = _") hoelzl@37489: proof- hoelzl@37489: have eq: "?S = cart_basis ` UNIV" by blast hoelzl@37489: show ?thesis unfolding eq using card_image[OF basis_inj] by simp hoelzl@37489: qed hoelzl@37489: hoelzl@37489: hoelzl@37489: lemma independent_stdbasis_lemma: hoelzl@37489: assumes x: "(x::real ^ 'n) \ span (cart_basis ` S)" hoelzl@37489: and iS: "i \ S" hoelzl@37489: shows "(x$i) = 0" hoelzl@37489: proof- hoelzl@37489: let ?U = "UNIV :: 'n set" hoelzl@37489: let ?B = "cart_basis ` S" hoelzl@37489: let ?P = "\(x::real^_). \i\ ?U. i \ S \ x$i =0" hoelzl@37489: {fix x::"real^_" assume xS: "x\ ?B" hoelzl@37489: from xS have "?P x" by auto} hoelzl@37489: moreover hoelzl@37489: have "subspace ?P" hoelzl@37489: by (auto simp add: subspace_def Collect_def mem_def) hoelzl@37489: ultimately show ?thesis hoelzl@37489: using x span_induct[of ?B ?P x] iS by blast hoelzl@37489: qed hoelzl@37489: hoelzl@37489: lemma independent_stdbasis: "independent {cart_basis i ::real^'n |i. i\ (UNIV :: 'n set)}" hoelzl@37489: proof- hoelzl@37489: let ?I = "UNIV :: 'n set" hoelzl@37489: let ?b = "cart_basis :: _ \ real ^'n" hoelzl@37489: let ?B = "?b ` ?I" hoelzl@37489: have eq: "{?b i|i. i \ ?I} = ?B" hoelzl@37489: by auto hoelzl@37489: {assume d: "dependent ?B" hoelzl@37489: then obtain k where k: "k \ ?I" "?b k \ span (?B - {?b k})" hoelzl@37489: unfolding dependent_def by auto hoelzl@37489: have eq1: "?B - {?b k} = ?B - ?b ` {k}" by simp hoelzl@37489: have eq2: "?B - {?b k} = ?b ` (?I - {k})" hoelzl@37489: unfolding eq1 hoelzl@37489: apply (rule inj_on_image_set_diff[symmetric]) hoelzl@37489: apply (rule basis_inj) using k(1) by auto hoelzl@37489: from k(2) have th0: "?b k \ span (?b ` (?I - {k}))" unfolding eq2 . hoelzl@37489: from independent_stdbasis_lemma[OF th0, of k, simplified] hoelzl@37489: have False by simp} hoelzl@37489: then show ?thesis unfolding eq dependent_def .. hoelzl@37489: qed hoelzl@37489: hoelzl@37489: lemma vector_sub_project_orthogonal_cart: "(b::real^'n) \ (x - ((b \ x) / (b \ b)) *s b) = 0" hoelzl@37489: unfolding inner_simps smult_conv_scaleR by auto hoelzl@37489: hoelzl@37489: lemma linear_eq_stdbasis_cart: hoelzl@37489: assumes lf: "linear (f::real^'m \ _)" and lg: "linear g" hoelzl@37489: and fg: "\i. f (cart_basis i) = g(cart_basis i)" hoelzl@37489: shows "f = g" hoelzl@37489: proof- hoelzl@37489: let ?U = "UNIV :: 'm set" hoelzl@37489: let ?I = "{cart_basis i:: real^'m|i. i \ ?U}" hoelzl@37489: {fix x assume x: "x \ (UNIV :: (real^'m) set)" hoelzl@37489: from equalityD2[OF span_stdbasis] hoelzl@37489: have IU: " (UNIV :: (real^'m) set) \ span ?I" by blast hoelzl@37489: from linear_eq[OF lf lg IU] fg x hoelzl@37489: have "f x = g x" unfolding Collect_def Ball_def mem_def by metis} hoelzl@37489: then show ?thesis by (auto intro: ext) hoelzl@37489: qed hoelzl@37489: hoelzl@37489: lemma bilinear_eq_stdbasis_cart: hoelzl@37489: assumes bf: "bilinear (f:: real^'m \ real^'n \ _)" hoelzl@37489: and bg: "bilinear g" hoelzl@37489: and fg: "\i j. f (cart_basis i) (cart_basis j) = g (cart_basis i) (cart_basis j)" hoelzl@37489: shows "f = g" hoelzl@37489: proof- hoelzl@37489: 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 hoelzl@37489: from bilinear_eq[OF bf bg equalityD2[OF span_stdbasis] equalityD2[OF span_stdbasis] th] show ?thesis by (blast intro: ext) hoelzl@37489: qed hoelzl@37489: hoelzl@37489: lemma left_invertible_transpose: hoelzl@37489: "(\(B). B ** transpose (A) = mat (1::'a::comm_semiring_1)) \ (\(B). A ** B = mat 1)" hoelzl@37489: by (metis matrix_transpose_mul transpose_mat transpose_transpose) hoelzl@37489: hoelzl@37489: lemma right_invertible_transpose: hoelzl@37489: "(\(B). transpose (A) ** B = mat (1::'a::comm_semiring_1)) \ (\(B). B ** A = mat 1)" hoelzl@37489: by (metis matrix_transpose_mul transpose_mat transpose_transpose) hoelzl@37489: hoelzl@37489: lemma matrix_left_invertible_injective: hoelzl@37489: "(\B. (B::real^'m^'n) ** (A::real^'n^'m) = mat 1) \ (\x y. A *v x = A *v y \ x = y)" hoelzl@37489: proof- hoelzl@37489: {fix B:: "real^'m^'n" and x y assume B: "B ** A = mat 1" and xy: "A *v x = A*v y" hoelzl@37489: from xy have "B*v (A *v x) = B *v (A*v y)" by simp hoelzl@37489: hence "x = y" hoelzl@37489: unfolding matrix_vector_mul_assoc B matrix_vector_mul_lid .} hoelzl@37489: moreover hoelzl@37489: {assume A: "\x y. A *v x = A *v y \ x = y" hoelzl@37489: hence i: "inj (op *v A)" unfolding inj_on_def by auto hoelzl@37489: from linear_injective_left_inverse[OF matrix_vector_mul_linear i] hoelzl@37489: obtain g where g: "linear g" "g o op *v A = id" by blast hoelzl@37489: have "matrix g ** A = mat 1" hoelzl@37489: unfolding matrix_eq matrix_vector_mul_lid matrix_vector_mul_assoc[symmetric] matrix_works[OF g(1)] hoelzl@37489: using g(2) by (simp add: o_def id_def stupid_ext) hoelzl@37489: then have "\B. (B::real ^'m^'n) ** A = mat 1" by blast} hoelzl@37489: ultimately show ?thesis by blast hoelzl@37489: qed hoelzl@37489: hoelzl@37489: lemma matrix_left_invertible_ker: hoelzl@37489: "(\B. (B::real ^'m^'n) ** (A::real^'n^'m) = mat 1) \ (\x. A *v x = 0 \ x = 0)" hoelzl@37489: unfolding matrix_left_invertible_injective hoelzl@37489: using linear_injective_0[OF matrix_vector_mul_linear, of A] hoelzl@37489: by (simp add: inj_on_def) hoelzl@37489: hoelzl@37489: lemma matrix_right_invertible_surjective: hoelzl@37489: "(\B. (A::real^'n^'m) ** (B::real^'m^'n) = mat 1) \ surj (\x. A *v x)" hoelzl@37489: proof- hoelzl@37489: {fix B :: "real ^'m^'n" assume AB: "A ** B = mat 1" hoelzl@37489: {fix x :: "real ^ 'm" hoelzl@37489: have "A *v (B *v x) = x" hoelzl@37489: by (simp add: matrix_vector_mul_lid matrix_vector_mul_assoc AB)} hoelzl@37489: hence "surj (op *v A)" unfolding surj_def by metis } hoelzl@37489: moreover hoelzl@37489: {assume sf: "surj (op *v A)" hoelzl@37489: from linear_surjective_right_inverse[OF matrix_vector_mul_linear sf] hoelzl@37489: obtain g:: "real ^'m \ real ^'n" where g: "linear g" "op *v A o g = id" hoelzl@37489: by blast hoelzl@37489: hoelzl@37489: have "A ** (matrix g) = mat 1" hoelzl@37489: unfolding matrix_eq matrix_vector_mul_lid hoelzl@37489: matrix_vector_mul_assoc[symmetric] matrix_works[OF g(1)] hoelzl@37489: using g(2) unfolding o_def stupid_ext[symmetric] id_def hoelzl@37489: . hoelzl@37489: hence "\B. A ** (B::real^'m^'n) = mat 1" by blast hoelzl@37489: } hoelzl@37489: ultimately show ?thesis unfolding surj_def by blast hoelzl@37489: qed hoelzl@37489: hoelzl@37489: lemma matrix_left_invertible_independent_columns: hoelzl@37489: fixes A :: "real^'n^'m" hoelzl@37489: 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))" hoelzl@37489: (is "?lhs \ ?rhs") hoelzl@37489: proof- hoelzl@37489: let ?U = "UNIV :: 'n set" hoelzl@37489: {assume k: "\x. A *v x = 0 \ x = 0" hoelzl@37489: {fix c i assume c: "setsum (\i. c i *s column i A) ?U = 0" hoelzl@37489: and i: "i \ ?U" hoelzl@37489: let ?x = "\ i. c i" hoelzl@37489: have th0:"A *v ?x = 0" hoelzl@37489: using c hoelzl@37489: unfolding matrix_mult_vsum Cart_eq hoelzl@37489: by auto hoelzl@37489: from k[rule_format, OF th0] i hoelzl@37489: have "c i = 0" by (vector Cart_eq)} hoelzl@37489: hence ?rhs by blast} hoelzl@37489: moreover hoelzl@37489: {assume H: ?rhs hoelzl@37489: {fix x assume x: "A *v x = 0" hoelzl@37489: let ?c = "\i. ((x$i ):: real)" hoelzl@37489: from H[rule_format, of ?c, unfolded matrix_mult_vsum[symmetric], OF x] hoelzl@37489: have "x = 0" by vector}} hoelzl@37489: ultimately show ?thesis unfolding matrix_left_invertible_ker by blast hoelzl@37489: qed hoelzl@37489: hoelzl@37489: lemma matrix_right_invertible_independent_rows: hoelzl@37489: fixes A :: "real^'n^'m" hoelzl@37489: 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))" hoelzl@37489: unfolding left_invertible_transpose[symmetric] hoelzl@37489: matrix_left_invertible_independent_columns hoelzl@37489: by (simp add: column_transpose) hoelzl@37489: hoelzl@37489: lemma matrix_right_invertible_span_columns: hoelzl@37489: "(\(B::real ^'n^'m). (A::real ^'m^'n) ** B = mat 1) \ span (columns A) = UNIV" (is "?lhs = ?rhs") hoelzl@37489: proof- hoelzl@37489: let ?U = "UNIV :: 'm set" hoelzl@37489: have fU: "finite ?U" by simp hoelzl@37489: have lhseq: "?lhs \ (\y. \(x::real^'m). setsum (\i. (x$i) *s column i A) ?U = y)" hoelzl@37489: unfolding matrix_right_invertible_surjective matrix_mult_vsum surj_def hoelzl@37489: apply (subst eq_commute) .. hoelzl@37489: have rhseq: "?rhs \ (\x. x \ span (columns A))" by blast hoelzl@37489: {assume h: ?lhs hoelzl@37489: {fix x:: "real ^'n" hoelzl@37489: from h[unfolded lhseq, rule_format, of x] obtain y:: "real ^'m" hoelzl@37489: where y: "setsum (\i. (y$i) *s column i A) ?U = x" by blast hoelzl@37489: have "x \ span (columns A)" hoelzl@37489: unfolding y[symmetric] hoelzl@37489: apply (rule span_setsum[OF fU]) hoelzl@37489: apply clarify hoelzl@37489: unfolding smult_conv_scaleR hoelzl@37489: apply (rule span_mul) hoelzl@37489: apply (rule span_superset) hoelzl@37489: unfolding columns_def hoelzl@37489: by blast} hoelzl@37489: then have ?rhs unfolding rhseq by blast} hoelzl@37489: moreover hoelzl@37489: {assume h:?rhs hoelzl@37489: let ?P = "\(y::real ^'n). \(x::real^'m). setsum (\i. (x$i) *s column i A) ?U = y" hoelzl@37489: {fix y have "?P y" hoelzl@37489: proof(rule span_induct_alt[of ?P "columns A", folded smult_conv_scaleR]) hoelzl@37489: show "\x\real ^ 'm. setsum (\i. (x$i) *s column i A) ?U = 0" hoelzl@37489: by (rule exI[where x=0], simp) hoelzl@37489: next hoelzl@37489: fix c y1 y2 assume y1: "y1 \ columns A" and y2: "?P y2" hoelzl@37489: from y1 obtain i where i: "i \ ?U" "y1 = column i A" hoelzl@37489: unfolding columns_def by blast hoelzl@37489: from y2 obtain x:: "real ^'m" where hoelzl@37489: x: "setsum (\i. (x$i) *s column i A) ?U = y2" by blast hoelzl@37489: let ?x = "(\ j. if j = i then c + (x$i) else (x$j))::real^'m" hoelzl@37489: show "?P (c*s y1 + y2)" hoelzl@37489: 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) hoelzl@37489: fix j hoelzl@37489: have th: "\xa \ ?U. (if xa = i then (c + (x$i)) * ((column xa A)$j) hoelzl@37489: 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) hoelzl@37489: by (simp add: field_simps) hoelzl@37489: have "setsum (\xa. if xa = i then (c + (x$i)) * ((column xa A)$j) hoelzl@37489: 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" hoelzl@37489: apply (rule setsum_cong[OF refl]) hoelzl@37489: using th by blast hoelzl@37489: 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" hoelzl@37489: by (simp add: setsum_addf) hoelzl@37489: also have "\ = c * ((column i A)$j) + setsum (\xa. ((x$xa) * ((column xa A)$j))) ?U" hoelzl@37489: unfolding setsum_delta[OF fU] hoelzl@37489: using i(1) by simp hoelzl@37489: finally show "setsum (\xa. if xa = i then (c + (x$i)) * ((column xa A)$j) hoelzl@37489: else (x$xa) * ((column xa A$j))) ?U = c * ((column i A)$j) + setsum (\xa. ((x$xa) * ((column xa A)$j))) ?U" . hoelzl@37489: qed hoelzl@37489: next hoelzl@37489: show "y \ span (columns A)" unfolding h by blast hoelzl@37489: qed} hoelzl@37489: then have ?lhs unfolding lhseq ..} hoelzl@37489: ultimately show ?thesis by blast hoelzl@37489: qed hoelzl@37489: hoelzl@37489: lemma matrix_left_invertible_span_rows: hoelzl@37489: "(\(B::real^'m^'n). B ** (A::real^'n^'m) = mat 1) \ span (rows A) = UNIV" hoelzl@37489: unfolding right_invertible_transpose[symmetric] hoelzl@37489: unfolding columns_transpose[symmetric] hoelzl@37489: unfolding matrix_right_invertible_span_columns hoelzl@37489: .. hoelzl@37489: hoelzl@37489: text {* The same result in terms of square matrices. *} hoelzl@37489: hoelzl@37489: lemma matrix_left_right_inverse: hoelzl@37489: fixes A A' :: "real ^'n^'n" hoelzl@37489: shows "A ** A' = mat 1 \ A' ** A = mat 1" hoelzl@37489: proof- hoelzl@37489: {fix A A' :: "real ^'n^'n" assume AA': "A ** A' = mat 1" hoelzl@37489: have sA: "surj (op *v A)" hoelzl@37489: unfolding surj_def hoelzl@37489: apply clarify hoelzl@37489: apply (rule_tac x="(A' *v y)" in exI) hoelzl@37489: by (simp add: matrix_vector_mul_assoc AA' matrix_vector_mul_lid) hoelzl@37489: from linear_surjective_isomorphism[OF matrix_vector_mul_linear sA] hoelzl@37489: obtain f' :: "real ^'n \ real ^'n" hoelzl@37489: where f': "linear f'" "\x. f' (A *v x) = x" "\x. A *v f' x = x" by blast hoelzl@37489: have th: "matrix f' ** A = mat 1" hoelzl@37489: by (simp add: matrix_eq matrix_works[OF f'(1)] matrix_vector_mul_assoc[symmetric] matrix_vector_mul_lid f'(2)[rule_format]) hoelzl@37489: hence "(matrix f' ** A) ** A' = mat 1 ** A'" by simp hoelzl@37489: hence "matrix f' = A'" by (simp add: matrix_mul_assoc[symmetric] AA' matrix_mul_rid matrix_mul_lid) hoelzl@37489: hence "matrix f' ** A = A' ** A" by simp hoelzl@37489: hence "A' ** A = mat 1" by (simp add: th)} hoelzl@37489: then show ?thesis by blast hoelzl@37489: qed hoelzl@37489: hoelzl@37489: text {* Considering an n-element vector as an n-by-1 or 1-by-n matrix. *} hoelzl@37489: hoelzl@37489: definition "rowvector v = (\ i j. (v$j))" hoelzl@37489: hoelzl@37489: definition "columnvector v = (\ i j. (v$i))" hoelzl@37489: hoelzl@37489: lemma transpose_columnvector: hoelzl@37489: "transpose(columnvector v) = rowvector v" hoelzl@37489: by (simp add: transpose_def rowvector_def columnvector_def Cart_eq) hoelzl@37489: hoelzl@37489: lemma transpose_rowvector: "transpose(rowvector v) = columnvector v" hoelzl@37489: by (simp add: transpose_def columnvector_def rowvector_def Cart_eq) hoelzl@37489: hoelzl@37489: lemma dot_rowvector_columnvector: hoelzl@37489: "columnvector (A *v v) = A ** columnvector v" hoelzl@37489: by (vector columnvector_def matrix_matrix_mult_def matrix_vector_mult_def) hoelzl@37489: hoelzl@37489: lemma dot_matrix_product: "(x::real^'n) \ y = (((rowvector x ::real^'n^1) ** (columnvector y :: real^1^'n))$1)$1" hoelzl@37489: by (vector matrix_matrix_mult_def rowvector_def columnvector_def inner_vector_def) hoelzl@37489: hoelzl@37489: lemma dot_matrix_vector_mul: hoelzl@37489: fixes A B :: "real ^'n ^'n" and x y :: "real ^'n" hoelzl@37489: shows "(A *v x) \ (B *v y) = hoelzl@37489: (((rowvector x :: real^'n^1) ** ((transpose A ** B) ** (columnvector y :: real ^1^'n)))$1)$1" hoelzl@37489: unfolding dot_matrix_product transpose_columnvector[symmetric] hoelzl@37489: dot_rowvector_columnvector matrix_transpose_mul matrix_mul_assoc .. hoelzl@37489: hoelzl@37489: hoelzl@37489: lemma infnorm_cart:"infnorm (x::real^'n) = Sup {abs(x$i) |i. i\ (UNIV :: 'n set)}" hoelzl@37489: unfolding infnorm_def apply(rule arg_cong[where f=Sup]) apply safe hoelzl@37489: apply(rule_tac x="\ i" in exI) defer hoelzl@37489: apply(rule_tac x="\' i" in exI) unfolding nth_conv_component using pi'_range by auto hoelzl@37489: hoelzl@37489: lemma infnorm_set_image_cart: hoelzl@37489: "{abs(x$i) |i. i\ (UNIV :: _ set)} = hoelzl@37489: (\i. abs(x$i)) ` (UNIV)" by blast hoelzl@37489: hoelzl@37489: lemma infnorm_set_lemma_cart: hoelzl@37489: shows "finite {abs((x::'a::abs ^'n)$i) |i. i\ (UNIV :: 'n set)}" hoelzl@37489: and "{abs(x$i) |i. i\ (UNIV :: 'n::finite set)} \ {}" hoelzl@37489: unfolding infnorm_set_image_cart nipkow@41030: by auto hoelzl@37489: hoelzl@37489: lemma component_le_infnorm_cart: hoelzl@37489: shows "\x$i\ \ infnorm (x::real^'n)" hoelzl@37489: unfolding nth_conv_component hoelzl@37489: using component_le_infnorm[of x] . hoelzl@37489: hoelzl@37489: instance cart :: (perfect_space, finite) perfect_space hoelzl@37489: proof hoelzl@37489: fix x :: "'a ^ 'b" hoelzl@37489: { hoelzl@37489: fix e :: real assume "0 < e" hoelzl@37489: def a \ "x $ undefined" hoelzl@37489: have "a islimpt UNIV" by (rule islimpt_UNIV) hoelzl@37489: with `0 < e` obtain b where "b \ a" and "dist b a < e" hoelzl@37489: unfolding islimpt_approachable by auto hoelzl@37489: def y \ "Cart_lambda ((Cart_nth x)(undefined := b))" hoelzl@37489: from `b \ a` have "y \ x" hoelzl@37489: unfolding a_def y_def by (simp add: Cart_eq) hoelzl@37489: from `dist b a < e` have "dist y x < e" hoelzl@37489: unfolding dist_vector_def a_def y_def hoelzl@37489: apply simp hoelzl@37489: apply (rule le_less_trans [OF setL2_le_setsum [OF zero_le_dist]]) hoelzl@37489: apply (subst setsum_diff1' [where a=undefined], simp, simp, simp) hoelzl@37489: done hoelzl@37489: from `y \ x` and `dist y x < e` hoelzl@37489: have "\y\UNIV. y \ x \ dist y x < e" by auto hoelzl@37489: } hoelzl@37489: then show "x islimpt UNIV" unfolding islimpt_approachable by blast hoelzl@37489: qed hoelzl@37489: hoelzl@37489: lemma closed_positive_orthant: "closed {x::real^'n. \i. 0 \x$i}" hoelzl@37489: proof- hoelzl@37489: let ?U = "UNIV :: 'n set" hoelzl@37489: let ?O = "{x::real^'n. \i. x$i\0}" hoelzl@37489: {fix x:: "real^'n" and i::'n assume H: "\e>0. \x'\?O. x' \ x \ dist x' x < e" hoelzl@37489: and xi: "x$i < 0" hoelzl@37489: from xi have th0: "-x$i > 0" by arith hoelzl@37489: from H[rule_format, OF th0] obtain x' where x': "x' \?O" "x' \ x" "dist x' x < -x $ i" by blast hoelzl@37489: have th:" \b a (x::real). abs x <= b \ b <= a ==> ~(a + x < 0)" by arith hoelzl@37489: have th': "\x (y::real). x < 0 \ 0 <= y ==> abs x <= abs (y - x)" by arith hoelzl@37489: have th1: "\x$i\ \ \(x' - x)$i\" using x'(1) xi hoelzl@37489: apply (simp only: vector_component) hoelzl@37489: by (rule th') auto hoelzl@37489: have th2: "\dist x x'\ \ \(x' - x)$i\" using component_le_norm_cart[of "x'-x" i] hoelzl@37489: apply (simp add: dist_norm) by norm hoelzl@37489: from th[OF th1 th2] x'(3) have False by (simp add: dist_commute) } hoelzl@37489: then show ?thesis unfolding closed_limpt islimpt_approachable hoelzl@37489: unfolding not_le[symmetric] by blast hoelzl@37489: qed hoelzl@37489: lemma Lim_component_cart: hoelzl@37489: fixes f :: "'a \ 'b::metric_space ^ 'n" hoelzl@37489: shows "(f ---> l) net \ ((\a. f a $i) ---> l$i) net" hoelzl@37489: unfolding tendsto_iff hoelzl@37489: apply (clarify) hoelzl@37489: apply (drule spec, drule (1) mp) hoelzl@37489: apply (erule eventually_elim1) hoelzl@37489: apply (erule le_less_trans [OF dist_nth_le_cart]) hoelzl@37489: done hoelzl@37489: hoelzl@37489: lemma bounded_component_cart: "bounded s \ bounded ((\x. x $ i) ` s)" hoelzl@37489: unfolding bounded_def hoelzl@37489: apply clarify hoelzl@37489: apply (rule_tac x="x $ i" in exI) hoelzl@37489: apply (rule_tac x="e" in exI) hoelzl@37489: apply clarify hoelzl@37489: apply (rule order_trans [OF dist_nth_le_cart], simp) hoelzl@37489: done hoelzl@37489: hoelzl@37489: lemma compact_lemma_cart: hoelzl@37489: fixes f :: "nat \ 'a::heine_borel ^ 'n" hoelzl@37489: assumes "bounded s" and "\n. f n \ s" hoelzl@37489: shows "\d. hoelzl@37489: \l r. subseq r \ hoelzl@37489: (\e>0. eventually (\n. \i\d. dist (f (r n) $ i) (l $ i) < e) sequentially)" hoelzl@37489: proof hoelzl@37489: fix d::"'n set" have "finite d" by simp hoelzl@37489: thus "\l::'a ^ 'n. \r. subseq r \ hoelzl@37489: (\e>0. eventually (\n. \i\d. dist (f (r n) $ i) (l $ i) < e) sequentially)" hoelzl@37489: proof(induct d) case empty thus ?case unfolding subseq_def by auto hoelzl@37489: next case (insert k d) hoelzl@37489: have s': "bounded ((\x. x $ k) ` s)" using `bounded s` by (rule bounded_component_cart) hoelzl@37489: 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" hoelzl@37489: using insert(3) by auto hoelzl@37489: have f': "\n. f (r1 n) $ k \ (\x. x $ k) ` s" using `\n. f n \ s` by simp hoelzl@37489: obtain l2 r2 where r2:"subseq r2" and lr2:"((\i. f (r1 (r2 i)) $ k) ---> l2) sequentially" hoelzl@37489: using bounded_imp_convergent_subsequence[OF s' f'] unfolding o_def by auto hoelzl@37489: def r \ "r1 \ r2" have r:"subseq r" hoelzl@37489: using r1 and r2 unfolding r_def o_def subseq_def by auto hoelzl@37489: moreover hoelzl@37489: def l \ "(\ i. if i = k then l2 else l1$i)::'a^'n" hoelzl@37489: { fix e::real assume "e>0" hoelzl@37489: from lr1 `e>0` have N1:"eventually (\n. \i\d. dist (f (r1 n) $ i) (l1 $ i) < e) sequentially" by blast hoelzl@37489: from lr2 `e>0` have N2:"eventually (\n. dist (f (r1 (r2 n)) $ k) l2 < e) sequentially" by (rule tendstoD) hoelzl@37489: from r2 N1 have N1': "eventually (\n. \i\d. dist (f (r1 (r2 n)) $ i) (l1 $ i) < e) sequentially" hoelzl@37489: by (rule eventually_subseq) hoelzl@37489: have "eventually (\n. \i\(insert k d). dist (f (r n) $ i) (l $ i) < e) sequentially" hoelzl@37489: using N1' N2 by (rule eventually_elim2, simp add: l_def r_def) hoelzl@37489: } hoelzl@37489: ultimately show ?case by auto hoelzl@37489: qed hoelzl@37489: qed hoelzl@37489: hoelzl@37489: instance cart :: (heine_borel, finite) heine_borel hoelzl@37489: proof hoelzl@37489: fix s :: "('a ^ 'b) set" and f :: "nat \ 'a ^ 'b" hoelzl@37489: assume s: "bounded s" and f: "\n. f n \ s" hoelzl@37489: then obtain l r where r: "subseq r" hoelzl@37489: and l: "\e>0. eventually (\n. \i\UNIV. dist (f (r n) $ i) (l $ i) < e) sequentially" hoelzl@37489: using compact_lemma_cart [OF s f] by blast hoelzl@37489: let ?d = "UNIV::'b set" hoelzl@37489: { fix e::real assume "e>0" hoelzl@37489: hence "0 < e / (real_of_nat (card ?d))" hoelzl@37489: using zero_less_card_finite using divide_pos_pos[of e, of "real_of_nat (card ?d)"] by auto hoelzl@37489: with l have "eventually (\n. \i. dist (f (r n) $ i) (l $ i) < e / (real_of_nat (card ?d))) sequentially" hoelzl@37489: by simp hoelzl@37489: moreover hoelzl@37489: { fix n assume n: "\i. dist (f (r n) $ i) (l $ i) < e / (real_of_nat (card ?d))" hoelzl@37489: have "dist (f (r n)) l \ (\i\?d. dist (f (r n) $ i) (l $ i))" hoelzl@37489: unfolding dist_vector_def using zero_le_dist by (rule setL2_le_setsum) hoelzl@37489: also have "\ < (\i\?d. e / (real_of_nat (card ?d)))" hoelzl@37489: by (rule setsum_strict_mono) (simp_all add: n) hoelzl@37489: finally have "dist (f (r n)) l < e" by simp hoelzl@37489: } hoelzl@37489: ultimately have "eventually (\n. dist (f (r n)) l < e) sequentially" hoelzl@37489: by (rule eventually_elim1) hoelzl@37489: } hoelzl@37489: hence *:"((f \ r) ---> l) sequentially" unfolding o_def tendsto_iff by simp hoelzl@37489: with r show "\l r. subseq r \ ((f \ r) ---> l) sequentially" by auto hoelzl@37489: qed hoelzl@37489: hoelzl@37489: lemma continuous_at_component: "continuous (at a) (\x. x $ i)" hoelzl@37489: unfolding continuous_at by (intro tendsto_intros) hoelzl@37489: hoelzl@37489: lemma continuous_on_component: "continuous_on s (\x. x $ i)" hoelzl@37489: unfolding continuous_on_def by (intro ballI tendsto_intros) hoelzl@37489: hoelzl@37489: lemma interval_cart: fixes a :: "'a::ord^'n" shows hoelzl@37489: "{a <..< b} = {x::'a^'n. \i. a$i < x$i \ x$i < b$i}" and hoelzl@37489: "{a .. b} = {x::'a^'n. \i. a$i \ x$i \ x$i \ b$i}" nipkow@39535: by (auto simp add: set_eq_iff vector_less_def vector_le_def) hoelzl@37489: hoelzl@37489: lemma mem_interval_cart: fixes a :: "'a::ord^'n" shows hoelzl@37489: "x \ {a<.. (\i. a$i < x$i \ x$i < b$i)" hoelzl@37489: "x \ {a .. b} \ (\i. a$i \ x$i \ x$i \ b$i)" nipkow@39535: using interval_cart[of a b] by(auto simp add: set_eq_iff vector_less_def vector_le_def) hoelzl@37489: hoelzl@37489: lemma interval_eq_empty_cart: fixes a :: "real^'n" shows hoelzl@37489: "({a <..< b} = {} \ (\i. b$i \ a$i))" (is ?th1) and hoelzl@37489: "({a .. b} = {} \ (\i. b$i < a$i))" (is ?th2) hoelzl@37489: proof- hoelzl@37489: { fix i x assume as:"b$i \ a$i" and x:"x\{a <..< b}" hoelzl@37489: hence "a $ i < x $ i \ x $ i < b $ i" unfolding mem_interval_cart by auto hoelzl@37489: hence "a$i < b$i" by auto hoelzl@37489: hence False using as by auto } hoelzl@37489: moreover hoelzl@37489: { assume as:"\i. \ (b$i \ a$i)" hoelzl@37489: let ?x = "(1/2) *\<^sub>R (a + b)" hoelzl@37489: { fix i hoelzl@37489: have "a$i < b$i" using as[THEN spec[where x=i]] by auto hoelzl@37489: hence "a$i < ((1/2) *\<^sub>R (a+b)) $ i" "((1/2) *\<^sub>R (a+b)) $ i < b$i" hoelzl@37489: unfolding vector_smult_component and vector_add_component hoelzl@37489: by auto } hoelzl@37489: hence "{a <..< b} \ {}" using mem_interval_cart(1)[of "?x" a b] by auto } hoelzl@37489: ultimately show ?th1 by blast hoelzl@37489: hoelzl@37489: { fix i x assume as:"b$i < a$i" and x:"x\{a .. b}" hoelzl@37489: hence "a $ i \ x $ i \ x $ i \ b $ i" unfolding mem_interval_cart by auto hoelzl@37489: hence "a$i \ b$i" by auto hoelzl@37489: hence False using as by auto } hoelzl@37489: moreover hoelzl@37489: { assume as:"\i. \ (b$i < a$i)" hoelzl@37489: let ?x = "(1/2) *\<^sub>R (a + b)" hoelzl@37489: { fix i hoelzl@37489: have "a$i \ b$i" using as[THEN spec[where x=i]] by auto hoelzl@37489: hence "a$i \ ((1/2) *\<^sub>R (a+b)) $ i" "((1/2) *\<^sub>R (a+b)) $ i \ b$i" hoelzl@37489: unfolding vector_smult_component and vector_add_component hoelzl@37489: by auto } hoelzl@37489: hence "{a .. b} \ {}" using mem_interval_cart(2)[of "?x" a b] by auto } hoelzl@37489: ultimately show ?th2 by blast hoelzl@37489: qed hoelzl@37489: hoelzl@37489: lemma interval_ne_empty_cart: fixes a :: "real^'n" shows hoelzl@37489: "{a .. b} \ {} \ (\i. a$i \ b$i)" and hoelzl@37489: "{a <..< b} \ {} \ (\i. a$i < b$i)" hoelzl@37489: unfolding interval_eq_empty_cart[of a b] by (auto simp add: not_less not_le) hoelzl@37489: (* BH: Why doesn't just "auto" work here? *) hoelzl@37489: hoelzl@37489: lemma subset_interval_imp_cart: fixes a :: "real^'n" shows hoelzl@37489: "(\i. a$i \ c$i \ d$i \ b$i) \ {c .. d} \ {a .. b}" and hoelzl@37489: "(\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 hoelzl@37489: "(\i. a$i \ c$i \ d$i \ b$i) \ {c<.. {a<.. {a<.. {a .. b}" hoelzl@37489: proof(simp add: subset_eq, rule) hoelzl@37489: fix x hoelzl@37489: assume x:"x \{a<.. x $ i" hoelzl@37489: using x order_less_imp_le[of "a$i" "x$i"] nipkow@39535: by(simp add: set_eq_iff vector_less_def vector_le_def Cart_eq) hoelzl@37489: } hoelzl@37489: moreover hoelzl@37489: { fix i hoelzl@37489: have "x $ i \ b $ i" hoelzl@37489: using x order_less_imp_le[of "x$i" "b$i"] nipkow@39535: by(simp add: set_eq_iff vector_less_def vector_le_def Cart_eq) hoelzl@37489: } hoelzl@37489: ultimately hoelzl@37489: show "a \ x \ x \ b" nipkow@39535: by(simp add: set_eq_iff vector_less_def vector_le_def Cart_eq) hoelzl@37489: qed hoelzl@37489: hoelzl@37489: lemma subset_interval_cart: fixes a :: "real^'n" shows hoelzl@37489: "{c .. d} \ {a .. b} \ (\i. c$i \ d$i) --> (\i. a$i \ c$i \ d$i \ b$i)" (is ?th1) and hoelzl@37489: "{c .. d} \ {a<.. (\i. c$i \ d$i) --> (\i. a$i < c$i \ d$i < b$i)" (is ?th2) and hoelzl@37489: "{c<.. {a .. b} \ (\i. c$i < d$i) --> (\i. a$i \ c$i \ d$i \ b$i)" (is ?th3) and hoelzl@37489: "{c<.. {a<.. (\i. c$i < d$i) --> (\i. a$i \ c$i \ d$i \ b$i)" (is ?th4) hoelzl@37489: using subset_interval[of c d a b] by (simp_all add: cart_simps real_euclidean_nth) hoelzl@37489: hoelzl@37489: lemma disjoint_interval_cart: fixes a::"real^'n" shows hoelzl@37489: "{a .. b} \ {c .. d} = {} \ (\i. (b$i < a$i \ d$i < c$i \ b$i < c$i \ d$i < a$i))" (is ?th1) and hoelzl@37489: "{a .. b} \ {c<.. (\i. (b$i < a$i \ d$i \ c$i \ b$i \ c$i \ d$i \ a$i))" (is ?th2) and hoelzl@37489: "{a<.. {c .. d} = {} \ (\i. (b$i \ a$i \ d$i < c$i \ b$i \ c$i \ d$i \ a$i))" (is ?th3) and hoelzl@37489: "{a<.. {c<.. (\i. (b$i \ a$i \ d$i \ c$i \ b$i \ c$i \ d$i \ a$i))" (is ?th4) hoelzl@37489: using disjoint_interval[of a b c d] by (simp_all add: cart_simps real_euclidean_nth) hoelzl@37489: hoelzl@37489: lemma inter_interval_cart: fixes a :: "'a::linorder^'n" shows hoelzl@37489: "{a .. b} \ {c .. d} = {(\ i. max (a$i) (c$i)) .. (\ i. min (b$i) (d$i))}" nipkow@39535: unfolding set_eq_iff and Int_iff and mem_interval_cart hoelzl@37489: by auto hoelzl@37489: hoelzl@37489: lemma closed_interval_left_cart: fixes b::"real^'n" hoelzl@37489: shows "closed {x::real^'n. \i. x$i \ b$i}" hoelzl@37489: proof- hoelzl@37489: { fix i hoelzl@37489: fix x::"real^'n" assume x:"\e>0. \x'\{x. \i. x $ i \ b $ i}. x' \ x \ dist x' x < e" hoelzl@37489: { assume "x$i > b$i" hoelzl@37489: then obtain y where "y $ i \ b $ i" "y \ x" "dist y x < x$i - b$i" using x[THEN spec[where x="x$i - b$i"]] by auto hoelzl@37489: hence False using component_le_norm_cart[of "y - x" i] unfolding dist_norm and vector_minus_component by auto } hoelzl@37489: hence "x$i \ b$i" by(rule ccontr)auto } hoelzl@37489: thus ?thesis unfolding closed_limpt unfolding islimpt_approachable by blast hoelzl@37489: qed hoelzl@37489: hoelzl@37489: lemma closed_interval_right_cart: fixes a::"real^'n" hoelzl@37489: shows "closed {x::real^'n. \i. a$i \ x$i}" hoelzl@37489: proof- hoelzl@37489: { fix i hoelzl@37489: fix x::"real^'n" assume x:"\e>0. \x'\{x. \i. a $ i \ x $ i}. x' \ x \ dist x' x < e" hoelzl@37489: { assume "a$i > x$i" hoelzl@37489: then obtain y where "a $ i \ y $ i" "y \ x" "dist y x < a$i - x$i" using x[THEN spec[where x="a$i - x$i"]] by auto hoelzl@37489: hence False using component_le_norm_cart[of "y - x" i] unfolding dist_norm and vector_minus_component by auto } hoelzl@37489: hence "a$i \ x$i" by(rule ccontr)auto } hoelzl@37489: thus ?thesis unfolding closed_limpt unfolding islimpt_approachable by blast hoelzl@37489: qed hoelzl@37489: hoelzl@37489: lemma is_interval_cart:"is_interval (s::(real^'n) set) \ hoelzl@37489: (\a\s. \b\s. \x. (\i. ((a$i \ x$i \ x$i \ b$i) \ (b$i \ x$i \ x$i \ a$i))) \ x \ s)" hoelzl@37489: unfolding is_interval_def Ball_def by (simp add: cart_simps real_euclidean_nth) hoelzl@37489: hoelzl@37489: lemma closed_halfspace_component_le_cart: hoelzl@37489: shows "closed {x::real^'n. x$i \ a}" hoelzl@37489: using closed_halfspace_le[of "(cart_basis i)::real^'n" a] unfolding inner_basis[OF assms] by auto hoelzl@37489: hoelzl@37489: lemma closed_halfspace_component_ge_cart: hoelzl@37489: shows "closed {x::real^'n. x$i \ a}" hoelzl@37489: using closed_halfspace_ge[of a "(cart_basis i)::real^'n"] unfolding inner_basis[OF assms] by auto hoelzl@37489: hoelzl@37489: lemma open_halfspace_component_lt_cart: hoelzl@37489: shows "open {x::real^'n. x$i < a}" hoelzl@37489: using open_halfspace_lt[of "(cart_basis i)::real^'n" a] unfolding inner_basis[OF assms] by auto hoelzl@37489: hoelzl@37489: lemma open_halfspace_component_gt_cart: hoelzl@37489: shows "open {x::real^'n. x$i > a}" hoelzl@37489: using open_halfspace_gt[of a "(cart_basis i)::real^'n"] unfolding inner_basis[OF assms] by auto hoelzl@37489: hoelzl@37489: lemma Lim_component_le_cart: fixes f :: "'a \ real^'n" hoelzl@37489: assumes "(f ---> l) net" "\ (trivial_limit net)" "eventually (\x. f(x)$i \ b) net" hoelzl@37489: shows "l$i \ b" hoelzl@37489: proof- hoelzl@37489: { fix x have "x \ {x::real^'n. inner (cart_basis i) x \ b} \ x$i \ b" unfolding inner_basis by auto } note * = this hoelzl@37489: show ?thesis using Lim_in_closed_set[of "{x. inner (cart_basis i) x \ b}" f net l] unfolding * hoelzl@37489: using closed_halfspace_le[of "(cart_basis i)::real^'n" b] and assms(1,2,3) by auto hoelzl@37489: qed hoelzl@37489: hoelzl@37489: lemma Lim_component_ge_cart: fixes f :: "'a \ real^'n" hoelzl@37489: assumes "(f ---> l) net" "\ (trivial_limit net)" "eventually (\x. b \ (f x)$i) net" hoelzl@37489: shows "b \ l$i" hoelzl@37489: proof- hoelzl@37489: { fix x have "x \ {x::real^'n. inner (cart_basis i) x \ b} \ x$i \ b" unfolding inner_basis by auto } note * = this hoelzl@37489: show ?thesis using Lim_in_closed_set[of "{x. inner (cart_basis i) x \ b}" f net l] unfolding * hoelzl@37489: using closed_halfspace_ge[of b "(cart_basis i)::real^'n"] and assms(1,2,3) by auto hoelzl@37489: qed hoelzl@37489: hoelzl@37489: lemma Lim_component_eq_cart: fixes f :: "'a \ real^'n" hoelzl@37489: assumes net:"(f ---> l) net" "~(trivial_limit net)" and ev:"eventually (\x. f(x)$i = b) net" hoelzl@37489: shows "l$i = b" hoelzl@37489: using ev[unfolded order_eq_iff eventually_and] using Lim_component_ge_cart[OF net, of b i] and hoelzl@37489: Lim_component_le_cart[OF net, of i b] by auto hoelzl@37489: hoelzl@37489: lemma connected_ivt_component_cart: fixes x::"real^'n" shows hoelzl@37489: "connected s \ x \ s \ y \ s \ x$k \ a \ a \ y$k \ (\z\s. z$k = a)" hoelzl@37489: using connected_ivt_hyperplane[of s x y "(cart_basis k)::real^'n" a] by (auto simp add: inner_basis) hoelzl@37489: hoelzl@37489: lemma subspace_substandard_cart: hoelzl@37489: "subspace {x::real^_. (\i. P i \ x$i = 0)}" hoelzl@37489: unfolding subspace_def by auto hoelzl@37489: hoelzl@37489: lemma closed_substandard_cart: hoelzl@37489: "closed {x::real^'n. \i. P i --> x$i = 0}" (is "closed ?A") hoelzl@37489: proof- hoelzl@37489: let ?D = "{i. P i}" hoelzl@37489: let ?Bs = "{{x::real^'n. inner (cart_basis i) x = 0}| i. i \ ?D}" hoelzl@37489: { fix x hoelzl@37489: { assume "x\?A" hoelzl@37489: hence x:"\i\?D. x $ i = 0" by auto hoelzl@37489: hence "x\ \ ?Bs" by(auto simp add: inner_basis x) } hoelzl@37489: moreover hoelzl@37489: { assume x:"x\\?Bs" hoelzl@37489: { fix i assume i:"i \ ?D" hoelzl@37489: then obtain B where BB:"B \ ?Bs" and B:"B = {x::real^'n. inner (cart_basis i) x = 0}" by auto hoelzl@37489: hence "x $ i = 0" unfolding B using x unfolding inner_basis by auto } hoelzl@37489: hence "x\?A" by auto } hoelzl@37489: ultimately have "x\?A \ x\ \?Bs" .. } hoelzl@37489: hence "?A = \ ?Bs" by auto hoelzl@37489: thus ?thesis by(auto simp add: closed_Inter closed_hyperplane) hoelzl@37489: qed hoelzl@37489: hoelzl@37489: lemma dim_substandard_cart: hoelzl@37489: shows "dim {x::real^'n. \i. i \ d \ x$i = 0} = card d" (is "dim ?A = _") hoelzl@37489: proof- have *:"{x. \i \' ` d \ x $$ i = 0} = hoelzl@37489: {x::real^'n. \i. i \ d \ x$i = 0}"apply safe hoelzl@37489: apply(erule_tac x="\' i" in allE) defer hoelzl@37489: apply(erule_tac x="\ i" in allE) hoelzl@37489: unfolding image_iff real_euclidean_nth[symmetric] by (auto simp: pi'_inj[THEN inj_eq]) hoelzl@37489: have " \' ` d \ {..'" d] using pi'_inj unfolding inj_on_def by auto hoelzl@37489: qed hoelzl@37489: hoelzl@37489: lemma affinity_inverses: hoelzl@37489: assumes m0: "m \ (0::'a::field)" hoelzl@37489: shows "(\x. m *s x + c) o (\x. inverse(m) *s x + (-(inverse(m) *s c))) = id" hoelzl@37489: "(\x. inverse(m) *s x + (-(inverse(m) *s c))) o (\x. m *s x + c) = id" hoelzl@37489: using m0 nipkow@39535: apply (auto simp add: fun_eq_iff vector_add_ldistrib) hoelzl@37489: by (simp_all add: vector_smult_lneg[symmetric] vector_smult_assoc vector_sneg_minus1[symmetric]) hoelzl@37489: hoelzl@37489: lemma vector_affinity_eq: hoelzl@37489: assumes m0: "(m::'a::field) \ 0" hoelzl@37489: shows "m *s x + c = y \ x = inverse m *s y + -(inverse m *s c)" hoelzl@37489: proof hoelzl@37489: assume h: "m *s x + c = y" hoelzl@37489: hence "m *s x = y - c" by (simp add: field_simps) hoelzl@37489: hence "inverse m *s (m *s x) = inverse m *s (y - c)" by simp hoelzl@37489: then show "x = inverse m *s y + - (inverse m *s c)" hoelzl@37489: using m0 by (simp add: vector_smult_assoc vector_ssub_ldistrib) hoelzl@37489: next hoelzl@37489: assume h: "x = inverse m *s y + - (inverse m *s c)" hoelzl@37489: show "m *s x + c = y" unfolding h diff_minus[symmetric] hoelzl@37489: using m0 by (simp add: vector_smult_assoc vector_ssub_ldistrib) hoelzl@37489: qed hoelzl@37489: hoelzl@37489: lemma vector_eq_affinity: hoelzl@37489: "(m::'a::field) \ 0 ==> (y = m *s x + c \ inverse(m) *s y + -(inverse(m) *s c) = x)" hoelzl@37489: using vector_affinity_eq[where m=m and x=x and y=y and c=c] hoelzl@37489: by metis hoelzl@37489: hoelzl@37489: lemma const_vector_cart:"((\ i. d)::real^'n) = (\\ i. d)" hoelzl@37489: apply(subst euclidean_eq) hoelzl@37489: proof safe case goal1 hoelzl@37489: hence *:"(basis i::real^'n) = cart_basis (\ i)" hoelzl@37489: unfolding basis_real_n[THEN sym] by auto hoelzl@37489: have "((\ i. d)::real^'n) $$ i = d" unfolding euclidean_component_def * hoelzl@37489: unfolding dot_basis by auto hoelzl@37489: thus ?case using goal1 by auto hoelzl@37489: qed hoelzl@37489: hoelzl@37489: section "Convex Euclidean Space" hoelzl@37489: hoelzl@37489: lemma Cart_1:"(1::real^'n) = (\\ i. 1)" hoelzl@37489: apply(subst euclidean_eq) hoelzl@37489: proof safe case goal1 thus ?case using nth_conv_component[THEN sym,where i1="\ i" and x1="1::real^'n"] by auto hoelzl@37489: qed hoelzl@37489: hoelzl@37489: declare vector_add_ldistrib[simp] vector_ssub_ldistrib[simp] vector_smult_assoc[simp] vector_smult_rneg[simp] hoelzl@37489: declare vector_sadd_rdistrib[simp] vector_sub_rdistrib[simp] hoelzl@37489: hoelzl@37489: lemmas vector_component_simps = vector_minus_component vector_smult_component vector_add_component vector_le_def Cart_lambda_beta basis_component vector_uminus_component hoelzl@37489: hoelzl@37489: lemma convex_box_cart: hoelzl@37489: assumes "\i. convex {x. P i x}" hoelzl@37489: shows "convex {x. \i. P i (x$i)}" hoelzl@37489: using assms unfolding convex_def by auto hoelzl@37489: hoelzl@37489: lemma convex_positive_orthant_cart: "convex {x::real^'n. (\i. 0 \ x$i)}" hoelzl@37489: by (rule convex_box_cart) (simp add: atLeast_def[symmetric] convex_real_interval) hoelzl@37489: hoelzl@37489: lemma unit_interval_convex_hull_cart: hoelzl@37489: "{0::real^'n .. 1} = convex hull {x. \i. (x$i = 0) \ (x$i = 1)}" (is "?int = convex hull ?points") hoelzl@37489: unfolding Cart_1 unit_interval_convex_hull[where 'a="real^'n"] nipkow@39535: apply(rule arg_cong[where f="\x. convex hull x"]) apply(rule set_eqI) unfolding mem_Collect_eq hoelzl@37489: apply safe apply(erule_tac x="\' i" in allE) unfolding nth_conv_component defer hoelzl@37489: apply(erule_tac x="\ i" in allE) by auto hoelzl@37489: hoelzl@37489: lemma cube_convex_hull_cart: hoelzl@37489: assumes "0 < d" obtains s::"(real^'n) set" where "finite s" "{x - (\ i. d) .. x + (\ i. d)} = convex hull s" hoelzl@37489: proof- from cube_convex_hull[OF assms, where 'a="real^'n" and x=x] guess s . note s=this hoelzl@37489: show thesis apply(rule that[OF s(1)]) unfolding s(2)[THEN sym] const_vector_cart .. hoelzl@37489: qed hoelzl@37489: hoelzl@37489: lemma std_simplex_cart: hoelzl@37489: "(insert (0::real^'n) { cart_basis i | i. i\UNIV}) = hoelzl@37489: (insert 0 { basis i | i. i i. u$i + (x$i - a$i) / (b$i - a$i) * (v$i - u$i))::real^'n)" hoelzl@37489: unfolding interval_bij_def apply(rule ext)+ apply safe hoelzl@37489: unfolding Cart_eq Cart_lambda_beta unfolding nth_conv_component hoelzl@37489: apply rule apply(subst euclidean_lambda_beta) using pi'_range by auto hoelzl@37489: hoelzl@37489: lemma interval_bij_affine_cart: hoelzl@37489: "interval_bij (a,b) (u,v) = (\x. (\ i. (v$i - u$i) / (b$i - a$i) * x$i) + hoelzl@37489: (\ i. u$i - (v$i - u$i) / (b$i - a$i) * a$i)::real^'n)" hoelzl@37489: apply rule unfolding Cart_eq interval_bij_cart vector_component_simps hoelzl@37489: by(auto simp add: field_simps add_divide_distrib[THEN sym]) hoelzl@37489: hoelzl@37489: subsection "Derivative" hoelzl@37489: hoelzl@37489: lemma has_derivative_vmul_component_cart: fixes c::"real^'a \ real^'b" and v::"real^'c" hoelzl@37489: assumes "(c has_derivative c') net" hoelzl@37489: shows "((\x. c(x)$k *\<^sub>R v) has_derivative (\x. (c' x)$k *\<^sub>R v)) net" hoelzl@37489: using has_derivative_vmul_component[OF assms] hoelzl@37489: unfolding nth_conv_component . hoelzl@37489: hoelzl@37489: lemma differentiable_at_imp_differentiable_on: "(\x\(s::(real^'n) set). f differentiable at x) \ f differentiable_on s" hoelzl@37489: unfolding differentiable_on_def by(auto intro!: differentiable_at_withinI) hoelzl@37489: hoelzl@37489: definition "jacobian f net = matrix(frechet_derivative f net)" hoelzl@37489: hoelzl@37489: lemma jacobian_works: "(f::(real^'a) \ (real^'b)) differentiable net \ (f has_derivative (\h. (jacobian f net) *v h)) net" hoelzl@37489: apply rule unfolding jacobian_def apply(simp only: matrix_works[OF linear_frechet_derivative]) defer hoelzl@37489: apply(rule differentiableI) apply assumption unfolding frechet_derivative_works by assumption hoelzl@37489: hoelzl@37489: subsection {* Component of the differential must be zero if it exists at a local *) hoelzl@37489: (* maximum or minimum for that corresponding component. *} hoelzl@37489: hoelzl@37489: lemma differential_zero_maxmin_component: fixes f::"real^'a \ real^'b" hoelzl@37489: assumes "0 < e" "((\y \ ball x e. (f y)$k \ (f x)$k) \ (\y\ball x e. (f x)$k \ (f y)$k))" hoelzl@37489: "f differentiable (at x)" shows "jacobian f (at x) $ k = 0" hoelzl@37489: (* FIXME: reuse proof of generic differential_zero_maxmin_component*) hoelzl@37489: hoelzl@37489: proof(rule ccontr) hoelzl@37489: def D \ "jacobian f (at x)" assume "jacobian f (at x) $ k \ 0" hoelzl@37489: then obtain j where j:"D$k$j \ 0" unfolding Cart_eq D_def by auto hoelzl@37489: hence *:"abs (jacobian f (at x) $ k $ j) / 2 > 0" unfolding D_def by auto hoelzl@37489: note as = assms(3)[unfolded jacobian_works has_derivative_at_alt] hoelzl@37489: guess e' using as[THEN conjunct2,rule_format,OF *] .. note e' = this hoelzl@37489: guess d using real_lbound_gt_zero[OF assms(1) e'[THEN conjunct1]] .. note d = this hoelzl@37489: { fix c assume "abs c \ d" hoelzl@37489: hence *:"norm (x + c *\<^sub>R cart_basis j - x) < e'" using norm_basis[of j] d by auto hoelzl@37489: 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))" hoelzl@37489: by(rule component_le_norm_cart) hoelzl@37489: 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 hoelzl@37489: 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 hoelzl@37489: hence "\f (x + c *\<^sub>R cart_basis j) $ k - f x $ k - c * D $ k $ j\ \ \D $ k $ j\ / 2 * \c\" hoelzl@37489: unfolding vector_component_simps matrix_vector_mul_component unfolding smult_conv_scaleR[symmetric] hoelzl@37489: unfolding inner_simps dot_basis smult_conv_scaleR by simp } note * = this hoelzl@37489: have "x + d *\<^sub>R cart_basis j \ ball x e" "x - d *\<^sub>R cart_basis j \ ball x e" hoelzl@37489: unfolding mem_ball dist_norm using norm_basis[of j] d by auto hoelzl@37489: 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) \ hoelzl@37489: ((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 hoelzl@37489: 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 hoelzl@37489: show False apply(rule ***[OF **, where dx="d * D $ k $ j" and d="\D $ k $ j\ / 2 * \d\"]) hoelzl@37489: using *[of "-d"] and *[of d] and d[THEN conjunct1] and j unfolding mult_minus_left hoelzl@37489: unfolding abs_mult diff_minus_eq_add scaleR.minus_left unfolding algebra_simps by (auto intro: mult_pos_pos) hoelzl@37489: qed hoelzl@37489: hoelzl@37494: subsection {* Lemmas for working on @{typ "real^1"} *} hoelzl@37489: hoelzl@37489: lemma forall_1[simp]: "(\i::1. P i) \ P 1" hoelzl@37489: by (metis num1_eq_iff) hoelzl@37489: hoelzl@37489: lemma ex_1[simp]: "(\x::1. P x) \ P 1" hoelzl@37489: by auto (metis num1_eq_iff) hoelzl@37489: hoelzl@37489: lemma exhaust_2: hoelzl@37489: fixes x :: 2 shows "x = 1 \ x = 2" hoelzl@37489: proof (induct x) hoelzl@37489: case (of_int z) hoelzl@37489: then have "0 <= z" and "z < 2" by simp_all hoelzl@37489: then have "z = 0 | z = 1" by arith hoelzl@37489: then show ?case by auto hoelzl@37489: qed hoelzl@37489: hoelzl@37489: lemma forall_2: "(\i::2. P i) \ P 1 \ P 2" hoelzl@37489: by (metis exhaust_2) hoelzl@37489: hoelzl@37489: lemma exhaust_3: hoelzl@37489: fixes x :: 3 shows "x = 1 \ x = 2 \ x = 3" hoelzl@37489: proof (induct x) hoelzl@37489: case (of_int z) hoelzl@37489: then have "0 <= z" and "z < 3" by simp_all hoelzl@37489: then have "z = 0 \ z = 1 \ z = 2" by arith hoelzl@37489: then show ?case by auto hoelzl@37489: qed hoelzl@37489: hoelzl@37489: lemma forall_3: "(\i::3. P i) \ P 1 \ P 2 \ P 3" hoelzl@37489: by (metis exhaust_3) hoelzl@37489: hoelzl@37489: lemma UNIV_1 [simp]: "UNIV = {1::1}" hoelzl@37489: by (auto simp add: num1_eq_iff) hoelzl@37489: hoelzl@37489: lemma UNIV_2: "UNIV = {1::2, 2::2}" hoelzl@37489: using exhaust_2 by auto hoelzl@37489: hoelzl@37489: lemma UNIV_3: "UNIV = {1::3, 2::3, 3::3}" hoelzl@37489: using exhaust_3 by auto hoelzl@37489: hoelzl@37489: lemma setsum_1: "setsum f (UNIV::1 set) = f 1" hoelzl@37489: unfolding UNIV_1 by simp hoelzl@37489: hoelzl@37489: lemma setsum_2: "setsum f (UNIV::2 set) = f 1 + f 2" hoelzl@37489: unfolding UNIV_2 by simp hoelzl@37489: hoelzl@37489: lemma setsum_3: "setsum f (UNIV::3 set) = f 1 + f 2 + f 3" hoelzl@37489: unfolding UNIV_3 by (simp add: add_ac) hoelzl@37489: hoelzl@37489: instantiation num1 :: cart_one begin hoelzl@37489: instance proof hoelzl@37489: show "CARD(1) = Suc 0" by auto hoelzl@37489: qed end hoelzl@37489: hoelzl@37489: (* "lift" from 'a to 'a^1 and "drop" from 'a^1 to 'a -- FIXME: potential use of transfer *) hoelzl@37489: hoelzl@37489: abbreviation vec1:: "'a \ 'a ^ 1" where "vec1 x \ vec x" hoelzl@37489: hoelzl@37489: abbreviation dest_vec1:: "'a ^1 \ 'a" hoelzl@37489: where "dest_vec1 x \ (x$1)" hoelzl@37489: hoelzl@37489: lemma vec1_dest_vec1[simp]: "vec1(dest_vec1 x) = x" "dest_vec1(vec1 y) = y" hoelzl@37489: by (simp_all add: Cart_eq) hoelzl@37489: hoelzl@37489: lemma vec1_component[simp]: "(vec1 x)$1 = x" hoelzl@37489: by (simp_all add: Cart_eq) hoelzl@37489: hoelzl@37489: declare vec1_dest_vec1(1) [simp] hoelzl@37489: hoelzl@37489: lemma forall_vec1: "(\x. P x) \ (\x. P (vec1 x))" hoelzl@37489: by (metis vec1_dest_vec1(1)) hoelzl@37489: hoelzl@37489: lemma exists_vec1: "(\x. P x) \ (\x. P(vec1 x))" hoelzl@37489: by (metis vec1_dest_vec1(1)) hoelzl@37489: hoelzl@37489: lemma vec1_eq[simp]: "vec1 x = vec1 y \ x = y" hoelzl@37489: by (metis vec1_dest_vec1(2)) hoelzl@37489: hoelzl@37489: lemma dest_vec1_eq[simp]: "dest_vec1 x = dest_vec1 y \ x = y" hoelzl@37489: by (metis vec1_dest_vec1(1)) hoelzl@37489: hoelzl@37489: subsection{* The collapse of the general concepts to dimension one. *} hoelzl@37489: hoelzl@37489: lemma vector_one: "(x::'a ^1) = (\ i. (x$1))" hoelzl@37489: by (simp add: Cart_eq) hoelzl@37489: hoelzl@37489: lemma forall_one: "(\(x::'a ^1). P x) \ (\x. P(\ i. x))" hoelzl@37489: apply auto hoelzl@37489: apply (erule_tac x= "x$1" in allE) hoelzl@37489: apply (simp only: vector_one[symmetric]) hoelzl@37489: done hoelzl@37489: hoelzl@37489: lemma norm_vector_1: "norm (x :: _^1) = norm (x$1)" hoelzl@37489: by (simp add: norm_vector_def) hoelzl@37489: hoelzl@37489: lemma norm_real: "norm(x::real ^ 1) = abs(x$1)" hoelzl@37489: by (simp add: norm_vector_1) hoelzl@37489: hoelzl@37489: lemma dist_real: "dist(x::real ^ 1) y = abs((x$1) - (y$1))" hoelzl@37489: by (auto simp add: norm_real dist_norm) hoelzl@37489: hoelzl@37489: subsection{* Explicit vector construction from lists. *} hoelzl@37489: hoelzl@37489: primrec from_nat :: "nat \ 'a::{monoid_add,one}" hoelzl@37489: where "from_nat 0 = 0" | "from_nat (Suc n) = 1 + from_nat n" hoelzl@37489: hoelzl@37489: lemma from_nat [simp]: "from_nat = of_nat" hoelzl@37489: by (rule ext, induct_tac x, simp_all) hoelzl@37489: hoelzl@37489: primrec hoelzl@37489: list_fun :: "nat \ _ list \ _ \ _" hoelzl@37489: where hoelzl@37489: "list_fun n [] = (\x. 0)" hoelzl@37489: | "list_fun n (x # xs) = fun_upd (list_fun (Suc n) xs) (from_nat n) x" hoelzl@37489: hoelzl@37489: definition "vector l = (\ i. list_fun 1 l i)" hoelzl@37489: (*definition "vector l = (\ i. if i <= length l then l ! (i - 1) else 0)"*) hoelzl@37489: hoelzl@37489: lemma vector_1: "(vector[x]) $1 = x" hoelzl@37489: unfolding vector_def by simp hoelzl@37489: hoelzl@37489: lemma vector_2: hoelzl@37489: "(vector[x,y]) $1 = x" hoelzl@37489: "(vector[x,y] :: 'a^2)$2 = (y::'a::zero)" hoelzl@37489: unfolding vector_def by simp_all hoelzl@37489: hoelzl@37489: lemma vector_3: hoelzl@37489: "(vector [x,y,z] ::('a::zero)^3)$1 = x" hoelzl@37489: "(vector [x,y,z] ::('a::zero)^3)$2 = y" hoelzl@37489: "(vector [x,y,z] ::('a::zero)^3)$3 = z" hoelzl@37489: unfolding vector_def by simp_all hoelzl@37489: hoelzl@37489: lemma forall_vector_1: "(\v::'a::zero^1. P v) \ (\x. P(vector[x]))" hoelzl@37489: apply auto hoelzl@37489: apply (erule_tac x="v$1" in allE) hoelzl@37489: apply (subgoal_tac "vector [v$1] = v") hoelzl@37489: apply simp hoelzl@37489: apply (vector vector_def) hoelzl@37489: apply simp hoelzl@37489: done hoelzl@37489: hoelzl@37489: lemma forall_vector_2: "(\v::'a::zero^2. P v) \ (\x y. P(vector[x, y]))" hoelzl@37489: apply auto hoelzl@37489: apply (erule_tac x="v$1" in allE) hoelzl@37489: apply (erule_tac x="v$2" in allE) hoelzl@37489: apply (subgoal_tac "vector [v$1, v$2] = v") hoelzl@37489: apply simp hoelzl@37489: apply (vector vector_def) hoelzl@37489: apply (simp add: forall_2) hoelzl@37489: done hoelzl@37489: hoelzl@37489: lemma forall_vector_3: "(\v::'a::zero^3. P v) \ (\x y z. P(vector[x, y, z]))" hoelzl@37489: apply auto hoelzl@37489: apply (erule_tac x="v$1" in allE) hoelzl@37489: apply (erule_tac x="v$2" in allE) hoelzl@37489: apply (erule_tac x="v$3" in allE) hoelzl@37489: apply (subgoal_tac "vector [v$1, v$2, v$3] = v") hoelzl@37489: apply simp hoelzl@37489: apply (vector vector_def) hoelzl@37489: apply (simp add: forall_3) hoelzl@37489: done hoelzl@37489: nipkow@39535: lemma range_vec1[simp]:"range vec1 = UNIV" apply(rule set_eqI,rule) unfolding image_iff defer hoelzl@37489: apply(rule_tac x="dest_vec1 x" in bexI) by auto hoelzl@37489: hoelzl@37489: lemma dest_vec1_lambda: "dest_vec1(\ i. x i) = x 1" hoelzl@37489: by (simp) hoelzl@37489: hoelzl@37489: lemma dest_vec1_vec: "dest_vec1(vec x) = x" hoelzl@37489: by (simp) hoelzl@37489: hoelzl@37489: lemma dest_vec1_sum: assumes fS: "finite S" hoelzl@37489: shows "dest_vec1(setsum f S) = setsum (dest_vec1 o f) S" hoelzl@37489: apply (induct rule: finite_induct[OF fS]) hoelzl@37489: apply simp hoelzl@37489: apply auto hoelzl@37489: done hoelzl@37489: hoelzl@37489: lemma norm_vec1 [simp]: "norm(vec1 x) = abs(x)" hoelzl@37489: by (simp add: vec_def norm_real) hoelzl@37489: hoelzl@37489: lemma dist_vec1: "dist(vec1 x) (vec1 y) = abs(x - y)" hoelzl@37489: by (simp only: dist_real vec1_component) hoelzl@37489: lemma abs_dest_vec1: "norm x = \dest_vec1 x\" hoelzl@37489: by (metis vec1_dest_vec1(1) norm_vec1) hoelzl@37489: hoelzl@37489: 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 hoelzl@37489: vec1_eq vec_cmul[THEN sym] smult_conv_scaleR[THEN sym] o_def dist_real_def norm_vec1 real_norm_def hoelzl@37489: hoelzl@37489: lemma bounded_linear_vec1:"bounded_linear (vec1::real\real^1)" hoelzl@37489: unfolding bounded_linear_def additive_def bounded_linear_axioms_def hoelzl@37489: unfolding smult_conv_scaleR[THEN sym] unfolding vec1_dest_vec1_simps hoelzl@37489: apply(rule conjI) defer apply(rule conjI) defer apply(rule_tac x=1 in exI) by auto hoelzl@37489: hoelzl@37489: lemma linear_vmul_dest_vec1: hoelzl@37489: fixes f:: "real^_ \ real^1" hoelzl@37489: shows "linear f \ linear (\x. dest_vec1(f x) *s v)" hoelzl@37489: unfolding smult_conv_scaleR hoelzl@37489: by (rule linear_vmul_component) hoelzl@37489: hoelzl@37489: lemma linear_from_scalars: hoelzl@37489: assumes lf: "linear (f::real^1 \ real^_)" hoelzl@37489: shows "f = (\x. dest_vec1 x *s column 1 (matrix f))" hoelzl@37489: unfolding smult_conv_scaleR hoelzl@37489: apply (rule ext) hoelzl@37489: apply (subst matrix_works[OF lf, symmetric]) hoelzl@37489: apply (auto simp add: Cart_eq matrix_vector_mult_def column_def mult_commute) hoelzl@37489: done hoelzl@37489: hoelzl@37489: lemma linear_to_scalars: assumes lf: "linear (f::real ^'n \ real^1)" hoelzl@37489: shows "f = (\x. vec1(row 1 (matrix f) \ x))" hoelzl@37489: apply (rule ext) hoelzl@37489: apply (subst matrix_works[OF lf, symmetric]) hoelzl@37489: apply (simp add: Cart_eq matrix_vector_mult_def row_def inner_vector_def mult_commute) hoelzl@37489: done hoelzl@37489: hoelzl@37489: lemma dest_vec1_eq_0: "dest_vec1 x = 0 \ x = 0" hoelzl@37489: by (simp add: dest_vec1_eq[symmetric]) hoelzl@37489: hoelzl@37489: lemma setsum_scalars: assumes fS: "finite S" hoelzl@37489: shows "setsum f S = vec1 (setsum (dest_vec1 o f) S)" hoelzl@37489: unfolding vec_setsum[OF fS] by simp hoelzl@37489: hoelzl@37489: 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" hoelzl@37489: apply (cases "dest_vec1 x \ dest_vec1 y") hoelzl@37489: apply simp hoelzl@37489: apply (subgoal_tac "dest_vec1 y \ dest_vec1 x") hoelzl@37489: apply (auto) hoelzl@37489: done hoelzl@37489: hoelzl@37489: text{* Lifting and dropping *} hoelzl@37489: hoelzl@37489: lemma continuous_on_o_dest_vec1: fixes f::"real \ 'a::real_normed_vector" hoelzl@37489: assumes "continuous_on {a..b::real} f" shows "continuous_on {vec1 a..vec1 b} (f o dest_vec1)" hoelzl@37489: using assms unfolding continuous_on_iff apply safe hoelzl@37489: apply(erule_tac x="x$1" in ballE,erule_tac x=e in allE) apply safe hoelzl@37489: apply(rule_tac x=d in exI) apply safe unfolding o_def dist_real_def dist_real hoelzl@37489: apply(erule_tac x="dest_vec1 x'" in ballE) by(auto simp add:vector_le_def) hoelzl@37489: hoelzl@37489: lemma continuous_on_o_vec1: fixes f::"real^1 \ 'a::real_normed_vector" hoelzl@37489: assumes "continuous_on {a..b} f" shows "continuous_on {dest_vec1 a..dest_vec1 b} (f o vec1)" hoelzl@37489: using assms unfolding continuous_on_iff apply safe hoelzl@37489: apply(erule_tac x="vec x" in ballE,erule_tac x=e in allE) apply safe hoelzl@37489: apply(rule_tac x=d in exI) apply safe unfolding o_def dist_real_def dist_real hoelzl@37489: apply(erule_tac x="vec1 x'" in ballE) by(auto simp add:vector_le_def) hoelzl@37489: hoelzl@37489: lemma continuous_on_vec1:"continuous_on A (vec1::real\real^1)" hoelzl@37489: by(rule linear_continuous_on[OF bounded_linear_vec1]) hoelzl@37489: hoelzl@37489: lemma mem_interval_1: fixes x :: "real^1" shows hoelzl@37489: "(x \ {a .. b} \ dest_vec1 a \ dest_vec1 x \ dest_vec1 x \ dest_vec1 b)" hoelzl@37489: "(x \ {a<.. dest_vec1 a < dest_vec1 x \ dest_vec1 x < dest_vec1 b)" hoelzl@37489: by(simp_all add: Cart_eq vector_less_def vector_le_def) hoelzl@37489: hoelzl@37489: lemma vec1_interval:fixes a::"real" shows hoelzl@37489: "vec1 ` {a .. b} = {vec1 a .. vec1 b}" hoelzl@37489: "vec1 ` {a<.. {a .. b} ==> x \ {a<.. (x = a) \ (x = b)" hoelzl@37489: unfolding Cart_eq vector_less_def vector_le_def mem_interval_cart by(auto simp del:dest_vec1_eq) hoelzl@37489: hoelzl@37489: lemma in_interval_1: fixes x :: "real^1" shows hoelzl@37489: "(x \ {a .. b} \ dest_vec1 a \ dest_vec1 x \ dest_vec1 x \ dest_vec1 b) \ hoelzl@37489: (x \ {a<.. dest_vec1 a < dest_vec1 x \ dest_vec1 x < dest_vec1 b)" hoelzl@37489: unfolding Cart_eq vector_less_def vector_le_def mem_interval_cart by(auto simp del:dest_vec1_eq) hoelzl@37489: hoelzl@37489: lemma interval_eq_empty_1: fixes a :: "real^1" shows hoelzl@37489: "{a .. b} = {} \ dest_vec1 b < dest_vec1 a" hoelzl@37489: "{a<.. dest_vec1 b \ dest_vec1 a" hoelzl@37489: unfolding interval_eq_empty_cart and ex_1 by auto hoelzl@37489: hoelzl@37489: lemma subset_interval_1: fixes a :: "real^1" shows hoelzl@37489: "({a .. b} \ {c .. d} \ dest_vec1 b < dest_vec1 a \ hoelzl@37489: dest_vec1 c \ dest_vec1 a \ dest_vec1 a \ dest_vec1 b \ dest_vec1 b \ dest_vec1 d)" hoelzl@37489: "({a .. b} \ {c<.. dest_vec1 b < dest_vec1 a \ hoelzl@37489: dest_vec1 c < dest_vec1 a \ dest_vec1 a \ dest_vec1 b \ dest_vec1 b < dest_vec1 d)" hoelzl@37489: "({a<.. {c .. d} \ dest_vec1 b \ dest_vec1 a \ hoelzl@37489: dest_vec1 c \ dest_vec1 a \ dest_vec1 a < dest_vec1 b \ dest_vec1 b \ dest_vec1 d)" hoelzl@37489: "({a<.. {c<.. dest_vec1 b \ dest_vec1 a \ hoelzl@37489: dest_vec1 c \ dest_vec1 a \ dest_vec1 a < dest_vec1 b \ dest_vec1 b \ dest_vec1 d)" hoelzl@37489: unfolding subset_interval_cart[of a b c d] unfolding forall_1 by auto hoelzl@37489: hoelzl@37489: lemma eq_interval_1: fixes a :: "real^1" shows hoelzl@37489: "{a .. b} = {c .. d} \ hoelzl@37489: dest_vec1 b < dest_vec1 a \ dest_vec1 d < dest_vec1 c \ hoelzl@37489: dest_vec1 a = dest_vec1 c \ dest_vec1 b = dest_vec1 d" hoelzl@37489: unfolding set_eq_subset[of "{a .. b}" "{c .. d}"] hoelzl@37489: unfolding subset_interval_1(1)[of a b c d] hoelzl@37489: unfolding subset_interval_1(1)[of c d a b] hoelzl@37489: by auto hoelzl@37489: hoelzl@37489: lemma disjoint_interval_1: fixes a :: "real^1" shows hoelzl@37489: "{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" hoelzl@37489: "{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" hoelzl@37489: "{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" hoelzl@37489: "{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" hoelzl@37489: unfolding disjoint_interval_cart and ex_1 by auto hoelzl@37489: hoelzl@37489: lemma open_closed_interval_1: fixes a :: "real^1" shows hoelzl@37489: "{a<.. dest_vec1 b ==> {a .. b} = {a<.. {a,b}" nipkow@39535: unfolding set_eq_iff apply simp unfolding vector_less_def and vector_le_def and forall_1 and dest_vec1_eq[THEN sym] by(auto simp del:dest_vec1_eq) hoelzl@37489: hoelzl@37489: lemma Lim_drop_le: fixes f :: "'a \ real^1" shows hoelzl@37489: "(f ---> l) net \ ~(trivial_limit net) \ eventually (\x. dest_vec1 (f x) \ b) net ==> dest_vec1 l \ b" hoelzl@37489: using Lim_component_le_cart[of f l net 1 b] by auto hoelzl@37489: hoelzl@37489: lemma Lim_drop_ge: fixes f :: "'a \ real^1" shows hoelzl@37489: "(f ---> l) net \ ~(trivial_limit net) \ eventually (\x. b \ dest_vec1 (f x)) net ==> b \ dest_vec1 l" hoelzl@37489: using Lim_component_ge_cart[of f l net b 1] by auto hoelzl@37489: hoelzl@37489: text{* Also more convenient formulations of monotone convergence. *} hoelzl@37489: hoelzl@37489: lemma bounded_increasing_convergent: fixes s::"nat \ real^1" hoelzl@37489: assumes "bounded {s n| n::nat. True}" "\n. dest_vec1(s n) \ dest_vec1(s(Suc n))" hoelzl@37489: shows "\l. (s ---> l) sequentially" hoelzl@37489: proof- hoelzl@37489: obtain a where a:"\n. \dest_vec1 (s n)\ \ a" using assms(1)[unfolded bounded_iff abs_dest_vec1] by auto hoelzl@37489: { fix m::nat hoelzl@37489: have "\ n. n\m \ dest_vec1 (s m) \ dest_vec1 (s n)" hoelzl@37489: apply(induct_tac n) apply simp using assms(2) apply(erule_tac x="na" in allE) by(auto simp add: not_less_eq_eq) } hoelzl@37489: hence "\m n. m \ n \ dest_vec1 (s m) \ dest_vec1 (s n)" by auto hoelzl@37489: 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 hoelzl@37489: thus ?thesis unfolding Lim_sequentially apply(rule_tac x="vec1 l" in exI) hoelzl@37489: unfolding dist_norm unfolding abs_dest_vec1 by auto hoelzl@37489: qed hoelzl@37489: hoelzl@37489: lemma dest_vec1_simps[simp]: fixes a::"real^1" hoelzl@37489: shows "a$1 = 0 \ a = 0" (*"a \ 1 \ dest_vec1 a \ 1" "0 \ a \ 0 \ dest_vec1 a"*) hoelzl@37489: "a \ b \ dest_vec1 a \ dest_vec1 b" "dest_vec1 (1::real^1) = 1" hoelzl@37489: by(auto simp add: vector_le_def Cart_eq) hoelzl@37489: hoelzl@37489: lemma dest_vec1_inverval: hoelzl@37489: "dest_vec1 ` {a .. b} = {dest_vec1 a .. dest_vec1 b}" hoelzl@37489: "dest_vec1 ` {a<.. b} = {dest_vec1 a<.. dest_vec1 b}" hoelzl@37489: "dest_vec1 ` {a ..x. dest_vec1 (f x)) S" hoelzl@37489: using dest_vec1_sum[OF assms] by auto hoelzl@37489: hoelzl@37489: lemma open_dest_vec1_vimage: "open S \ open (dest_vec1 -` S)" hoelzl@37489: unfolding open_vector_def forall_1 by auto hoelzl@37489: hoelzl@37489: lemma tendsto_dest_vec1 [tendsto_intros]: hoelzl@37489: "(f ---> l) net \ ((\x. dest_vec1 (f x)) ---> dest_vec1 l) net" hoelzl@37489: by(rule tendsto_Cart_nth) hoelzl@37489: hoelzl@37489: lemma continuous_dest_vec1: "continuous net f \ continuous net (\x. dest_vec1 (f x))" hoelzl@37489: unfolding continuous_def by (rule tendsto_dest_vec1) hoelzl@37489: hoelzl@37489: lemma forall_dest_vec1: "(\x. P x) \ (\x. P(dest_vec1 x))" hoelzl@37489: apply safe defer apply(erule_tac x="vec1 x" in allE) by auto hoelzl@37489: hoelzl@37489: lemma forall_of_dest_vec1: "(\v. P (\x. dest_vec1 (v x))) \ (\x. P x)" hoelzl@37489: apply rule apply rule apply(erule_tac x="(vec1 \ x)" in allE) unfolding o_def vec1_dest_vec1 by auto hoelzl@37489: hoelzl@37489: lemma forall_of_dest_vec1': "(\v. P (dest_vec1 v)) \ (\x. P x)" hoelzl@37489: apply rule apply rule apply(erule_tac x="(vec1 x)" in allE) defer apply rule hoelzl@37489: apply(erule_tac x="dest_vec1 v" in allE) unfolding o_def vec1_dest_vec1 by auto hoelzl@37489: hoelzl@37489: lemma dist_vec1_0[simp]: "dist(vec1 (x::real)) 0 = norm x" unfolding dist_norm by auto hoelzl@37489: hoelzl@37489: lemma bounded_linear_vec1_dest_vec1: fixes f::"real \ real" hoelzl@37489: shows "linear (vec1 \ f \ dest_vec1) = bounded_linear f" (is "?l = ?r") proof- hoelzl@37489: { assume ?l guess K using linear_bounded[OF `?l`] .. hoelzl@37489: hence "\K. \x. \f x\ \ \x\ * K" apply(rule_tac x=K in exI) hoelzl@37489: unfolding vec1_dest_vec1_simps by (auto simp add:field_simps) } hoelzl@37489: thus ?thesis unfolding linear_def bounded_linear_def additive_def bounded_linear_axioms_def o_def hoelzl@37489: unfolding vec1_dest_vec1_simps by auto qed hoelzl@37489: hoelzl@37489: lemma vec1_le[simp]:fixes a::real shows "vec1 a \ vec1 b \ a \ b" hoelzl@37489: unfolding vector_le_def by auto hoelzl@37489: lemma vec1_less[simp]:fixes a::real shows "vec1 a < vec1 b \ a < b" hoelzl@37489: unfolding vector_less_def by auto hoelzl@37489: hoelzl@37489: hoelzl@37489: subsection {* Derivatives on real = Derivatives on @{typ "real^1"} *} hoelzl@37489: hoelzl@37489: lemma has_derivative_within_vec1_dest_vec1: fixes f::"real\real" shows hoelzl@37489: "((vec1 \ f \ dest_vec1) has_derivative (vec1 \ f' \ dest_vec1)) (at (vec1 x) within vec1 ` s) hoelzl@37489: = (f has_derivative f') (at x within s)" hoelzl@37489: unfolding has_derivative_within unfolding bounded_linear_vec1_dest_vec1[unfolded linear_conv_bounded_linear] hoelzl@37489: unfolding o_def Lim_within Ball_def unfolding forall_vec1 hoelzl@37489: unfolding vec1_dest_vec1_simps dist_vec1_0 image_iff by auto hoelzl@37489: hoelzl@37489: lemma has_derivative_at_vec1_dest_vec1: fixes f::"real\real" shows hoelzl@37489: "((vec1 \ f \ dest_vec1) has_derivative (vec1 \ f' \ dest_vec1)) (at (vec1 x)) = (f has_derivative f') (at x)" hoelzl@37489: using has_derivative_within_vec1_dest_vec1[where s=UNIV, unfolded range_vec1 within_UNIV] by auto hoelzl@37489: hoelzl@37489: lemma bounded_linear_vec1': fixes f::"'a::real_normed_vector\real" hoelzl@37489: shows "bounded_linear f = bounded_linear (vec1 \ f)" hoelzl@37489: unfolding bounded_linear_def additive_def bounded_linear_axioms_def o_def hoelzl@37489: unfolding vec1_dest_vec1_simps by auto hoelzl@37489: hoelzl@37489: lemma bounded_linear_dest_vec1: fixes f::"real\'a::real_normed_vector" hoelzl@37489: shows "bounded_linear f = bounded_linear (f \ dest_vec1)" hoelzl@37489: unfolding bounded_linear_def additive_def bounded_linear_axioms_def o_def hoelzl@37489: unfolding vec1_dest_vec1_simps by auto hoelzl@37489: hoelzl@37489: lemma has_derivative_at_vec1: fixes f::"'a::real_normed_vector\real" shows hoelzl@37489: "(f has_derivative f') (at x) = ((vec1 \ f) has_derivative (vec1 \ f')) (at x)" hoelzl@37489: unfolding has_derivative_at unfolding bounded_linear_vec1'[unfolded linear_conv_bounded_linear] hoelzl@37489: unfolding o_def Lim_at unfolding vec1_dest_vec1_simps dist_vec1_0 by auto hoelzl@37489: hoelzl@37489: lemma has_derivative_within_dest_vec1:fixes f::"real\'a::real_normed_vector" shows hoelzl@37489: "((f \ dest_vec1) has_derivative (f' \ dest_vec1)) (at (vec1 x) within vec1 ` s) = (f has_derivative f') (at x within s)" hoelzl@37489: unfolding has_derivative_within bounded_linear_dest_vec1 unfolding o_def Lim_within Ball_def hoelzl@37489: unfolding vec1_dest_vec1_simps dist_vec1_0 image_iff by auto hoelzl@37489: hoelzl@37489: lemma has_derivative_at_dest_vec1:fixes f::"real\'a::real_normed_vector" shows hoelzl@37489: "((f \ dest_vec1) has_derivative (f' \ dest_vec1)) (at (vec1 x)) = (f has_derivative f') (at x)" hoelzl@37489: using has_derivative_within_dest_vec1[where s=UNIV] by(auto simp add:within_UNIV) hoelzl@37489: hoelzl@37489: subsection {* In particular if we have a mapping into @{typ "real^1"}. *} hoelzl@37489: hoelzl@37489: lemma onorm_vec1: fixes f::"real \ real" hoelzl@37489: shows "onorm (\x. vec1 (f (dest_vec1 x))) = onorm f" proof- hoelzl@37489: have "\x::real^1. norm x = 1 \ x\{vec1 -1, vec1 (1::real)}" unfolding forall_vec1 by(auto simp add:Cart_eq) hoelzl@37489: hence 1:"{x. norm x = 1} = {vec1 -1, vec1 (1::real)}" by auto hoelzl@37489: 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 hoelzl@37489: have "\x::real. norm x = 1 \ x\{-1, 1}" by auto hence 3:"{x. norm x = 1} = {-1, (1::real)}" by auto hoelzl@37489: have 4:"{norm (f x) |x. norm x = 1} = (\x. norm (f x)) ` {x. norm x=1}" by auto hoelzl@37489: show ?thesis unfolding onorm_def 1 2 3 4 by(simp add:Sup_finite_Max) qed hoelzl@37489: hoelzl@37489: lemma convex_vec1:"convex (vec1 ` s) = convex (s::real set)" hoelzl@37489: unfolding convex_def Ball_def forall_vec1 unfolding vec1_dest_vec1_simps image_iff by auto hoelzl@37489: hoelzl@37489: lemma bounded_linear_component_cart[intro]: "bounded_linear (\x::real^'n. x $ k)" hoelzl@37489: apply(rule bounded_linearI[where K=1]) hoelzl@37489: using component_le_norm_cart[of _ k] unfolding real_norm_def by auto hoelzl@37489: hoelzl@37489: lemma bounded_vec1[intro]: "bounded s \ bounded (vec1 ` (s::real set))" hoelzl@37489: unfolding bounded_def apply safe apply(rule_tac x="vec1 x" in exI,rule_tac x=e in exI) hoelzl@37489: by(auto simp add: dist_real dist_real_def) hoelzl@37489: hoelzl@37489: (*lemma content_closed_interval_cases_cart: hoelzl@37489: "content {a..b::real^'n} = hoelzl@37489: (if {a..b} = {} then 0 else setprod (\i. b$i - a$i) UNIV)" hoelzl@37489: proof(cases "{a..b} = {}") hoelzl@37489: case True thus ?thesis unfolding content_def by auto hoelzl@37489: next case Falsethus ?thesis unfolding content_def unfolding if_not_P[OF False] hoelzl@37489: proof(cases "\i. a $ i \ b $ i") hoelzl@37489: case False thus ?thesis unfolding content_def using t by auto hoelzl@37489: next case True note interval_eq_empty hoelzl@37489: apply auto hoelzl@37489: hoelzl@37489: sorry*) hoelzl@37489: hoelzl@37489: lemma integral_component_eq_cart[simp]: fixes f::"'n::ordered_euclidean_space \ real^'m" hoelzl@37489: assumes "f integrable_on s" shows "integral s (\x. f x $ k) = integral s f $ k" hoelzl@37489: using integral_linear[OF assms(1) bounded_linear_component_cart,unfolded o_def] . hoelzl@37489: hoelzl@37489: lemma interval_split_cart: hoelzl@37489: "{a..b::real^'n} \ {x. x$k \ c} = {a .. (\ i. if i = k then min (b$k) c else b$i)}" hoelzl@37489: "{a..b} \ {x. x$k \ c} = {(\ i. if i = k then max (a$k) c else a$i) .. b}" nipkow@39535: apply(rule_tac[!] set_eqI) unfolding Int_iff mem_interval_cart mem_Collect_eq hoelzl@37489: unfolding Cart_lambda_beta by auto hoelzl@37489: hoelzl@37489: (*lemma content_split_cart: hoelzl@37489: "content {a..b::real^'n} = content({a..b} \ {x. x$k \ c}) + content({a..b} \ {x. x$k >= c})" hoelzl@37489: proof- note simps = interval_split_cart content_closed_interval_cases_cart Cart_lambda_beta vector_le_def hoelzl@37489: { presume "a\b \ ?thesis" thus ?thesis apply(cases "a\b") unfolding simps by auto } hoelzl@37489: have *:"UNIV = insert k (UNIV - {k})" "\x. finite (UNIV-{x::'n})" "\x. x\UNIV-{x}" by auto hoelzl@37489: have *:"\X Y Z. (\i\UNIV. Z i (if i = k then X else Y i)) = Z k X * (\i\UNIV-{k}. Z i (Y i))" hoelzl@37489: "(\i\UNIV. b$i - a$i) = (\i\UNIV-{k}. b$i - a$i) * (b$k - a$k)" hoelzl@37489: apply(subst *(1)) defer apply(subst *(1)) unfolding setprod_insert[OF *(2-)] by auto hoelzl@37489: assume as:"a\b" moreover have "\x. min (b $ k) c = max (a $ k) c hoelzl@37489: \ x* (b$k - a$k) = x*(max (a $ k) c - a $ k) + x*(b $ k - max (a $ k) c)" hoelzl@37489: by (auto simp add:field_simps) hoelzl@37489: moreover have "\ a $ k \ c \ \ c \ b $ k \ False" hoelzl@37489: unfolding not_le using as[unfolded vector_le_def,rule_format,of k] by auto hoelzl@37489: ultimately show ?thesis hoelzl@37489: unfolding simps unfolding *(1)[of "\i x. b$i - x"] *(1)[of "\i x. x - a$i"] *(2) by(auto) hoelzl@37489: qed*) hoelzl@37489: hoelzl@37489: lemma has_integral_vec1: assumes "(f has_integral k) {a..b}" hoelzl@37489: shows "((\x. vec1 (f x)) has_integral (vec1 k)) {a..b}" hoelzl@37489: 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)" hoelzl@37489: unfolding vec_sub Cart_eq by(auto simp add: split_beta) hoelzl@37489: show ?thesis using assms unfolding has_integral apply safe hoelzl@37489: apply(erule_tac x=e in allE,safe) apply(rule_tac x=d in exI,safe) hoelzl@37489: apply(erule_tac x=p in allE,safe) unfolding * norm_vector_1 by auto qed hoelzl@37489: hoelzl@37489: end