# HG changeset patch # User paulson # Date 1487697130 0 # Node ID 2cf841ff23be327778bc7feab941f04b842a5b1c # Parent ab7e11730ad8515e4e320d857c2ba4a52d51bcf3 some new material, also recasting some theorems using “obtains” diff -r ab7e11730ad8 -r 2cf841ff23be src/HOL/Analysis/Cauchy_Integral_Theorem.thy --- a/src/HOL/Analysis/Cauchy_Integral_Theorem.thy Tue Feb 21 15:04:01 2017 +0000 +++ b/src/HOL/Analysis/Cauchy_Integral_Theorem.thy Tue Feb 21 17:12:10 2017 +0000 @@ -38,6 +38,53 @@ finally show ?thesis . qed +lemma path_connected_arc_complement: + fixes \ :: "real \ 'a::euclidean_space" + assumes "arc \" "2 \ DIM('a)" + shows "path_connected(- path_image \)" +proof - + have "path_image \ homeomorphic {0..1::real}" + by (simp add: assms homeomorphic_arc_image_interval) + then + show ?thesis + apply (rule path_connected_complement_homeomorphic_convex_compact) + apply (auto simp: assms) + done +qed + +lemma connected_arc_complement: + fixes \ :: "real \ 'a::euclidean_space" + assumes "arc \" "2 \ DIM('a)" + shows "connected(- path_image \)" + by (simp add: assms path_connected_arc_complement path_connected_imp_connected) + +lemma inside_arc_empty: + fixes \ :: "real \ 'a::euclidean_space" + assumes "arc \" + shows "inside(path_image \) = {}" +proof (cases "DIM('a) = 1") + case True + then show ?thesis + using assms connected_arc_image connected_convex_1_gen inside_convex by blast +next + case False + show ?thesis + proof (rule inside_bounded_complement_connected_empty) + show "connected (- path_image \)" + apply (rule connected_arc_complement [OF assms]) + using False + by (metis DIM_ge_Suc0 One_nat_def Suc_1 not_less_eq_eq order_class.order.antisym) + show "bounded (path_image \)" + by (simp add: assms bounded_arc_image) + qed +qed + +lemma inside_simple_curve_imp_closed: + fixes \ :: "real \ 'a::euclidean_space" + shows "\simple_path \; x \ inside(path_image \)\ \ pathfinish \ = pathstart \" + using arc_simple_path inside_arc_empty by blast + + subsection \Piecewise differentiable functions\ definition piecewise_differentiable_on @@ -3998,29 +4045,29 @@ lemma winding_number_constant: assumes \: "path \" and loop: "pathfinish \ = pathstart \" and cs: "connected s" and sg: "s \ path_image \ = {}" - shows "\k. \z \ s. winding_number \ z = k" + obtains k where "\z. z \ s \ winding_number \ z = k" proof - - { fix y z - assume ne: "winding_number \ y \ winding_number \ z" - assume "y \ s" "z \ s" - then have "winding_number \ y \ \" "winding_number \ z \ \" - using integer_winding_number [OF \ loop] sg \y \ s\ by auto - with ne have "1 \ cmod (winding_number \ y - winding_number \ z)" + have *: "1 \ cmod (winding_number \ y - winding_number \ z)" + if ne: "winding_number \ y \ winding_number \ z" and "y \ s" "z \ s" for y z + proof - + have "winding_number \ y \ \" "winding_number \ z \ \" + using that integer_winding_number [OF \ loop] sg \y \ s\ by auto + with ne show ?thesis by (auto simp: Ints_def of_int_diff [symmetric] simp del: of_int_diff) - } note * = this - show ?thesis - apply (rule continuous_discrete_range_constant [OF cs]) + qed + have cont: "continuous_on s (\w. winding_number \ w)" using continuous_on_winding_number [OF \] sg - apply (metis Diff_Compl Diff_eq_empty_iff continuous_on_subset) - apply (rule_tac x=1 in exI) - apply (auto simp: *) - done + by (meson continuous_on_subset disjoint_eq_subset_Compl) + show ?thesis + apply (rule continuous_discrete_range_constant [OF cs cont]) + using "*" zero_less_one apply blast + by (simp add: that) qed lemma winding_number_eq: "\path \; pathfinish \ = pathstart \; w \ s; z \ s; connected s; s \ path_image \ = {}\ \ winding_number \ w = winding_number \ z" -using winding_number_constant by fastforce + using winding_number_constant by blast lemma open_winding_number_levelsets: assumes \: "path \" and loop: "pathfinish \ = pathstart \" diff -r ab7e11730ad8 -r 2cf841ff23be src/HOL/Analysis/Further_Topology.thy --- a/src/HOL/Analysis/Further_Topology.thy Tue Feb 21 15:04:01 2017 +0000 +++ b/src/HOL/Analysis/Further_Topology.thy Tue Feb 21 17:12:10 2017 +0000 @@ -4314,7 +4314,7 @@ qed next case False - have "\a. \x\S \ T. g x - h x = a" + obtain a where a: "\x. x \ S \ T \ g x - h x = a" proof (rule continuous_discrete_range_constant [OF ST]) show "continuous_on (S \ T) (\x. g x - h x)" apply (intro continuous_intros) @@ -4338,15 +4338,14 @@ then show ?thesis by (rule_tac x="2*pi" in exI) (fastforce simp add: algebra_simps) qed - qed - then obtain a where a: "\x. x \ S \ T \ g x - h x = a" by blast + qed blast with False have "exp a = 1" by (metis IntI disjoint_iff_not_equal divide_self_if exp_diff exp_not_eq_zero fg fh) with a show ?thesis apply (rule_tac x="\x. if x \ S then g x else a + h x" in exI) apply (intro continuous_on_cases_local_open opeS opeT contg conth continuous_intros conjI) apply (auto simp: algebra_simps fg fh exp_add) - done + done qed qed @@ -4383,7 +4382,7 @@ qed next case False - have "\a. \x\S \ T. g x - h x = a" + obtain a where a: "\x. x \ S \ T \ g x - h x = a" proof (rule continuous_discrete_range_constant [OF ST]) show "continuous_on (S \ T) (\x. g x - h x)" apply (intro continuous_intros) @@ -4407,15 +4406,14 @@ then show ?thesis by (rule_tac x="2*pi" in exI) (fastforce simp add: algebra_simps) qed - qed - then obtain a where a: "\x. x \ S \ T \ g x - h x = a" by blast + qed blast with False have "exp a = 1" by (metis IntI disjoint_iff_not_equal divide_self_if exp_diff exp_not_eq_zero fg fh) with a show ?thesis apply (rule_tac x="\x. if x \ S then g x else a + h x" in exI) apply (intro continuous_on_cases_local cloS cloT contg conth continuous_intros conjI) apply (auto simp: algebra_simps fg fh exp_add) - done + done qed qed @@ -4444,7 +4442,9 @@ then obtain f' where f': "\y. y \ T \ f' y \ S \ f (f' y) = y" by metis have *: "\a. \x \ {x. x \ S \ f x = y}. h x - h(f' y) = a" if "y \ T" for y - proof (rule continuous_discrete_range_constant, simp_all add: algebra_simps) + proof (rule continuous_discrete_range_constant [OF conn [OF that], of "\x. h x - h (f' y)"], simp_all add: algebra_simps) + show "continuous_on {x \ S. f x = y} (\x. h x - h (f' y))" + by (intro continuous_intros continuous_on_subset [OF conth]) auto show "\e>0. \u. u \ S \ f u = y \ h u \ h x \ e \ cmod (h u - h x)" if x: "x \ S \ f x = y" for x proof - @@ -4460,9 +4460,7 @@ then show ?thesis by (rule_tac x="2*pi" in exI) (fastforce simp add: algebra_simps) qed - show "continuous_on {x \ S. f x = y} (\x. h x - h (f' y))" - by (intro continuous_intros continuous_on_subset [OF conth]) auto - qed (simp add: conn that) + qed have "h x = h (f' (f x))" if "x \ S" for x using * [of "f x"] fim that apply clarsimp by (metis f' imageI right_minus_eq) diff -r ab7e11730ad8 -r 2cf841ff23be src/HOL/Analysis/Topology_Euclidean_Space.thy --- a/src/HOL/Analysis/Topology_Euclidean_Space.thy Tue Feb 21 15:04:01 2017 +0000 +++ b/src/HOL/Analysis/Topology_Euclidean_Space.thy Tue Feb 21 17:12:10 2017 +0000 @@ -7527,47 +7527,97 @@ text \Continuity implies uniform continuity on a compact domain.\ -lemma compact_uniformly_continuous: +subsection \Continuity implies uniform continuity on a compact domain.\ + +text\From the proof of the Heine-Borel theorem: Lemma 2 in section 3.7, page 69 of +J. C. Burkill and H. Burkill. A Second Course in Mathematical Analysis (CUP, 2002)\ + +lemma Heine_Borel_lemma: + assumes "compact S" and Ssub: "S \ \\" and op: "\G. G \ \ \ open G" + obtains e where "0 < e" "\x. x \ S \ \G \ \. ball x e \ G" +proof - + have False if neg: "\e. 0 < e \ \x \ S. \G \ \. \ ball x e \ G" + proof - + have "\x \ S. \G \ \. \ ball x (1 / Suc n) \ G" for n + using neg by simp + then obtain f where "\n. f n \ S" and fG: "\G n. G \ \ \ \ ball (f n) (1 / Suc n) \ G" + by metis + then obtain l r where "l \ S" "subseq r" and to_l: "(f \ r) \ l" + using \compact S\ compact_def that by metis + then obtain G where "l \ G" "G \ \" + using Ssub by auto + then obtain e where "0 < e" and e: "\z. dist z l < e \ z \ G" + using op open_dist by blast + obtain N1 where N1: "\n. n \ N1 \ dist (f (r n)) l < e/2" + using to_l apply (simp add: lim_sequentially) + using \0 < e\ half_gt_zero that by blast + obtain N2 where N2: "of_nat N2 > 2/e" + using reals_Archimedean2 by blast + obtain x where "x \ ball (f (r (max N1 N2))) (1 / real (Suc (r (max N1 N2))))" and "x \ G" + using fG [OF \G \ \\, of "r (max N1 N2)"] by blast + then have "dist (f (r (max N1 N2))) x < 1 / real (Suc (r (max N1 N2)))" + by simp + also have "... \ 1 / real (Suc (max N1 N2))" + apply (simp add: divide_simps del: max.bounded_iff) + using \subseq r\ seq_suble by blast + also have "... \ 1 / real (Suc N2)" + by (simp add: field_simps) + also have "... < e/2" + using N2 \0 < e\ by (simp add: field_simps) + finally have "dist (f (r (max N1 N2))) x < e / 2" . + moreover have "dist (f (r (max N1 N2))) l < e/2" + using N1 max.cobounded1 by blast + ultimately have "dist x l < e" + using dist_triangle_half_r by blast + then show ?thesis + using e \x \ G\ by blast + qed + then show ?thesis + by (meson that) +qed + +lemma compact_uniformly_equicontinuous: + assumes "compact S" + and cont: "\x e. \x \ S; 0 < e\ + \ \d. 0 < d \ + (\f \ \. \x' \ S. dist x' x < d \ dist (f x') (f x) < e)" + and "0 < e" + obtains d where "0 < d" + "\f x x'. \f \ \; x \ S; x' \ S; dist x' x < d\ \ dist (f x') (f x) < e" +proof - + obtain d where d_pos: "\x e. \x \ S; 0 < e\ \ 0 < d x e" + and d_dist : "\x x' e f. \dist x' x < d x e; x \ S; x' \ S; 0 < e; f \ \\ \ dist (f x') (f x) < e" + using cont by metis + let ?\ = "((\x. ball x (d x (e / 2))) ` S)" + have Ssub: "S \ \ ?\" + by clarsimp (metis d_pos \0 < e\ dist_self half_gt_zero_iff) + then obtain k where "0 < k" and k: "\x. x \ S \ \G \ ?\. ball x k \ G" + by (rule Heine_Borel_lemma [OF \compact S\]) auto + moreover have "dist (f v) (f u) < e" if "f \ \" "u \ S" "v \ S" "dist v u < k" for f u v + proof - + obtain G where "G \ ?\" "u \ G" "v \ G" + using k that + by (metis \dist v u < k\ \u \ S\ \0 < k\ centre_in_ball subsetD dist_commute mem_ball) + then obtain w where w: "dist w u < d w (e / 2)" "dist w v < d w (e / 2)" "w \ S" + by auto + with that d_dist have "dist (f w) (f v) < e/2" + by (metis \0 < e\ dist_commute half_gt_zero) + moreover + have "dist (f w) (f u) < e/2" + using that d_dist w by (metis \0 < e\ dist_commute divide_pos_pos zero_less_numeral) + ultimately show ?thesis + using dist_triangle_half_r by blast + qed + ultimately show ?thesis using that by blast +qed + +corollary compact_uniformly_continuous: fixes f :: "'a :: metric_space \ 'b :: metric_space" - assumes f: "continuous_on s f" - and s: "compact s" - shows "uniformly_continuous_on s f" - unfolding uniformly_continuous_on_def -proof (cases, safe) - fix e :: real - assume "0 < e" "s \ {}" - define R where [simp]: - "R = {(y, d). y \ s \ 0 < d \ ball y d \ s \ {x \ s. f x \ ball (f y) (e/2)}}" - let ?b = "(\(y, d). ball y (d/2))" - have "(\r\R. open (?b r))" "s \ (\r\R. ?b r)" - proof safe - fix y - assume "y \ s" - from continuous_openin_preimage_gen[OF f open_ball] - obtain T where "open T" and T: "{x \ s. f x \ ball (f y) (e/2)} = T \ s" - unfolding openin_subtopology open_openin by metis - then obtain d where "ball y d \ T" "0 < d" - using \0 < e\ \y \ s\ by (auto elim!: openE) - with T \y \ s\ show "y \ (\r\R. ?b r)" - by (intro UN_I[of "(y, d)"]) auto - qed auto - with s obtain D where D: "finite D" "D \ R" "s \ (\(y, d)\D. ball y (d/2))" - by (rule compactE_image) - with \s \ {}\ have [simp]: "\x. x < Min (snd ` D) \ (\(y, d)\D. x < d)" - by (subst Min_gr_iff) auto - show "\d>0. \x\s. \x'\s. dist x' x < d \ dist (f x') (f x) < e" - proof (rule, safe) - fix x x' - assume in_s: "x' \ s" "x \ s" - with D obtain y d where x: "x \ ball y (d/2)" "(y, d) \ D" - by blast - moreover assume "dist x x' < Min (snd`D) / 2" - ultimately have "dist y x' < d" - by (intro dist_triangle_half_r[of x _ d]) (auto simp: dist_commute) - with D x in_s show "dist (f x) (f x') < e" - by (intro dist_triangle_half_r[of "f y" _ e]) (auto simp: dist_commute subset_eq) - qed (insert D, auto) -qed auto + assumes f: "continuous_on S f" and S: "compact S" + shows "uniformly_continuous_on S f" + using f + unfolding continuous_on_iff uniformly_continuous_on_def + by (force intro: compact_uniformly_equicontinuous [OF S, of "{f}"]) subsection \Topological stuff about the set of Reals\ @@ -10207,6 +10257,46 @@ "\S. open S \ \U. U \ \ \ S = \U" by (metis ex_countable_basis topological_basis_def) +lemma subset_second_countable: + obtains \ :: "'a:: euclidean_space set set" + where "countable \" + "{} \ \" + "\C. C \ \ \ openin(subtopology euclidean S) C" + "\T. openin(subtopology euclidean S) T \ \\. \ \ \ \ T = \\" +proof - + obtain \ :: "'a set set" + where "countable \" + and opeB: "\C. C \ \ \ openin(subtopology euclidean S) C" + and \: "\T. openin(subtopology euclidean S) T \ \\. \ \ \ \ T = \\" + proof - + obtain \ :: "'a set set" + where "countable \" and ope: "\C. C \ \ \ open C" + and \: "\S. open S \ \U. U \ \ \ S = \U" + by (metis univ_second_countable that) + show ?thesis + proof + show "countable ((\C. S \ C) ` \)" + by (simp add: \countable \\) + show "\C. C \ op \ S ` \ \ openin (subtopology euclidean S) C" + using ope by auto + show "\T. openin (subtopology euclidean S) T \ \\\op \ S ` \. T = \\" + by (metis \ image_mono inf_Sup openin_open) + qed + qed + show ?thesis + proof + show "countable (\ - {{}})" + using \countable \\ by blast + show "\C. \C \ \ - {{}}\ \ openin (subtopology euclidean S) C" + by (simp add: \\C. C \ \ \ openin (subtopology euclidean S) C\) + show "\\\\ - {{}}. T = \\" if "openin (subtopology euclidean S) T" for T + using \ [OF that] + apply clarify + apply (rule_tac x="\ - {{}}" in exI, auto) + done + qed auto +qed + lemma univ_second_countable_sequence: obtains B :: "nat \ 'a::euclidean_space set" where "inj B" "\n. open(B n)" "\S. open S \ \k. S = \{B n |n. n \ k}" @@ -10244,6 +10334,52 @@ done qed +proposition separable: + fixes S :: "'a:: euclidean_space set" + obtains T where "countable T" "T \ S" "S \ closure T" +proof - + obtain \ :: "'a:: euclidean_space set set" + where "countable \" + and "{} \ \" + and ope: "\C. C \ \ \ openin(subtopology euclidean S) C" + and if_ope: "\T. openin(subtopology euclidean S) T \ \\. \ \ \ \ T = \\" + by (meson subset_second_countable) + then obtain f where f: "\C. C \ \ \ f C \ C" + by (metis equals0I) + show ?thesis + proof + show "countable (f ` \)" + by (simp add: \countable \\) + show "f ` \ \ S" + using ope f openin_imp_subset by blast + show "S \ closure (f ` \)" + proof (clarsimp simp: closure_approachable) + fix x and e::real + assume "x \ S" "0 < e" + have "openin (subtopology euclidean S) (S \ ball x e)" + by (simp add: openin_Int_open) + with if_ope obtain \ where \: "\ \ \" "S \ ball x e = \\" + by meson + show "\C \ \. dist (f C) x < e" + proof (cases "\ = {}") + case True + then show ?thesis + using \0 < e\ \ \x \ S\ by auto + next + case False + then obtain C where "C \ \" by blast + show ?thesis + proof + show "dist (f C) x < e" + by (metis Int_iff Union_iff \ \C \ \\ dist_commute f mem_ball subsetCE) + show "C \ \" + using \\ \ \\ \C \ \\ by blast + qed + qed + qed + qed +qed + proposition Lindelof: fixes \ :: "'a::euclidean_space set set" assumes \: "\S. S \ \ \ open S" @@ -10789,21 +10925,21 @@ qed lemma continuous_disconnected_range_constant_eq: - "(connected s \ + "(connected S \ (\f::'a::topological_space \ 'b::real_normed_algebra_1. - \t. continuous_on s f \ f ` s \ t \ (\y \ t. connected_component_set t y = {y}) - \ (\a::'b. \x \ s. f x = a)))" (is ?thesis1) + \t. continuous_on S f \ f ` S \ t \ (\y \ t. connected_component_set t y = {y}) + \ (\a::'b. \x \ S. f x = a)))" (is ?thesis1) and continuous_discrete_range_constant_eq: - "(connected s \ + "(connected S \ (\f::'a::topological_space \ 'b::real_normed_algebra_1. - continuous_on s f \ - (\x \ s. \e. 0 < e \ (\y. y \ s \ (f y \ f x) \ e \ norm(f y - f x))) - \ (\a::'b. \x \ s. f x = a)))" (is ?thesis2) + continuous_on S f \ + (\x \ S. \e. 0 < e \ (\y. y \ S \ (f y \ f x) \ e \ norm(f y - f x))) + \ (\a::'b. \x \ S. f x = a)))" (is ?thesis2) and continuous_finite_range_constant_eq: - "(connected s \ + "(connected S \ (\f::'a::topological_space \ 'b::real_normed_algebra_1. - continuous_on s f \ finite (f ` s) - \ (\a::'b. \x \ s. f x = a)))" (is ?thesis3) + continuous_on S f \ finite (f ` S) + \ (\a::'b. \x \ S. f x = a)))" (is ?thesis3) proof - have *: "\s t u v. \s \ t; t \ u; u \ v; v \ s\ \ (s \ t) \ (s \ u) \ (s \ v)" @@ -10822,19 +10958,19 @@ lemma continuous_discrete_range_constant: fixes f :: "'a::topological_space \ 'b::real_normed_algebra_1" - assumes s: "connected s" - and "continuous_on s f" - and "\x. x \ s \ \e>0. \y. y \ s \ f y \ f x \ e \ norm (f y - f x)" - shows "\a. \x \ s. f x = a" - using continuous_discrete_range_constant_eq [THEN iffD1, OF s] assms + assumes S: "connected S" + and "continuous_on S f" + and "\x. x \ S \ \e>0. \y. y \ S \ f y \ f x \ e \ norm (f y - f x)" + obtains a where "\x. x \ S \ f x = a" + using continuous_discrete_range_constant_eq [THEN iffD1, OF S] assms by blast lemma continuous_finite_range_constant: fixes f :: "'a::topological_space \ 'b::real_normed_algebra_1" - assumes "connected s" - and "continuous_on s f" - and "finite (f ` s)" - shows "\a. \x \ s. f x = a" + assumes "connected S" + and "continuous_on S f" + and "finite (f ` S)" + obtains a where "\x. x \ S \ f x = a" using assms continuous_finite_range_constant_eq by blast diff -r ab7e11730ad8 -r 2cf841ff23be src/HOL/Analysis/Uniform_Limit.thy --- a/src/HOL/Analysis/Uniform_Limit.thy Tue Feb 21 15:04:01 2017 +0000 +++ b/src/HOL/Analysis/Uniform_Limit.thy Tue Feb 21 17:12:10 2017 +0000 @@ -519,7 +519,7 @@ using lte by (force intro: eventually_mono) qed -lemma uniform_lim_div: +lemma uniform_lim_divide: fixes f :: "'a \ 'b \ 'c::real_normed_field" assumes f: "uniform_limit S f l F" and g: "uniform_limit S g m F"