huffman@30019: (* Title: HOL/Library/Product_Vector.thy huffman@30019: Author: Brian Huffman huffman@30019: *) huffman@30019: huffman@30019: header {* Cartesian Products as Vector Spaces *} huffman@30019: huffman@30019: theory Product_Vector huffman@30019: imports Inner_Product Product_plus huffman@30019: begin huffman@30019: huffman@30019: subsection {* Product is a real vector space *} huffman@30019: haftmann@37678: instantiation prod :: (real_vector, real_vector) real_vector huffman@30019: begin huffman@30019: huffman@30019: definition scaleR_prod_def: huffman@30019: "scaleR r A = (scaleR r (fst A), scaleR r (snd A))" huffman@30019: huffman@30019: lemma fst_scaleR [simp]: "fst (scaleR r A) = scaleR r (fst A)" huffman@30019: unfolding scaleR_prod_def by simp huffman@30019: huffman@30019: lemma snd_scaleR [simp]: "snd (scaleR r A) = scaleR r (snd A)" huffman@30019: unfolding scaleR_prod_def by simp huffman@30019: huffman@30019: lemma scaleR_Pair [simp]: "scaleR r (a, b) = (scaleR r a, scaleR r b)" huffman@30019: unfolding scaleR_prod_def by simp huffman@30019: huffman@30019: instance proof huffman@30019: fix a b :: real and x y :: "'a \ 'b" huffman@30019: show "scaleR a (x + y) = scaleR a x + scaleR a y" huffman@44066: by (simp add: prod_eq_iff scaleR_right_distrib) huffman@30019: show "scaleR (a + b) x = scaleR a x + scaleR b x" huffman@44066: by (simp add: prod_eq_iff scaleR_left_distrib) huffman@30019: show "scaleR a (scaleR b x) = scaleR (a * b) x" huffman@44066: by (simp add: prod_eq_iff) huffman@30019: show "scaleR 1 x = x" huffman@44066: by (simp add: prod_eq_iff) huffman@30019: qed huffman@30019: huffman@30019: end huffman@30019: huffman@31415: subsection {* Product is a topological space *} huffman@31415: haftmann@37678: instantiation prod :: (topological_space, topological_space) topological_space huffman@31415: begin huffman@31415: huffman@31492: definition open_prod_def: huffman@31492: "open (S :: ('a \ 'b) set) \ huffman@31492: (\x\S. \A B. open A \ open B \ x \ A \ B \ A \ B \ S)" huffman@31415: huffman@36332: lemma open_prod_elim: huffman@36332: assumes "open S" and "x \ S" huffman@36332: obtains A B where "open A" and "open B" and "x \ A \ B" and "A \ B \ S" huffman@36332: using assms unfolding open_prod_def by fast huffman@36332: huffman@36332: lemma open_prod_intro: huffman@36332: assumes "\x. x \ S \ \A B. open A \ open B \ x \ A \ B \ A \ B \ S" huffman@36332: shows "open S" huffman@36332: using assms unfolding open_prod_def by fast huffman@36332: huffman@31415: instance proof huffman@31492: show "open (UNIV :: ('a \ 'b) set)" huffman@31492: unfolding open_prod_def by auto huffman@31415: next huffman@31415: fix S T :: "('a \ 'b) set" huffman@36332: assume "open S" "open T" huffman@36332: show "open (S \ T)" huffman@36332: proof (rule open_prod_intro) huffman@36332: fix x assume x: "x \ S \ T" huffman@36332: from x have "x \ S" by simp huffman@36332: obtain Sa Sb where A: "open Sa" "open Sb" "x \ Sa \ Sb" "Sa \ Sb \ S" huffman@36332: using `open S` and `x \ S` by (rule open_prod_elim) huffman@36332: from x have "x \ T" by simp huffman@36332: obtain Ta Tb where B: "open Ta" "open Tb" "x \ Ta \ Tb" "Ta \ Tb \ T" huffman@36332: using `open T` and `x \ T` by (rule open_prod_elim) huffman@36332: let ?A = "Sa \ Ta" and ?B = "Sb \ Tb" huffman@36332: have "open ?A \ open ?B \ x \ ?A \ ?B \ ?A \ ?B \ S \ T" huffman@36332: using A B by (auto simp add: open_Int) huffman@36332: thus "\A B. open A \ open B \ x \ A \ B \ A \ B \ S \ T" huffman@36332: by fast huffman@36332: qed huffman@31415: next huffman@31492: fix K :: "('a \ 'b) set set" huffman@31492: assume "\S\K. open S" thus "open (\K)" huffman@31492: unfolding open_prod_def by fast huffman@31415: qed huffman@31415: huffman@31415: end huffman@31415: huffman@31562: lemma open_Times: "open S \ open T \ open (S \ T)" huffman@31562: unfolding open_prod_def by auto huffman@31562: huffman@31562: lemma fst_vimage_eq_Times: "fst -` S = S \ UNIV" huffman@31562: by auto huffman@31562: huffman@31562: lemma snd_vimage_eq_Times: "snd -` S = UNIV \ S" huffman@31562: by auto huffman@31562: huffman@31562: lemma open_vimage_fst: "open S \ open (fst -` S)" huffman@31562: by (simp add: fst_vimage_eq_Times open_Times) huffman@31562: huffman@31562: lemma open_vimage_snd: "open S \ open (snd -` S)" huffman@31562: by (simp add: snd_vimage_eq_Times open_Times) huffman@31562: huffman@31568: lemma closed_vimage_fst: "closed S \ closed (fst -` S)" huffman@31568: unfolding closed_open vimage_Compl [symmetric] huffman@31568: by (rule open_vimage_fst) huffman@31568: huffman@31568: lemma closed_vimage_snd: "closed S \ closed (snd -` S)" huffman@31568: unfolding closed_open vimage_Compl [symmetric] huffman@31568: by (rule open_vimage_snd) huffman@31568: huffman@31568: lemma closed_Times: "closed S \ closed T \ closed (S \ T)" huffman@31568: proof - huffman@31568: have "S \ T = (fst -` S) \ (snd -` T)" by auto huffman@31568: thus "closed S \ closed T \ closed (S \ T)" huffman@31568: by (simp add: closed_vimage_fst closed_vimage_snd closed_Int) huffman@31568: qed huffman@31568: huffman@34110: lemma openI: (* TODO: move *) huffman@34110: assumes "\x. x \ S \ \T. open T \ x \ T \ T \ S" huffman@34110: shows "open S" huffman@34110: proof - huffman@34110: have "open (\{T. open T \ T \ S})" by auto huffman@34110: moreover have "\{T. open T \ T \ S} = S" by (auto dest!: assms) huffman@34110: ultimately show "open S" by simp huffman@34110: qed huffman@34110: huffman@34110: lemma subset_fst_imageI: "A \ B \ S \ y \ B \ A \ fst ` S" huffman@34110: unfolding image_def subset_eq by force huffman@34110: huffman@34110: lemma subset_snd_imageI: "A \ B \ S \ x \ A \ B \ snd ` S" huffman@34110: unfolding image_def subset_eq by force huffman@34110: huffman@34110: lemma open_image_fst: assumes "open S" shows "open (fst ` S)" huffman@34110: proof (rule openI) huffman@34110: fix x assume "x \ fst ` S" huffman@34110: then obtain y where "(x, y) \ S" by auto huffman@34110: then obtain A B where "open A" "open B" "x \ A" "y \ B" "A \ B \ S" huffman@34110: using `open S` unfolding open_prod_def by auto huffman@34110: from `A \ B \ S` `y \ B` have "A \ fst ` S" by (rule subset_fst_imageI) huffman@34110: with `open A` `x \ A` have "open A \ x \ A \ A \ fst ` S" by simp huffman@34110: then show "\T. open T \ x \ T \ T \ fst ` S" by - (rule exI) huffman@34110: qed huffman@34110: huffman@34110: lemma open_image_snd: assumes "open S" shows "open (snd ` S)" huffman@34110: proof (rule openI) huffman@34110: fix y assume "y \ snd ` S" huffman@34110: then obtain x where "(x, y) \ S" by auto huffman@34110: then obtain A B where "open A" "open B" "x \ A" "y \ B" "A \ B \ S" huffman@34110: using `open S` unfolding open_prod_def by auto huffman@34110: from `A \ B \ S` `x \ A` have "B \ snd ` S" by (rule subset_snd_imageI) huffman@34110: with `open B` `y \ B` have "open B \ y \ B \ B \ snd ` S" by simp huffman@34110: then show "\T. open T \ y \ T \ T \ snd ` S" by - (rule exI) huffman@34110: qed huffman@31568: huffman@44575: subsubsection {* Continuity of operations *} huffman@44575: huffman@44575: lemma tendsto_fst [tendsto_intros]: huffman@44575: assumes "(f ---> a) F" huffman@44575: shows "((\x. fst (f x)) ---> fst a) F" huffman@44575: proof (rule topological_tendstoI) huffman@44575: fix S assume "open S" and "fst a \ S" huffman@44575: then have "open (fst -` S)" and "a \ fst -` S" huffman@44575: by (simp_all add: open_vimage_fst) huffman@44575: with assms have "eventually (\x. f x \ fst -` S) F" huffman@44575: by (rule topological_tendstoD) huffman@44575: then show "eventually (\x. fst (f x) \ S) F" huffman@44575: by simp huffman@44575: qed huffman@44575: huffman@44575: lemma tendsto_snd [tendsto_intros]: huffman@44575: assumes "(f ---> a) F" huffman@44575: shows "((\x. snd (f x)) ---> snd a) F" huffman@44575: proof (rule topological_tendstoI) huffman@44575: fix S assume "open S" and "snd a \ S" huffman@44575: then have "open (snd -` S)" and "a \ snd -` S" huffman@44575: by (simp_all add: open_vimage_snd) huffman@44575: with assms have "eventually (\x. f x \ snd -` S) F" huffman@44575: by (rule topological_tendstoD) huffman@44575: then show "eventually (\x. snd (f x) \ S) F" huffman@44575: by simp huffman@44575: qed huffman@44575: huffman@44575: lemma tendsto_Pair [tendsto_intros]: huffman@44575: assumes "(f ---> a) F" and "(g ---> b) F" huffman@44575: shows "((\x. (f x, g x)) ---> (a, b)) F" huffman@44575: proof (rule topological_tendstoI) huffman@44575: fix S assume "open S" and "(a, b) \ S" huffman@44575: then obtain A B where "open A" "open B" "a \ A" "b \ B" "A \ B \ S" huffman@44575: unfolding open_prod_def by fast huffman@44575: have "eventually (\x. f x \ A) F" huffman@44575: using `(f ---> a) F` `open A` `a \ A` huffman@44575: by (rule topological_tendstoD) huffman@44575: moreover huffman@44575: have "eventually (\x. g x \ B) F" huffman@44575: using `(g ---> b) F` `open B` `b \ B` huffman@44575: by (rule topological_tendstoD) huffman@44575: ultimately huffman@44575: show "eventually (\x. (f x, g x) \ S) F" huffman@44575: by (rule eventually_elim2) huffman@44575: (simp add: subsetD [OF `A \ B \ S`]) huffman@44575: qed huffman@44575: huffman@44575: lemma isCont_fst [simp]: "isCont f a \ isCont (\x. fst (f x)) a" huffman@44575: unfolding isCont_def by (rule tendsto_fst) huffman@44575: huffman@44575: lemma isCont_snd [simp]: "isCont f a \ isCont (\x. snd (f x)) a" huffman@44575: unfolding isCont_def by (rule tendsto_snd) huffman@44575: huffman@44575: lemma isCont_Pair [simp]: huffman@44575: "\isCont f a; isCont g a\ \ isCont (\x. (f x, g x)) a" huffman@44575: unfolding isCont_def by (rule tendsto_Pair) huffman@44575: huffman@44575: subsubsection {* Separation axioms *} huffman@44214: huffman@44214: lemma mem_Times_iff: "x \ A \ B \ fst x \ A \ snd x \ B" huffman@44214: by (induct x) simp (* TODO: move elsewhere *) huffman@44214: huffman@44214: instance prod :: (t0_space, t0_space) t0_space huffman@44214: proof huffman@44214: fix x y :: "'a \ 'b" assume "x \ y" huffman@44214: hence "fst x \ fst y \ snd x \ snd y" huffman@44214: by (simp add: prod_eq_iff) huffman@44214: thus "\U. open U \ (x \ U) \ (y \ U)" huffman@44214: apply (rule disjE) huffman@44214: apply (drule t0_space, erule exE, rule_tac x="U \ UNIV" in exI) huffman@44214: apply (simp add: open_Times mem_Times_iff) huffman@44214: apply (drule t0_space, erule exE, rule_tac x="UNIV \ U" in exI) huffman@44214: apply (simp add: open_Times mem_Times_iff) huffman@44214: done huffman@44214: qed huffman@44214: huffman@44214: instance prod :: (t1_space, t1_space) t1_space huffman@44214: proof huffman@44214: fix x y :: "'a \ 'b" assume "x \ y" huffman@44214: hence "fst x \ fst y \ snd x \ snd y" huffman@44214: by (simp add: prod_eq_iff) huffman@44214: thus "\U. open U \ x \ U \ y \ U" huffman@44214: apply (rule disjE) huffman@44214: apply (drule t1_space, erule exE, rule_tac x="U \ UNIV" in exI) huffman@44214: apply (simp add: open_Times mem_Times_iff) huffman@44214: apply (drule t1_space, erule exE, rule_tac x="UNIV \ U" in exI) huffman@44214: apply (simp add: open_Times mem_Times_iff) huffman@44214: done huffman@44214: qed huffman@44214: huffman@44214: instance prod :: (t2_space, t2_space) t2_space huffman@44214: proof huffman@44214: fix x y :: "'a \ 'b" assume "x \ y" huffman@44214: hence "fst x \ fst y \ snd x \ snd y" huffman@44214: by (simp add: prod_eq_iff) huffman@44214: thus "\U V. open U \ open V \ x \ U \ y \ V \ U \ V = {}" huffman@44214: apply (rule disjE) huffman@44214: apply (drule hausdorff, clarify) huffman@44214: apply (rule_tac x="U \ UNIV" in exI, rule_tac x="V \ UNIV" in exI) huffman@44214: apply (simp add: open_Times mem_Times_iff disjoint_iff_not_equal) huffman@44214: apply (drule hausdorff, clarify) huffman@44214: apply (rule_tac x="UNIV \ U" in exI, rule_tac x="UNIV \ V" in exI) huffman@44214: apply (simp add: open_Times mem_Times_iff disjoint_iff_not_equal) huffman@44214: done huffman@44214: qed huffman@44214: huffman@31339: subsection {* Product is a metric space *} huffman@31339: haftmann@37678: instantiation prod :: (metric_space, metric_space) metric_space huffman@31339: begin huffman@31339: huffman@31339: definition dist_prod_def: huffman@44214: "dist x y = sqrt ((dist (fst x) (fst y))\ + (dist (snd x) (snd y))\)" huffman@31339: huffman@31339: lemma dist_Pair_Pair: "dist (a, b) (c, d) = sqrt ((dist a c)\ + (dist b d)\)" huffman@31339: unfolding dist_prod_def by simp huffman@31339: huffman@36332: lemma dist_fst_le: "dist (fst x) (fst y) \ dist x y" huffman@36332: unfolding dist_prod_def by (rule real_sqrt_sum_squares_ge1) huffman@36332: huffman@36332: lemma dist_snd_le: "dist (snd x) (snd y) \ dist x y" huffman@36332: unfolding dist_prod_def by (rule real_sqrt_sum_squares_ge2) huffman@36332: huffman@31339: instance proof huffman@31339: fix x y :: "'a \ 'b" huffman@31339: show "dist x y = 0 \ x = y" huffman@44066: unfolding dist_prod_def prod_eq_iff by simp huffman@31339: next huffman@31339: fix x y z :: "'a \ 'b" huffman@31339: show "dist x y \ dist x z + dist y z" huffman@31339: unfolding dist_prod_def huffman@31563: by (intro order_trans [OF _ real_sqrt_sum_squares_triangle_ineq] huffman@31563: real_sqrt_le_mono add_mono power_mono dist_triangle2 zero_le_dist) huffman@31415: next huffman@31492: fix S :: "('a \ 'b) set" huffman@31492: show "open S \ (\x\S. \e>0. \y. dist y x < e \ y \ S)" huffman@31563: proof huffman@36332: assume "open S" show "\x\S. \e>0. \y. dist y x < e \ y \ S" huffman@36332: proof huffman@36332: fix x assume "x \ S" huffman@36332: obtain A B where "open A" "open B" "x \ A \ B" "A \ B \ S" huffman@36332: using `open S` and `x \ S` by (rule open_prod_elim) huffman@36332: obtain r where r: "0 < r" "\y. dist y (fst x) < r \ y \ A" huffman@36332: using `open A` and `x \ A \ B` unfolding open_dist by auto huffman@36332: obtain s where s: "0 < s" "\y. dist y (snd x) < s \ y \ B" huffman@36332: using `open B` and `x \ A \ B` unfolding open_dist by auto huffman@36332: let ?e = "min r s" huffman@36332: have "0 < ?e \ (\y. dist y x < ?e \ y \ S)" huffman@36332: proof (intro allI impI conjI) huffman@36332: show "0 < min r s" by (simp add: r(1) s(1)) huffman@36332: next huffman@36332: fix y assume "dist y x < min r s" huffman@36332: hence "dist y x < r" and "dist y x < s" huffman@36332: by simp_all huffman@36332: hence "dist (fst y) (fst x) < r" and "dist (snd y) (snd x) < s" huffman@36332: by (auto intro: le_less_trans dist_fst_le dist_snd_le) huffman@36332: hence "fst y \ A" and "snd y \ B" huffman@36332: by (simp_all add: r(2) s(2)) huffman@36332: hence "y \ A \ B" by (induct y, simp) huffman@36332: with `A \ B \ S` show "y \ S" .. huffman@36332: qed huffman@36332: thus "\e>0. \y. dist y x < e \ y \ S" .. huffman@36332: qed huffman@31563: next huffman@44575: assume *: "\x\S. \e>0. \y. dist y x < e \ y \ S" show "open S" huffman@44575: proof (rule open_prod_intro) huffman@44575: fix x assume "x \ S" huffman@44575: then obtain e where "0 < e" and S: "\y. dist y x < e \ y \ S" huffman@44575: using * by fast huffman@44575: def r \ "e / sqrt 2" and s \ "e / sqrt 2" huffman@44575: from `0 < e` have "0 < r" and "0 < s" huffman@44575: unfolding r_def s_def by (simp_all add: divide_pos_pos) huffman@44575: from `0 < e` have "e = sqrt (r\ + s\)" huffman@44575: unfolding r_def s_def by (simp add: power_divide) huffman@44575: def A \ "{y. dist (fst x) y < r}" and B \ "{y. dist (snd x) y < s}" huffman@44575: have "open A" and "open B" huffman@44575: unfolding A_def B_def by (simp_all add: open_ball) huffman@44575: moreover have "x \ A \ B" huffman@44575: unfolding A_def B_def mem_Times_iff huffman@44575: using `0 < r` and `0 < s` by simp huffman@44575: moreover have "A \ B \ S" huffman@44575: proof (clarify) huffman@44575: fix a b assume "a \ A" and "b \ B" huffman@44575: hence "dist a (fst x) < r" and "dist b (snd x) < s" huffman@44575: unfolding A_def B_def by (simp_all add: dist_commute) huffman@44575: hence "dist (a, b) x < e" huffman@44575: unfolding dist_prod_def `e = sqrt (r\ + s\)` huffman@44575: by (simp add: add_strict_mono power_strict_mono) huffman@44575: thus "(a, b) \ S" huffman@44575: by (simp add: S) huffman@44575: qed huffman@44575: ultimately show "\A B. open A \ open B \ x \ A \ B \ A \ B \ S" by fast huffman@44575: qed huffman@31563: qed huffman@31339: qed huffman@31339: huffman@31339: end huffman@31339: huffman@31405: lemma Cauchy_fst: "Cauchy X \ Cauchy (\n. fst (X n))" huffman@31405: unfolding Cauchy_def by (fast elim: le_less_trans [OF dist_fst_le]) huffman@31405: huffman@31405: lemma Cauchy_snd: "Cauchy X \ Cauchy (\n. snd (X n))" huffman@31405: unfolding Cauchy_def by (fast elim: le_less_trans [OF dist_snd_le]) huffman@31405: huffman@31405: lemma Cauchy_Pair: huffman@31405: assumes "Cauchy X" and "Cauchy Y" huffman@31405: shows "Cauchy (\n. (X n, Y n))" huffman@31405: proof (rule metric_CauchyI) huffman@31405: fix r :: real assume "0 < r" huffman@31405: then have "0 < r / sqrt 2" (is "0 < ?s") huffman@31405: by (simp add: divide_pos_pos) huffman@31405: obtain M where M: "\m\M. \n\M. dist (X m) (X n) < ?s" huffman@31405: using metric_CauchyD [OF `Cauchy X` `0 < ?s`] .. huffman@31405: obtain N where N: "\m\N. \n\N. dist (Y m) (Y n) < ?s" huffman@31405: using metric_CauchyD [OF `Cauchy Y` `0 < ?s`] .. huffman@31405: have "\m\max M N. \n\max M N. dist (X m, Y m) (X n, Y n) < r" huffman@31405: using M N by (simp add: real_sqrt_sum_squares_less dist_Pair_Pair) huffman@31405: then show "\n0. \m\n0. \n\n0. dist (X m, Y m) (X n, Y n) < r" .. huffman@31405: qed huffman@31405: huffman@31405: subsection {* Product is a complete metric space *} huffman@31405: haftmann@37678: instance prod :: (complete_space, complete_space) complete_space huffman@31405: proof huffman@31405: fix X :: "nat \ 'a \ 'b" assume "Cauchy X" huffman@31405: have 1: "(\n. fst (X n)) ----> lim (\n. fst (X n))" huffman@31405: using Cauchy_fst [OF `Cauchy X`] huffman@31405: by (simp add: Cauchy_convergent_iff convergent_LIMSEQ_iff) huffman@31405: have 2: "(\n. snd (X n)) ----> lim (\n. snd (X n))" huffman@31405: using Cauchy_snd [OF `Cauchy X`] huffman@31405: by (simp add: Cauchy_convergent_iff convergent_LIMSEQ_iff) huffman@31405: have "X ----> (lim (\n. fst (X n)), lim (\n. snd (X n)))" huffman@36660: using tendsto_Pair [OF 1 2] by simp huffman@31405: then show "convergent X" huffman@31405: by (rule convergentI) huffman@31405: qed huffman@31405: huffman@30019: subsection {* Product is a normed vector space *} huffman@30019: haftmann@37678: instantiation prod :: (real_normed_vector, real_normed_vector) real_normed_vector huffman@30019: begin huffman@30019: huffman@30019: definition norm_prod_def: huffman@30019: "norm x = sqrt ((norm (fst x))\ + (norm (snd x))\)" huffman@30019: huffman@30019: definition sgn_prod_def: huffman@30019: "sgn (x::'a \ 'b) = scaleR (inverse (norm x)) x" huffman@30019: huffman@30019: lemma norm_Pair: "norm (a, b) = sqrt ((norm a)\ + (norm b)\)" huffman@30019: unfolding norm_prod_def by simp huffman@30019: huffman@30019: instance proof huffman@30019: fix r :: real and x y :: "'a \ 'b" huffman@30019: show "norm x = 0 \ x = 0" huffman@30019: unfolding norm_prod_def huffman@44066: by (simp add: prod_eq_iff) huffman@30019: show "norm (x + y) \ norm x + norm y" huffman@30019: unfolding norm_prod_def huffman@30019: apply (rule order_trans [OF _ real_sqrt_sum_squares_triangle_ineq]) huffman@30019: apply (simp add: add_mono power_mono norm_triangle_ineq) huffman@30019: done huffman@30019: show "norm (scaleR r x) = \r\ * norm x" huffman@30019: unfolding norm_prod_def huffman@31587: apply (simp add: power_mult_distrib) webertj@49962: apply (simp add: distrib_left [symmetric]) huffman@30019: apply (simp add: real_sqrt_mult_distrib) huffman@30019: done huffman@30019: show "sgn x = scaleR (inverse (norm x)) x" huffman@30019: by (rule sgn_prod_def) huffman@31290: show "dist x y = norm (x - y)" huffman@31339: unfolding dist_prod_def norm_prod_def huffman@31339: by (simp add: dist_norm) huffman@30019: qed huffman@30019: huffman@30019: end huffman@30019: haftmann@37678: instance prod :: (banach, banach) banach .. huffman@31405: huffman@44575: subsubsection {* Pair operations are linear *} huffman@30019: huffman@44282: lemma bounded_linear_fst: "bounded_linear fst" huffman@44127: using fst_add fst_scaleR huffman@44127: by (rule bounded_linear_intro [where K=1], simp add: norm_prod_def) huffman@30019: huffman@44282: lemma bounded_linear_snd: "bounded_linear snd" huffman@44127: using snd_add snd_scaleR huffman@44127: by (rule bounded_linear_intro [where K=1], simp add: norm_prod_def) huffman@30019: huffman@30019: text {* TODO: move to NthRoot *} huffman@30019: lemma sqrt_add_le_add_sqrt: huffman@30019: assumes x: "0 \ x" and y: "0 \ y" huffman@30019: shows "sqrt (x + y) \ sqrt x + sqrt y" huffman@30019: apply (rule power2_le_imp_le) huffman@44749: apply (simp add: power2_sum x y) huffman@30019: apply (simp add: mult_nonneg_nonneg x y) huffman@44126: apply (simp add: x y) huffman@30019: done huffman@30019: huffman@30019: lemma bounded_linear_Pair: huffman@30019: assumes f: "bounded_linear f" huffman@30019: assumes g: "bounded_linear g" huffman@30019: shows "bounded_linear (\x. (f x, g x))" huffman@30019: proof huffman@30019: interpret f: bounded_linear f by fact huffman@30019: interpret g: bounded_linear g by fact huffman@30019: fix x y and r :: real huffman@30019: show "(f (x + y), g (x + y)) = (f x, g x) + (f y, g y)" huffman@30019: by (simp add: f.add g.add) huffman@30019: show "(f (r *\<^sub>R x), g (r *\<^sub>R x)) = r *\<^sub>R (f x, g x)" huffman@30019: by (simp add: f.scaleR g.scaleR) huffman@30019: obtain Kf where "0 < Kf" and norm_f: "\x. norm (f x) \ norm x * Kf" huffman@30019: using f.pos_bounded by fast huffman@30019: obtain Kg where "0 < Kg" and norm_g: "\x. norm (g x) \ norm x * Kg" huffman@30019: using g.pos_bounded by fast huffman@30019: have "\x. norm (f x, g x) \ norm x * (Kf + Kg)" huffman@30019: apply (rule allI) huffman@30019: apply (simp add: norm_Pair) huffman@30019: apply (rule order_trans [OF sqrt_add_le_add_sqrt], simp, simp) webertj@49962: apply (simp add: distrib_left) huffman@30019: apply (rule add_mono [OF norm_f norm_g]) huffman@30019: done huffman@30019: then show "\K. \x. norm (f x, g x) \ norm x * K" .. huffman@30019: qed huffman@30019: huffman@44575: subsubsection {* Frechet derivatives involving pairs *} huffman@30019: huffman@30019: lemma FDERIV_Pair: huffman@30019: assumes f: "FDERIV f x :> f'" and g: "FDERIV g x :> g'" huffman@30019: shows "FDERIV (\x. (f x, g x)) x :> (\h. (f' h, g' h))" huffman@44575: proof (rule FDERIV_I) huffman@44575: show "bounded_linear (\h. (f' h, g' h))" huffman@44575: using f g by (intro bounded_linear_Pair FDERIV_bounded_linear) huffman@44575: let ?Rf = "\h. f (x + h) - f x - f' h" huffman@44575: let ?Rg = "\h. g (x + h) - g x - g' h" huffman@44575: let ?R = "\h. ((f (x + h), g (x + h)) - (f x, g x) - (f' h, g' h))" huffman@44575: show "(\h. norm (?R h) / norm h) -- 0 --> 0" huffman@44575: proof (rule real_LIM_sandwich_zero) huffman@44575: show "(\h. norm (?Rf h) / norm h + norm (?Rg h) / norm h) -- 0 --> 0" huffman@44575: using f g by (intro tendsto_add_zero FDERIV_D) huffman@44575: fix h :: 'a assume "h \ 0" huffman@44575: thus "0 \ norm (?R h) / norm h" huffman@44575: by (simp add: divide_nonneg_pos) huffman@44575: show "norm (?R h) / norm h \ norm (?Rf h) / norm h + norm (?Rg h) / norm h" huffman@44575: unfolding add_divide_distrib [symmetric] huffman@44575: by (simp add: norm_Pair divide_right_mono huffman@44575: order_trans [OF sqrt_add_le_add_sqrt]) huffman@44575: qed huffman@44575: qed huffman@44575: huffman@44575: subsection {* Product is an inner product space *} huffman@44575: huffman@44575: instantiation prod :: (real_inner, real_inner) real_inner huffman@44575: begin huffman@44575: huffman@44575: definition inner_prod_def: huffman@44575: "inner x y = inner (fst x) (fst y) + inner (snd x) (snd y)" huffman@44575: huffman@44575: lemma inner_Pair [simp]: "inner (a, b) (c, d) = inner a c + inner b d" huffman@44575: unfolding inner_prod_def by simp huffman@44575: huffman@44575: instance proof huffman@44575: fix r :: real huffman@44575: fix x y z :: "'a::real_inner \ 'b::real_inner" huffman@44575: show "inner x y = inner y x" huffman@44575: unfolding inner_prod_def huffman@44575: by (simp add: inner_commute) huffman@44575: show "inner (x + y) z = inner x z + inner y z" huffman@44575: unfolding inner_prod_def huffman@44575: by (simp add: inner_add_left) huffman@44575: show "inner (scaleR r x) y = r * inner x y" huffman@44575: unfolding inner_prod_def webertj@49962: by (simp add: distrib_left) huffman@44575: show "0 \ inner x x" huffman@44575: unfolding inner_prod_def huffman@44575: by (intro add_nonneg_nonneg inner_ge_zero) huffman@44575: show "inner x x = 0 \ x = 0" huffman@44575: unfolding inner_prod_def prod_eq_iff huffman@44575: by (simp add: add_nonneg_eq_0_iff) huffman@44575: show "norm x = sqrt (inner x x)" huffman@44575: unfolding norm_prod_def inner_prod_def huffman@44575: by (simp add: power2_norm_eq_inner) huffman@44575: qed huffman@30019: huffman@30019: end huffman@44575: huffman@44575: end