# HG changeset patch # User huffman # Date 1314631869 25200 # Node ID c5e42b8590dd52b385ad25e35aae82040e1bcee6 # Parent bd91b77c4cd69f5ee7ece54b7239cc233f76e2a0 Product_Vector.thy: clean up some proofs diff -r bd91b77c4cd6 -r c5e42b8590dd src/HOL/Library/Product_Vector.thy --- a/src/HOL/Library/Product_Vector.thy Sun Aug 28 20:56:49 2011 -0700 +++ b/src/HOL/Library/Product_Vector.thy Mon Aug 29 08:31:09 2011 -0700 @@ -154,7 +154,65 @@ then show "\T. open T \ y \ T \ T \ snd ` S" by - (rule exI) qed -text {* Product preserves separation axioms. *} +subsubsection {* Continuity of operations *} + +lemma tendsto_fst [tendsto_intros]: + assumes "(f ---> a) F" + shows "((\x. fst (f x)) ---> fst a) F" +proof (rule topological_tendstoI) + fix S assume "open S" and "fst a \ S" + then have "open (fst -` S)" and "a \ fst -` S" + by (simp_all add: open_vimage_fst) + with assms have "eventually (\x. f x \ fst -` S) F" + by (rule topological_tendstoD) + then show "eventually (\x. fst (f x) \ S) F" + by simp +qed + +lemma tendsto_snd [tendsto_intros]: + assumes "(f ---> a) F" + shows "((\x. snd (f x)) ---> snd a) F" +proof (rule topological_tendstoI) + fix S assume "open S" and "snd a \ S" + then have "open (snd -` S)" and "a \ snd -` S" + by (simp_all add: open_vimage_snd) + with assms have "eventually (\x. f x \ snd -` S) F" + by (rule topological_tendstoD) + then show "eventually (\x. snd (f x) \ S) F" + by simp +qed + +lemma tendsto_Pair [tendsto_intros]: + assumes "(f ---> a) F" and "(g ---> b) F" + shows "((\x. (f x, g x)) ---> (a, b)) F" +proof (rule topological_tendstoI) + fix S assume "open S" and "(a, b) \ S" + then obtain A B where "open A" "open B" "a \ A" "b \ B" "A \ B \ S" + unfolding open_prod_def by fast + have "eventually (\x. f x \ A) F" + using `(f ---> a) F` `open A` `a \ A` + by (rule topological_tendstoD) + moreover + have "eventually (\x. g x \ B) F" + using `(g ---> b) F` `open B` `b \ B` + by (rule topological_tendstoD) + ultimately + show "eventually (\x. (f x, g x) \ S) F" + by (rule eventually_elim2) + (simp add: subsetD [OF `A \ B \ S`]) +qed + +lemma isCont_fst [simp]: "isCont f a \ isCont (\x. fst (f x)) a" + unfolding isCont_def by (rule tendsto_fst) + +lemma isCont_snd [simp]: "isCont f a \ isCont (\x. snd (f x)) a" + unfolding isCont_def by (rule tendsto_snd) + +lemma isCont_Pair [simp]: + "\isCont f a; isCont g a\ \ isCont (\x. (f x, g x)) a" + unfolding isCont_def by (rule tendsto_Pair) + +subsubsection {* Separation axioms *} lemma mem_Times_iff: "x \ A \ B \ fst x \ A \ snd x \ B" by (induct x) simp (* TODO: move elsewhere *) @@ -231,9 +289,6 @@ by (intro order_trans [OF _ real_sqrt_sum_squares_triangle_ineq] real_sqrt_le_mono add_mono power_mono dist_triangle2 zero_le_dist) next - (* FIXME: long proof! *) - (* Maybe it would be easier to define topological spaces *) - (* in terms of neighborhoods instead of open sets? *) fix S :: "('a \ 'b) set" show "open S \ (\x\S. \e>0. \y. dist y x < e \ y \ S)" proof @@ -264,111 +319,40 @@ thus "\e>0. \y. dist y x < e \ y \ S" .. qed next - assume "\x\S. \e>0. \y. dist y x < e \ y \ S" thus "open S" - unfolding open_prod_def open_dist - apply safe - apply (drule (1) bspec) - apply clarify - apply (subgoal_tac "\r>0. \s>0. e = sqrt (r\ + s\)") - apply clarify - apply (rule_tac x="{y. dist y a < r}" in exI) - apply (rule_tac x="{y. dist y b < s}" in exI) - apply (rule conjI) - apply clarify - apply (rule_tac x="r - dist x a" in exI, rule conjI, simp) - apply clarify - apply (simp add: less_diff_eq) - apply (erule le_less_trans [OF dist_triangle]) - apply (rule conjI) - apply clarify - apply (rule_tac x="s - dist x b" in exI, rule conjI, simp) - apply clarify - apply (simp add: less_diff_eq) - apply (erule le_less_trans [OF dist_triangle]) - apply (rule conjI) - apply simp - apply (clarify, rename_tac c d) - apply (drule spec, erule mp) - apply (simp add: dist_Pair_Pair add_strict_mono power_strict_mono) - apply (rule_tac x="e / sqrt 2" in exI, simp add: divide_pos_pos) - apply (rule_tac x="e / sqrt 2" in exI, simp add: divide_pos_pos) - apply (simp add: power_divide) - done + assume *: "\x\S. \e>0. \y. dist y x < e \ y \ S" show "open S" + proof (rule open_prod_intro) + fix x assume "x \ S" + then obtain e where "0 < e" and S: "\y. dist y x < e \ y \ S" + using * by fast + def r \ "e / sqrt 2" and s \ "e / sqrt 2" + from `0 < e` have "0 < r" and "0 < s" + unfolding r_def s_def by (simp_all add: divide_pos_pos) + from `0 < e` have "e = sqrt (r\ + s\)" + unfolding r_def s_def by (simp add: power_divide) + def A \ "{y. dist (fst x) y < r}" and B \ "{y. dist (snd x) y < s}" + have "open A" and "open B" + unfolding A_def B_def by (simp_all add: open_ball) + moreover have "x \ A \ B" + unfolding A_def B_def mem_Times_iff + using `0 < r` and `0 < s` by simp + moreover have "A \ B \ S" + proof (clarify) + fix a b assume "a \ A" and "b \ B" + hence "dist a (fst x) < r" and "dist b (snd x) < s" + unfolding A_def B_def by (simp_all add: dist_commute) + hence "dist (a, b) x < e" + unfolding dist_prod_def `e = sqrt (r\ + s\)` + by (simp add: add_strict_mono power_strict_mono) + thus "(a, b) \ S" + by (simp add: S) + qed + ultimately show "\A B. open A \ open B \ x \ A \ B \ A \ B \ S" by fast + qed qed qed end -subsection {* Continuity of operations *} - -lemma tendsto_fst [tendsto_intros]: - assumes "(f ---> a) net" - shows "((\x. fst (f x)) ---> fst a) net" -proof (rule topological_tendstoI) - fix S assume "open S" "fst a \ S" - then have "open (fst -` S)" "a \ fst -` S" - unfolding open_prod_def - apply simp_all - apply clarify - apply (rule exI, erule conjI) - apply (rule exI, rule conjI [OF open_UNIV]) - apply auto - done - with assms have "eventually (\x. f x \ fst -` S) net" - by (rule topological_tendstoD) - then show "eventually (\x. fst (f x) \ S) net" - by simp -qed - -lemma tendsto_snd [tendsto_intros]: - assumes "(f ---> a) net" - shows "((\x. snd (f x)) ---> snd a) net" -proof (rule topological_tendstoI) - fix S assume "open S" "snd a \ S" - then have "open (snd -` S)" "a \ snd -` S" - unfolding open_prod_def - apply simp_all - apply clarify - apply (rule exI, rule conjI [OF open_UNIV]) - apply (rule exI, erule conjI) - apply auto - done - with assms have "eventually (\x. f x \ snd -` S) net" - by (rule topological_tendstoD) - then show "eventually (\x. snd (f x) \ S) net" - by simp -qed - -lemma tendsto_Pair [tendsto_intros]: - assumes "(f ---> a) net" and "(g ---> b) net" - shows "((\x. (f x, g x)) ---> (a, b)) net" -proof (rule topological_tendstoI) - fix S assume "open S" "(a, b) \ S" - then obtain A B where "open A" "open B" "a \ A" "b \ B" "A \ B \ S" - unfolding open_prod_def by auto - have "eventually (\x. f x \ A) net" - using `(f ---> a) net` `open A` `a \ A` - by (rule topological_tendstoD) - moreover - have "eventually (\x. g x \ B) net" - using `(g ---> b) net` `open B` `b \ B` - by (rule topological_tendstoD) - ultimately - show "eventually (\x. (f x, g x) \ S) net" - by (rule eventually_elim2) - (simp add: subsetD [OF `A \ B \ S`]) -qed - -lemma isCont_fst [simp]: "isCont f a \ isCont (\x. fst (f x)) a" - unfolding isCont_def by (rule tendsto_fst) - -lemma isCont_snd [simp]: "isCont f a \ isCont (\x. snd (f x)) a" - unfolding isCont_def by (rule tendsto_snd) - -lemma isCont_Pair [simp]: - "\isCont f a; isCont g a\ \ isCont (\x. (f x, g x)) a" - unfolding isCont_def by (rule tendsto_Pair) - lemma Cauchy_fst: "Cauchy X \ Cauchy (\n. fst (X n))" unfolding Cauchy_def by (fast elim: le_less_trans [OF dist_fst_le]) @@ -451,43 +435,7 @@ instance prod :: (banach, banach) banach .. -subsection {* Product is an inner product space *} - -instantiation prod :: (real_inner, real_inner) real_inner -begin - -definition inner_prod_def: - "inner x y = inner (fst x) (fst y) + inner (snd x) (snd y)" - -lemma inner_Pair [simp]: "inner (a, b) (c, d) = inner a c + inner b d" - unfolding inner_prod_def by simp - -instance proof - fix r :: real - fix x y z :: "'a::real_inner * 'b::real_inner" - show "inner x y = inner y x" - unfolding inner_prod_def - by (simp add: inner_commute) - show "inner (x + y) z = inner x z + inner y z" - unfolding inner_prod_def - by (simp add: inner_add_left) - show "inner (scaleR r x) y = r * inner x y" - unfolding inner_prod_def - by (simp add: right_distrib) - show "0 \ inner x x" - unfolding inner_prod_def - by (intro add_nonneg_nonneg inner_ge_zero) - show "inner x x = 0 \ x = 0" - unfolding inner_prod_def prod_eq_iff - by (simp add: add_nonneg_eq_0_iff) - show "norm x = sqrt (inner x x)" - unfolding norm_prod_def inner_prod_def - by (simp add: power2_norm_eq_inner) -qed - -end - -subsection {* Pair operations are linear *} +subsubsection {* Pair operations are linear *} lemma bounded_linear_fst: "bounded_linear fst" using fst_add fst_scaleR @@ -533,29 +481,65 @@ then show "\K. \x. norm (f x, g x) \ norm x * K" .. qed -subsection {* Frechet derivatives involving pairs *} +subsubsection {* Frechet derivatives involving pairs *} lemma FDERIV_Pair: assumes f: "FDERIV f x :> f'" and g: "FDERIV g x :> g'" shows "FDERIV (\x. (f x, g x)) x :> (\h. (f' h, g' h))" -apply (rule FDERIV_I) -apply (rule bounded_linear_Pair) -apply (rule FDERIV_bounded_linear [OF f]) -apply (rule FDERIV_bounded_linear [OF g]) -apply (simp add: norm_Pair) -apply (rule real_LIM_sandwich_zero) -apply (rule tendsto_add_zero) -apply (rule FDERIV_D [OF f]) -apply (rule FDERIV_D [OF g]) -apply (rename_tac h) -apply (simp add: divide_nonneg_pos) -apply (rename_tac h) -apply (subst add_divide_distrib [symmetric]) -apply (rule divide_right_mono [OF _ norm_ge_zero]) -apply (rule order_trans [OF sqrt_add_le_add_sqrt]) -apply simp -apply simp -apply simp -done +proof (rule FDERIV_I) + show "bounded_linear (\h. (f' h, g' h))" + using f g by (intro bounded_linear_Pair FDERIV_bounded_linear) + let ?Rf = "\h. f (x + h) - f x - f' h" + let ?Rg = "\h. g (x + h) - g x - g' h" + let ?R = "\h. ((f (x + h), g (x + h)) - (f x, g x) - (f' h, g' h))" + show "(\h. norm (?R h) / norm h) -- 0 --> 0" + proof (rule real_LIM_sandwich_zero) + show "(\h. norm (?Rf h) / norm h + norm (?Rg h) / norm h) -- 0 --> 0" + using f g by (intro tendsto_add_zero FDERIV_D) + fix h :: 'a assume "h \ 0" + thus "0 \ norm (?R h) / norm h" + by (simp add: divide_nonneg_pos) + show "norm (?R h) / norm h \ norm (?Rf h) / norm h + norm (?Rg h) / norm h" + unfolding add_divide_distrib [symmetric] + by (simp add: norm_Pair divide_right_mono + order_trans [OF sqrt_add_le_add_sqrt]) + qed +qed + +subsection {* Product is an inner product space *} + +instantiation prod :: (real_inner, real_inner) real_inner +begin + +definition inner_prod_def: + "inner x y = inner (fst x) (fst y) + inner (snd x) (snd y)" + +lemma inner_Pair [simp]: "inner (a, b) (c, d) = inner a c + inner b d" + unfolding inner_prod_def by simp + +instance proof + fix r :: real + fix x y z :: "'a::real_inner \ 'b::real_inner" + show "inner x y = inner y x" + unfolding inner_prod_def + by (simp add: inner_commute) + show "inner (x + y) z = inner x z + inner y z" + unfolding inner_prod_def + by (simp add: inner_add_left) + show "inner (scaleR r x) y = r * inner x y" + unfolding inner_prod_def + by (simp add: right_distrib) + show "0 \ inner x x" + unfolding inner_prod_def + by (intro add_nonneg_nonneg inner_ge_zero) + show "inner x x = 0 \ x = 0" + unfolding inner_prod_def prod_eq_iff + by (simp add: add_nonneg_eq_0_iff) + show "norm x = sqrt (inner x x)" + unfolding norm_prod_def inner_prod_def + by (simp add: power2_norm_eq_inner) +qed end + +end